Introducing Euclidean Relations to Mizar
Total Page:16
File Type:pdf, Size:1020Kb
Proceedings of the Federated Conference on DOI: 10.15439/2017F368 Computer Science and Information Systems pp. 245–248 ISSN 2300-5963 ACSIS, Vol. 11 Introducing Euclidean Relations to Mizar Adam Naumowicz, Artur Korniłowicz Institute of Informatics, University of Białystok ul. Ciołkowskiego 1 M, 15-245 Białystok, Poland Email: {adamn, arturk}@math.uwb.edu.pl Abstract—In this paper we present the methodology of imple- property has been analyzed and implemented most recently menting a new enhancement of the Mizar proof checker based on [15]. enabling special processing of Euclidean predicates, i.e. binary Table I presents how many registrations of predicate prop- predicates which fulfill a specific variant of transitivity postulated by Euclid. Typically, every proof step in formal mathematical erties are used in the MML. The data has been collected reasoning is associated with a formula to be proved and a list with Mizar Version 8.1.05 working with the MML Version of references used to justify the formula. With the proposed 5.37.1275.2 enhancement, the Euclidean property of given relations can be Property Occurrences Articles registered during their definition, and so the verification of some reflexivity 138 91 proof steps related to these relations can be automated to avoid irreflexivity 11 10 explicit referencing. symmetry 122 82 asymmetry 6 6 I. INTRODUCTION connectedness 4 4 total 281 119 HE Mizar system [1], [2], [3] is a computer proof- Table I T assistant system used for encoding formal mathematical PREDICATE PROPERTIES OCCURRENCES data (definitions and theorems) and verifying mathematical proofs. The system comprises a dedicated formal computer language – the Mizar language, a collection of command-line On the other hand, Table II shows the impact of registering tools including the VERIFIER and a repository of formal texts predicate properties for proofs stored in the library as the – Mizar Mathematical Library (MML) that have been written number of errors occurring after removing registrations of the in the Mizar language and machine-verified for their logical properties from texts, and the numbers of articles with such correctness by the VERIFIER. errors. The Mizar language preserves many features of natu- Property Errors Affected articles ral language mathematical writing. It is designed to write reflexivity 356 44 declarative-style documents both readable for humans and irreflexivity 9 2 effectively processed by computers.1 The feature-rich language symmetry 498 47 asymmetry 6 4 enables producing rigorous and semantically unambiguous connectedness 65 4 texts. Apart from rules for writing traditional mathematical total 934 73 items (e.g. definitions, lemmas, theorems, proof steps, etc.) Table II it also provides syntactic constructions to launch specialized PREDICATE PROPERTIES IMPACT algorithms for processing particular mechanisms (e.g. term identifications, term reductions [7], flexary connectives [8]) increasing computational power of VERIFIER (e.g. equational In this paper we propose strengthening the Mizar system calculus [9], [10]). There have also been experiments on using by implementing the representation of two new properties – 3 specialized external systems to increase computational power rightEuclidean and leftEuclidean. The properties of the Mizar system in selected domains [11], [12], [13]. are described in Section II. In Section III we present some Mizar allows registering various properties of predicates examples of Euclidean properties found in the current MML. and functors [14] when defining new notions. Let us briefly In Section IV we indicate several directions of further devel- recall that the set of currently implemented properties in- opment of processing properties in Mizar. cludes involutiveness and projectivity for unary II. THE EUCLIDEAN PROPERTY operations, as well as commutativity and idempotence An example of a property which does not have its corre- for binary operations. As far as binary predicates are con- sponding representation in Mizar yet is the Euclidean property. cerned, which is directly related to the topic of this pa- per, the current version of the Mizar system supports reg- 2Computations were carried out at the Computer Center of University of istering reflexivity, irreflexivity, symmetry, Białystok http://uco.uwb.edu.pl 3An experimental version of the VERIFIER executable file for the Linux asymmetry and connectedness. The transitivity platform implementing the new features as well as other supplementary resources required for it to work with the current MML are available at 1The legibility of proofs is a subject of ongoing research [4], [5], [6]. http://alioth.uwb.edu.pl/~artur/euclidean/euclidean_ver.zip. IEEE Catalog Number: CFP1785N-ART c 2017, PTI 245 246 PROCEEDINGS OF THE FEDCSIS. PRAGUE, 2017 This property appears, most notably, in the set of axiomatic It should be noted that the form of definitions and a corre- statements given at the start of Book I of Euclid’s The sponding correctness proof can in general be more complicated Elements [16] as Common Notion 1, which states that: Things when the definition has several cases. Below is a formal repre- which are equal to the same thing are also equal to each sentation of the correctness proof structure in the case of a def- 4 other. Formally speaking, a binary relation R on a set X is inition with three explicit cases (Γ1(x1, x2, . , xn, y1, y2), Euclidean (sometimes called right Euclidean) if it satisfies the Γ2(x1, x2, . , xn, y1, y2) and Γ3(x1, x2, . , xn, y1, y2)) as following condition: ∀a, b, c ∈ X (a R b ∧ a R c → b R c). well as the default case (introduced by otherwise): R X Dually, a relation on may be called left Euclidean if definition ∀a, b, c ∈ X (b R a ∧ c R a → b R c). let x1 be θ1, x2 be θ2, . , xn be θn, y1, y2 be θn+1; The property of being Euclidean is a variant of transitivity, pred π(y1, y2) means :ident: x , x , . , xn, y , y if x , x , . , xn, y , y , but the properties are indeed different. A transitive relation Φ1( 1 2 1 2) Γ1( 1 2 1 2) Φ2(x1, x2, . , xn, y1, y2) if Γ2(x1, x2, . , xn, y1, y2), is Euclidean only if it is also symmetric. Only a symmetric Φ3(x1, x2, . , xn, y1, y2) if Γ3(x1, x2, . , xn, y1, y2) Euclidean relation is transitive. A relation which is both otherwise Φn(x1, x2, . , xn, y1, y2); consistency; Euclidean and reflexive is also symmetric and therefore it is rightEuclidean an equivalence relation. proof In the sequel we describe an enhancement of the Mizar sys- thus for a, b, c being θn+1 st ( tem supporting automatic processing of Euclidean predicates. (Γ1(x1, x2, . , xn, a, b) implies Φ1(x1, x2, . , xn, a, b)) & The automation involves specific computations performed dur- (Γ2(x1, x2, . , xn, a, b) implies Φ2(x1, x2, . , xn, a, b)) & ing the verification process.5 (Γ3(x1, x2, . , xn, a, b) implies Φ3(x1, x2, . , xn, a, b)) & (not Γ1(x1, x2, . , xn, a, b) & To enable such an automation, when a new predicate is not Γ2(x1, x2, . , xn, a, b) & being defined, if it is Euclidean, it should be declared as such. not Γ3(x1, x2, . , xn, a, b) implies The declaration is placed within a definitional block with the Φn(x1, x2, . , xn, a, b)) & (Γ1(x1, x2, . , xn, a, c) implies Φ1(x1, x2, . , xn, a, c)) & following Mizar syntax for right Euclidean: (Γ2(x1, x2, . , xn, a, c) implies Φ2(x1, x2, . , xn, a, c)) & (Γ3(x1, x2, . , xn, a, c) implies Φ3(x1, x2, . , xn, a, c)) & definition (not Γ1(x1, x2, . , xn, a, c) & n n n let x1 be θ1, x2 be θ2, . , x be θ , y1, y2 be θ +1; not Γ2(x1, x2, . , xn, a, c) & ident pred π(y1, y2) means : : not Γ3(x1, x2, . , xn, a, c) implies n Φ(x1, x2, . , x , y1, y2); Φn(x1, x2, . , xn, a, c)) rightEuclidean ) holds proof ( , , n thus for a b c being θ +1 (Γ1(x1, x2, . , xn, b, c) implies Φ1(x1, x2, . , xn, b, c)) & n n st Φ(x1, x2, . , x , a, b) & Φ(x1, x2, . , x , a, c) (Γ2(x1, x2, . , xn, b, c) implies Φ2(x1, x2, . , xn, b, c)) & n holds Φ(x1, x2, . , x , b, c); (Γ3(x1, x2, . , xn, b, c) implies Φ3(x1, x2, . , xn, b, c)) & end; (not Γ1(x1, x2, . , xn, b, c) & end; not Γ2(x1, x2, . , xn, b, c) & not Γ3(x1, x2, . , xn, b, c) implies and left Euclidean, respectively: Φn(x1, x2, . , xn, b, c)) ); definition end; let x1 be θ1, x2 be θ2, . , xn be θn, y1, y2 be θn+1; end; pred π(y1, y2) means :ident: Φ(x1, x2, . , xn, y1, y2); The proof for a left Euclidean predicate with an analogous leftEuclidean proof set of conditions would look like this: thus for a, b, c being θn+1 definition st Φ(x1, x2, . , xn, b, a) & Φ(x1, x2, . , xn, c, a) let x be θ , x be θ , . , xn be θn, y , y be θn ; holds Φ(x1, x2, . , xn, b, c); 1 1 2 2 1 2 +1 ident end; pred π(y1, y2) means : : end; Φ1(x1, x2, . , xn, y1, y2) if Γ1(x1, x2, . , xn, y1, y2), Φ2(x1, x2, . , xn, y1, y2) if Γ2(x1, x2, . , xn, y1, y2), The statements of the formulas of correctness proofs shown Φ3(x1, x2, . , xn, y1, y2) if Γ3(x1, x2, . , xn, y1, y2) otherwise Φn(x1, x2, . , xn, y1, y2); in both the above definitions must be proved according to consistency; a special formula expressing the corresponding property of leftEuclidean the defined predicate. For example, having such a definition, proof thus for a, b, c being θn+1 st whenever VERIFIER meets a conjunction of formulas π(a, b) ( and π(a, c) within an inference, and the predicate π is known (Γ1(x1, x2, .