Basic Elements of Computer Algebra in MIZAR

Similar documents
Integrity of MML - Mizar Mathematical Library. We had a paper on integrity at MKM 03, but many rough edges have been ironed out by now.

How to Define Terms in Mizar Effectively

Properties of Subsets

Free Many Sorted Universal Algebra

Euler s Polyhedron Formula

Correctness of Dijkstra s Shortest Path and Prim s Minimum Spanning Tree Algorithms 1

Content Development for Distance Education in Advanced University Mathematics Using Mizar

CSC Discrete Math I, Spring Sets

CHAPTER 8. Copyright Cengage Learning. All rights reserved.

Writing a Mizar article in nine easy steps

Discrete Mathematics Lecture 4. Harper Langston New York University

Basic Functions and Operations on Functions

Relational Database: The Relational Data Model; Operations on Database Relations

SOFTWARE ENGINEERING DESIGN I

This is already grossly inconvenient in present formalisms. Why do we want to make this convenient? GENERAL GOALS

Families of Sets. Beata Padlewska 1 Warsaw University Bia lystok

Sets 1. The things in a set are called the elements of it. If x is an element of the set S, we say

Semantics via Syntax. f (4) = if define f (x) =2 x + 55.

3.4 Deduction and Evaluation: Tools Conditional-Equational Logic

CSCI.6962/4962 Software Verification Fundamental Proof Methods in Computer Science (Arkoudas and Musser) Sections p.

Formalization and Verification of Number Theoretic Algorithms Using the Mizar Proof Checker

The Further Mathematics Support Programme

One of the most important areas where quantifier logic is used is formal specification of computer programs.

A set with only one member is called a SINGLETON. A set with no members is called the EMPTY SET or 2 N

CSC 501 Semantics of Programming Languages

A Simplified Abstract Syntax for the Dataflow Algebra. A. J. Cowling

Introduction to Homotopy Type Theory

(a) (4 pts) Prove that if a and b are rational, then ab is rational. Since a and b are rational they can be written as the ratio of integers a 1

Lecture 5. Logic I. Statement Logic

Variable, Complement, and Literal are terms used in Boolean Algebra.

Boolean Algebra. P1. The OR operation is closed for all x, y B x + y B

Computing Fundamentals 2 Introduction to CafeOBJ

9/19/12. Why Study Discrete Math? What is discrete? Sets (Rosen, Chapter 2) can be described by discrete math TOPICS

BOOLEAN ALGEBRA AND CIRCUITS

Reasoning About Programs Panagiotis Manolios

CHAPTER 7. Copyright Cengage Learning. All rights reserved.

Programming Languages Third Edition

9.5 Equivalence Relations

Towards a Logical Reconstruction of Relational Database Theory

Propositional Calculus: Boolean Algebra and Simplification. CS 270: Mathematical Foundations of Computer Science Jeremy Johnson

Mathematically Rigorous Software Design Review of mathematical prerequisites

ROUGH MEMBERSHIP FUNCTIONS: A TOOL FOR REASONING WITH UNCERTAINTY

Axiomatic Specification. Al-Said, Apcar, Jerejian

Mizar Users Group An Outline of PC Mizar Micha l Muzalewski Series Editor Roman Matuszewski Fondation Philippe le Hodey Brussels, 1993

Lecture Notes on Real-world SMT

Appendix 1. Description Logic Terminology

Introduction to Sets and Logic (MATH 1190)

Appendix 1. Description Logic Terminology

THREE LECTURES ON BASIC TOPOLOGY. 1. Basic notions.

Propositional Logic. Part I

CSCI.6962/4962 Software Verification Fundamental Proof Methods in Computer Science (Arkoudas and Musser) Chapter p. 1/27

Complexity Theory. Compiled By : Hari Prasad Pokhrel Page 1 of 20. ioenotes.edu.np

1. M,M sequential composition: try tactic M; if it succeeds try tactic M. sequential composition (, )

Lecture 5: The Halting Problem. Michael Beeson

MATH 22 MORE ABOUT FUNCTIONS. Lecture M: 10/14/2003. Form follows function. Louis Henri Sullivan

CS 6110 S14 Lecture 38 Abstract Interpretation 30 April 2014

arxiv: v1 [cs.ms] 7 May 2015

CS 1200 Discrete Math Math Preliminaries. A.R. Hurson 323 CS Building, Missouri S&T

An Introduction to Programming and Proving in Agda (incomplete draft)

What Is A Relation? Example. is a relation from A to B.

An Evolution of Mathematical Tools

A.1 Numbers, Sets and Arithmetic

Slides for Faculty Oxford University Press All rights reserved.

Abstract Interpretation

Summary of Course Coverage

Lecture (04) Boolean Algebra and Logic Gates

Lecture (04) Boolean Algebra and Logic Gates By: Dr. Ahmed ElShafee

CSE 20 DISCRETE MATH. Fall

Rational Numbers CHAPTER Introduction

Natural Numbers. We will use natural numbers to illustrate several ideas that will apply to Haskell data types in general.

Jaykov Foukzon. Israel Institute of Technology, Haifa, Israel.

Notes on metric spaces and topology. Math 309: Topics in geometry. Dale Rolfsen. University of British Columbia

[Ch 6] Set Theory. 1. Basic Concepts and Definitions. 400 lecture note #4. 1) Basics

Algebraic Expressions

Categorical models of type theory

Propositional Calculus. CS 270: Mathematical Foundations of Computer Science Jeremy Johnson

Evaluation Trees for Proposition Algebra

THE FOUNDATIONS OF MATHEMATICS

Rewriting. Andreas Rümpel Faculty of Computer Science Technische Universität Dresden Dresden, Germany.

AXIOMS FOR THE INTEGERS

A fuzzy subset of a set A is any mapping f : A [0, 1], where [0, 1] is the real unit closed interval. the degree of membership of x to f

Foundations of Computer Science Spring Mathematical Preliminaries

Math 221 Final Exam Review

Chapter 4. Relations & Graphs. 4.1 Relations. Exercises For each of the relations specified below:

From Types to Sets in Isabelle/HOL

Lecture Notes on Program Equivalence

Associative Operations on a Three-Element Set

EDAA40 At home exercises 1

4. Simplicial Complexes and Simplicial Homology

Post and Jablonsky Algebras of Compositions (Superpositions)

Copyright 2011 Pearson Education, Inc. Publishing as Pearson Addison-Wesley. Chapter 6 Outline. Unary Relational Operations: SELECT and

Denotational Semantics. Domain Theory

Information Granulation and Approximation in a Decision-theoretic Model of Rough Sets

CSE 20 DISCRETE MATH. Winter

1 Introduction CHAPTER ONE: SETS

An Improved Upper Bound for the Sum-free Subset Constant

PROGRAMMING IN HASKELL. Chapter 2 - First Steps

A GRAPH THEORETICAL APPROACH TO MONOGENIC AND STRONGLY MONOGENIC RIGHT TERNARY N-GROUPS

The three faces of homotopy type theory. Type theory and category theory. Minicourse plan. Typing judgments. Michael Shulman.

Algebraic Processors

Transcription:

Basic Elements of Computer Algebra in MIZAR Adam Naumowicz Czeslaw Bylinski Institute of Computer Science University of Bialystok, Poland adamn@math.uwb.edu.pl Section of Computer Networks University of Bialystok, Poland bylinski@math.uwb.edu.pl Abstract - In this paper we describe special features of the Mizar system which provide some elements of computer algebra and present how they strengthen the capabilities of the Mizar checker. Keywords Mizar checker, computer algebra, requirements directive, properties of Mizar functors and predicates. Introduction The original goal of the Mizar project was to design and implement a software environment that supports writing traditional mathematics papers. Mathematical practice shows that even in formal proofs some easy background reasoning can be reduced. There are many powerful systems that efficiently process numeric and symbolic computation. Similar techniques incorporated into the Mizar system would considerably benefit the Mizar user community. At the moment, the inference checker uses model elimination with stress on processing speed, not power. However, its power can be extended in several ways. In this paper we discuss how properties that can be associated with Mizar s and the requirements directive can strengthen the process of inference justification in Mizar. Both these features influence how equality classes are generated in the EQUALIZER - the module responsible for the equality calculus in the Mizar checker (cf. [3]). Their effects can substantially reduce the amount of justification an author must provide in a proof. Used in connection with suitable management utilities these features stimulate the growth and evolution of the Mizar Mathematical Library (MML).. Properties As described in [0], there are four main kinds of constructors in Mizar: predicates, functors, modes and attributes. The Mizar system allows for special automated processing of certain properties of the first two types. The properties currently implemented for predicates (constructors of formulae) include: symmetry, asymmetry, reflexivity, irreflexivity, and connectedness. The properties for functors (constructors of terms) are: commutativity, idempotence, involutiveness, and projectivity. When Manuscript received October, 00; revised July 3, 00. The MML is the data base of Mizar articles. The systematic collection started in 989. At the time of this writing it contains 74 articles (about 54 MB of Mizar texts).

included in a of a predicate or a functor, the above-mentioned properties can be automatically used by the Mizar checker in every inference step which concerns that constructor. In that case, corresponding statements and references to these statements become superfluous. The properties are paired with a justification of suitable correctness conditions which we describe below. We also discuss the restrictions which are necessary to avoid a collapse of system consistency.. Predicate Properties In general, a Mizar predicate with properties is defined as: let x be θ ; let y be θ ; pred Example_Pred x,y means δ (x,y); :: the definiens of the predicate predicate-property-symbol proof...... let x be θ ; let y be θ ; pred Example_Pred x,y means δ (x,y); :: the definiens of the predicate predicate-property-symbol proof...... where predicate-property-symbol is one of the following: asymmetry, symmetry, reflexivity, irreflexivity, and connectedness. The properties are accepted only when the types θ and θ are equal. The following table contains a summary of predicate properties with suitable justification formulae. Examples of all properties taken from MML are presented below. Predicate property Formula to be proved as justification asymmetry for x,y being θ holds δ (x,y) implies not δ (y,x) symmetry for x,y being θ holds δ (x,y) implies δ (y,x) reflexivity for x being θ holds δ (x,x) irreflexivity for x being θ holds not δ (x,x) connectedness for x,y being θ holds not δ (x,y) implies δ (y,x) We illustrate asymmetry with the Mizar primitive in predicate. This predicate has no accompanying justification because it is built into the Mizar system. The article HIDDEN ([6]) documents built-in notions. let x,x be set; pred x in X;

asymmetry; As an example of the symmetry property, we show a predicate satisfied whenever two sets have an empty intersection (XBOOLE_0:def 7, [4]). It sometimes happens, as in this example, that the condition is obvious for the checker and no justification is needed. let X,Y be set; pred X misses Y means :Def7: X Y = {}; symmetry; antonym X meets Y; An example of reflexivity is the divisibility relation for natural numbers (NAT_:def 3, []) presented below: let k,l be natural number; pred k divides l means :Def3: ex t being natural number st l = k * t; reflexivity proof let i be natural number; i = i * ; hence thesis; An example of a predicate with irreflexivity is the proper inclusion of sets (XBOOLE_0:def 8, [4]). let X,Y be set; pred X c< Y means :Def8: X c= Y & X <> Y; irreflexivity; We demonstrate connectedness with the re of inclusion for ordinal numbers (ORDINAL, []). let A,B be Ordinal; redefine pred A c= B; connectedness proof let A,B be Ordinal; A in B or A = B or B in A by Th4; hence thesis by Def; The phrase Article-Identifier:def Definition-Number follows the convention which identifies all Mizar s in the MML.

Here, Th4 and Def refer to: theorem Th4: for A,B being Ordinal holds A in B or A = B or B in A let X be set; attr X is epsilon-transitive means :Def: for x being set st x in X holds x c= X; We note that a similar concept could also be implemented for modes since they are in fact special kinds of predicates. For example, reflexivity seems useful for a mode constructor like Subset of. Also, the set of currently implemented predicate properties is not purely accidental. Since every Mizar predicate can have an antonym, each property has a counterpart related to the antonym. For example, reflexivity automatically means irreflexivity for an antonym and vice versa. The same can be said for the pair connectedness and asymmetry. Obviously, symmetry of an original constructor and its antonym are equivalent.. Functor Properties The properties of binary functors in Mizar are commutativity and idempotence. In general, we define a binary functor with properties in the following form: let x be θ ; let y be θ ; func Example_Func(x,y) -> θ means 3 δ (it,x,y); binary-functor-property-symbol proof...... where binary-functor-property-symbol is commutativity or idempotence, and the Mizar reserved word 'it' in the definiens denotes the value of the functor being defined. Binary functor property Formula to be proved as justification commutativity for x being θ, y being θ, z being θ 3 holds δ (x,y,z) implies δ (x,z,y) idempotence for x being θ holds δ (x,x,x) An example showing both binary functor properties is the set theoretical join operator (XBOOLE_0:def, [4]). let X,Y be set; func X Y -> set means :Def: x in it iff x in X or x in Y; existence proof... uniqueness proof...

commutativity; idempotence; With the current implementation, commutativity is only applicable to functors for which the result type is invariant under swapping arguments. Furthermore, idempotence requires that the result type be wider than the type of the argument (or equal to it). The Mizar unary functor with properties uses the form below: let x be θ ; func Example_Func(x) -> θ means δ (it,x); unary-functor-property-symbol proof...... where unary-functor-property-symbol is involutiveness or projectivity. The system consistency is protected by the restriction that types θ and θ be equal. Unary functor property Formula to be proved as justification involutiveness for x,y being θ holds δ (x,y) implies δ (y,x) projectivity for x,y being θ holds δ (x,y) implies δ (x,x) The involutiveness property is used with the inverse relation (RELAT_:def 7, [4]). let R be Relation; func R~ -> Relation means :Def7: [x,y] in it iff [y,x] in R; existence proof... uniqueness proof... involutiveness; As an example of projectivity we give the functor for generating the absolute value of a real number (ABSVALUE:def, [9]). let x be real number; func abs x -> real number equals :Def: x if 0 <= x otherwise -x; coherence; consistency; projectivity by REAL_:66; Here, REAL_:66 ([8]) refers to: theorem :: REAL_:66 for x being real number holds x < 0 iff 0 < -x;

Due to some problems in implementation, the idempotence, involutiveness, and projectivity properties are not available for redefined objects as yet. 3. Requirements The requirements directive, which is comparatively new in Mizar 3 allows for special processing of selected constructors. Unlike the properties described in Section, it concerns the environ part of a Mizar article (cf. [0]). With the requirements directive, some built-in concepts for selected constructors will be imported during the accommodation stage of processing an article. In the MML database they are encoded in special files with extension '.dre'. As yet, the special files in use are: HIDDEN, BOOLE, SUBSET, ARYTM, and REAL. We describe how they assist the Mizar checker with the work of reasoning so that the amount of justification an author must provide can be reduced. 3. requirements HIDDEN This directive is automatically included during accommodation of every article and therefore does not need to be used explicitly. It identifies the objects defined in the axiomatic file HIDDEN, i.e., the mode set followed by the =' and 'in' predicates (HIDDEN:def - HIDDEN:def 3, [6]). Mode set is the most general Mizar mode and every other mode widens to it. Thanks to the identification provided by requirements HIDDEN it is used internally wherever the most basic Mizar type is needed, e.g., while generating various correctness conditions. The fundamental equality predicate '=' is extensional which means that two objects of the same kind (atomic formulae, types, functors, attributes) are equal when their arguments are equal. This particular property is used frequently by the Mizar checker. The '=' relation is also symmetric, reflexive, and transitive. Predicate 'in' plays an important role in the unfolding of sentences with the Fraenkel operator in positive and negative contexts. This allows sentences of the form ex y being θ st x=y & P[y] to be true whenever x in {y being θ : P[y]} is true and vice versa. Various features of the 'in' predicate are considered in conjunction with other requirements (see Sections 3., 3.3). 3. requirements BOOLE When processing an article with requirements BOOLE, Mizar treats specially the constructors provided for: the empty set ({}), attribute empty, set theoretical join ( ), meet ( ), difference ( ), and symmetric difference ( + ) given in s XBOOLE_0:def -XBOOLE_0:def 6, [4]). It allows the following frequently used equations to be accepted without any external justification: X {} = X, X {} = {}, X {} = X, {} X = {}, and {} + X = X. The empty set also gets additional properties: x is empty implies x = {}, and similarly x in X implies X is non empty. Additional features concerning the empty set are also described in the next section 4. 3 Historically, the first requirements directive was ARYTM, introduced in 995. The most recent is BOOLE, implemented in 00. 4 Recently, the Library Committee decided to provide a special article covering the proofs of requirements which can be formulated as Mizar statements. The first article of that series is BOOLE, [5].

3.3 requirements SUBSET This requirements directive concerns the of inclusion (TARSKI:def 3, []), the power set (ZFMISC_:def, [3]) and also mode 'Element of' (SUBSET_:def ) with a following re, []). With this directive X c= Y automatically yields X is Subset of Y and vice versa. The property of the form x in X & X c= Y implies x in Y is incorporated as well 5. When BOOLE is also applied in the requirements directive, the formula x in X is equivalent to x is Element of X & X is non empty. 3.4 requirements ARYTM Specification of requirements ARYTM concerns the s provided in the article ARYTM ([7]): the set REAL (ARYTM:def ), the re of the set NAT as a Subset of REAL, the real addition and multiplication operations (ARYTM:def 3, def 4) and the natural ordering of real numbers (ARYTM:def 5). It provides the correspondence between numerals and numbers defined in the MML. Without requirements ARYTM and SUBSET, numerals are just names for (not fixed) sets. With them, numerals obtain internally the type Element of NAT and appropriate values stored as rational numbers. These values are also used to assign a proper order between numerals. It also makes basic addition and multiplication operations on rational numbers accepted by the checker with no additional justification. The numerators and denominators of Mizar numerals must not be greater than 3767 (the maximum value for a 6-bit signed integer), although all internal calculations are based on 3-bit integers. For example, the following equalities can be easily calculated and therefore they are obvious: x + 0 = x, x * 0 = 0, x * = x. More processing capabilities for other arithmetic operations are introduced by the requirements REAL directive (see Section 3.5). 3.5 requirements REAL This enables special processing of real expressions based on the constructors REAL_:def - REAL_:def 4, [8]. As with ARYTM, the Mizar checker uses the rational value associated with real variables and a built-in GCD routine to evaluate new equalities. In particular, the following equalities are directly calculated at this stage: x / = x, x - 0 = x, etc. 4 Conclusions The above considerations show that there are quite efficient mechanisms in Mizar that provide some elements of computer algebra. The idea of implementing such features was based on statistical observations showing the extensive use of special constructs. The introduction of these techniques has a strong influence on the maintenance of the MML. However, the distinction between the function of properties and requirements is not always clear. In some cases, it is hard to decide which of these techniques is the best implementation. Requirements are much more flexible, but on the other hand, properties are a regular language construct and are not so system dependent. Every Mizar user can decide whether or not to use properties for newly created s while a new requirements directive yields a partial reimplementation of the system. Some of the features currently implemented as requirements could be transformed into some kind of 5 Formerly, it was an extensively used MML theorem BOOLE:.

properties. In particular, this would concern the properties of neutral elements as described in Sections 3.4 and 3.5. There is still discussion on what would be the best syntax for such a neutrality property. Another area of interest is the implementation of the associativity and transitivity properties. However, this work still remains in the to-do list due to some problems with finding an approach that will generate an efficient implementation. Acknowledgment: The authors would like to express their gratitude to Pauline N. Kawamoto for her kind help in preparation of this paper. References [] Grzegorz Bancerek, The fundamental properties of natural numbers, Journal of Formalized Mathematics, http://mizar.org/jfm/vol/nat_.html. [] Grzegorz Bancerek, The ordinal numbers, Journal of Formalized Mathematics, http://mizar.org/jfm/vol/ordinal.html. [3] Czeslaw Bylinski, Some basic properties of sets, Journal of Formalized Mathematics, http://mizar.org/jfm/vol/zfmisc_.html. [4] Library Committee, Boolean Properties of Sets Definitions, Journal of Formalized Mathematics, Encyclopedia of Mathematics in Mizar, 00, http://mizar.org/jfm/emm/xboole_0.html. [5] Library Committee, Boolean Properties of Sets Requirements, Journal of Formalized Mathematics, Encyclopedia of Mathematics in Mizar, 00, http://mizar.org/jfm/emm/boole.html. [6] Library Committee, Mizar built-in notions, Journal of Formalized Mathematics, Axiomatics, 989, http://mizar.org/jfm/axiomatics/hidden.html. [7] Library Committee, Preliminaries to arithmetic, Journal of Formalized Mathematics, Addenda, 995, http://mizar.org/jfm/addenda/arytm.html. [8] Krzysztof Hryniewiecki, Basic properties of real numbers, Journal of Formalized Mathematics, http://mizar.org/jfm/vol/real_.html. [9] Jan Popiolek, Some properties of functions modul and signum, Journal of Formalized Mathematics, http://mizar. org/jfm/vol/absvalue.html. [0] Piotr Rudnicki and Andrzej Trybulec, On Equivalents of Well-Foundedness. An Experiment in MIZAR, Journal of Automated Reasoning, 3, 999, pp. 97-34. [] Andrzej Trybulec, Tarski Grothendieck set theory, Journal of Formalized Mathematics, http://mizar.org/jfm/axiomatics/tarski.html. [] Zinaida Trybulec, Properties of subsets, Journal of Formalized Mathematics, http://mizar.org/jfm/vol/subset_.html. [3] Freek Wiedijk, Checker, available on WWW: http://www.cs.kun.nl/~freek/notes/by.ps.gz. [4] Edmund Woronowicz, Relations and their basic properties, Journal of Formalized Mathematics, http://mizar.org/jfm/vol/relat_.html.