# HG changeset patch # User blanchet # Date 1236161129 -3600 # Node ID aea5d7fa7ef5c14c4caf42ba85c4682d90484755 # Parent 3a1aef73b2b20df70467434d8ca057e1027e2cf2# Parent e70dae49dc57df4ec78a4af7c97545aedc35fc54 Merge. diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 Admin/CHECKLIST --- a/Admin/CHECKLIST Wed Mar 04 11:05:02 2009 +0100 +++ b/Admin/CHECKLIST Wed Mar 04 11:05:29 2009 +0100 @@ -1,7 +1,7 @@ Checklist for official releases =============================== -- test alice, mosml, polyml-5.0, polyml-4.1.3, polyml-4.1.4, polyml-4.2.0, x86-solaris, x86-cygwin; +- test mosml, polyml-5.0, polyml-4.1.3, polyml-4.1.4, polyml-4.2.0, x86-solaris, x86-cygwin; - test ProofGeneral; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 Admin/isatest/isatest-stats --- a/Admin/isatest/isatest-stats Wed Mar 04 11:05:02 2009 +0100 +++ b/Admin/isatest/isatest-stats Wed Mar 04 11:05:29 2009 +0100 @@ -16,6 +16,7 @@ HOL-Algebra \ HOL-Auth \ HOL-Bali \ + HOL-Decision_Procs \ HOL-Extraction \ HOL-Hoare \ HOL-HoareParallel \ diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 Admin/isatest/settings/sun-poly --- a/Admin/isatest/settings/sun-poly Wed Mar 04 11:05:02 2009 +0100 +++ b/Admin/isatest/settings/sun-poly Wed Mar 04 11:05:29 2009 +0100 @@ -4,7 +4,7 @@ ML_SYSTEM="polyml-5.1" ML_PLATFORM="sparc-solaris" ML_HOME="$POLYML_HOME/$ML_PLATFORM" - ML_OPTIONS="-H 1500" + ML_OPTIONS="-H 800" ISABELLE_HOME_USER=/tmp/isabelle-sun-poly diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 Admin/makedist --- a/Admin/makedist Wed Mar 04 11:05:02 2009 +0100 +++ b/Admin/makedist Wed Mar 04 11:05:29 2009 +0100 @@ -4,7 +4,7 @@ ## global settings -REPOS="https://isabelle.in.tum.de/repos/isabelle" +REPOS="http://isabelle.in.tum.de/repos/isabelle" DISTPREFIX=${DISTPREFIX:-~/tmp/isadist} @@ -156,7 +156,7 @@ rm doc/codegen_process.pdf rm -rf doc-src -mkdir contrib +mkdir -p contrib cp doc/isabelle*.eps lib/logo diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 CONTRIBUTORS --- a/CONTRIBUTORS Wed Mar 04 11:05:02 2009 +0100 +++ b/CONTRIBUTORS Wed Mar 04 11:05:29 2009 +0100 @@ -7,10 +7,16 @@ Contributions to this Isabelle version -------------------------------------- -* February 2008: Timothy Bourke, NICTA +* February 2009: Filip Maric, Univ. of Belgrade + A Serbian theory. + +* February 2009: Jasmin Christian Blanchette, TUM + Misc cleanup of HOL/refute. + +* February 2009: Timothy Bourke, NICTA New find_consts command. -* February 2008: Timothy Bourke, NICTA +* February 2009: Timothy Bourke, NICTA "solves" criterion for find_theorems and auto_solve option * December 2008: Clemens Ballarin, TUM @@ -31,6 +37,9 @@ processes. Additional ATP wrappers, including remote SystemOnTPTP services. +* September 2008: Stefan Berghofer, TUM and Marc Bezem, Univ. Bergen + Prover for coherent logic. + * August 2008: Fabian Immler, TUM Vampire wrapper script for remote SystemOnTPTP service. @@ -46,7 +55,7 @@ HOLCF library improvements. * 2007/2008: Stefan Berghofer, TUM - HOL-Nominal package improvements. + HOL-Nominal package improvements. * March 2008: Markus Reiter, TUM HOL/Library/RBT: red-black trees. diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 NEWS --- a/NEWS Wed Mar 04 11:05:02 2009 +0100 +++ b/NEWS Wed Mar 04 11:05:29 2009 +0100 @@ -6,6 +6,10 @@ *** General *** +* The main reference manuals (isar-ref, implementation, system) have +been updated and extended. Formally checked references as hyperlinks +are now available in uniform manner. + * Simplified main Isabelle executables, with less surprises on case-insensitive file-systems (such as Mac OS). @@ -47,9 +51,6 @@ regular 4-core machine, if the initial heap space is made reasonably large (cf. Poly/ML option -H). [Poly/ML 5.2.1 or later] -* The Isabelle System Manual (system) has been updated, with formally -checked references as hyperlinks. - * Generalized Isar history, with support for linear undo, direct state addressing etc. @@ -63,6 +64,8 @@ * There is a new syntactic category "float_const" for signed decimal fractions (e.g. 123.45 or -123.45). +* New prover for coherent logic (see src/Tools/coherent.ML). + *** Pure *** @@ -111,30 +114,32 @@ unify_trace_bound = 50 (formerly 25) unify_search_bound = 60 (formerly 30) -* Different bookkeeping for code equations: - a) On theory merge, the last set of code equations for a particular constant - is taken (in accordance with the policy applied by other parts of the - code generator framework). - b) Code equations stemming from explicit declarations (e.g. code attribute) - gain priority over default code equations stemming from definition, primrec, - fun etc. - INCOMPATIBILITY. - -* Global versions of theorems stemming from classes do not carry -a parameter prefix any longer. INCOMPATIBILITY. +* Different bookkeeping for code equations (INCOMPATIBILITY): + + a) On theory merge, the last set of code equations for a particular + constant is taken (in accordance with the policy applied by other + parts of the code generator framework). + + b) Code equations stemming from explicit declarations (e.g. code + attribute) gain priority over default code equations stemming + from definition, primrec, fun etc. + +* Global versions of theorems stemming from classes do not carry a +parameter prefix any longer. INCOMPATIBILITY. * Dropped locale element "includes". This is a major INCOMPATIBILITY. In existing theorem specifications replace the includes element by the -respective context elements of the included locale, omitting those that -are already present in the theorem specification. Multiple assume -elements of a locale should be replaced by a single one involving the -locale predicate. In the proof body, declarations (most notably -theorems) may be regained by interpreting the respective locales in the -proof context as required (command "interpret"). +respective context elements of the included locale, omitting those +that are already present in the theorem specification. Multiple +assume elements of a locale should be replaced by a single one +involving the locale predicate. In the proof body, declarations (most +notably theorems) may be regained by interpreting the respective +locales in the proof context as required (command "interpret"). + If using "includes" in replacement of a target solely because the parameter types in the theorem are not as general as in the target, -consider declaring a new locale with additional type constraints on the -parameters (context element "constrains"). +consider declaring a new locale with additional type constraints on +the parameters (context element "constrains"). * Dropped "locale (open)". INCOMPATIBILITY. @@ -145,9 +150,9 @@ * Interpretation commands no longer accept interpretation attributes. INCOMPATBILITY. -* Complete re-implementation of locales. INCOMPATIBILITY. -The most important changes are listed below. See documentation -(forthcoming) and tutorial (also forthcoming) for details. +* Complete re-implementation of locales. INCOMPATIBILITY. The most +important changes are listed below. See documentation (forthcoming) +and tutorial (also forthcoming) for details. - In locale expressions, instantiation replaces renaming. Parameters must be declared in a for clause. To aid compatibility with previous @@ -161,15 +166,15 @@ - More flexible mechanisms to qualify names generated by locale expressions. Qualifiers (prefixes) may be specified in locale -expressions. Available are normal qualifiers (syntax "name:") and strict -qualifiers (syntax "name!:"). The latter must occur in name references -and are useful to avoid accidental hiding of names, the former are -optional. Qualifiers derived from the parameter names of a locale are no -longer generated. - -- "sublocale l < e" replaces "interpretation l < e". The instantiation -clause in "interpretation" and "interpret" (square brackets) is no -longer available. Use locale expressions. +expressions. Available are normal qualifiers (syntax "name:") and +strict qualifiers (syntax "name!:"). The latter must occur in name +references and are useful to avoid accidental hiding of names, the +former are optional. Qualifiers derived from the parameter names of a +locale are no longer generated. + +- "sublocale l < e" replaces "interpretation l < e". The +instantiation clause in "interpretation" and "interpret" (square +brackets) is no longer available. Use locale expressions. - When converting proof scripts, be sure to replace qualifiers in "interpretation" and "interpret" by strict qualifiers. Qualifiers in @@ -183,8 +188,8 @@ * The 'axiomatization' command now only works within a global theory context. INCOMPATIBILITY. -* New find_theorems criterion "solves" matching theorems that -directly solve the current goal. Try "find_theorems solves". +* New find_theorems criterion "solves" matching theorems that directly +solve the current goal. Try "find_theorems solves". * Added an auto solve option, which can be enabled through the ProofGeneral Isabelle settings menu (disabled by default). @@ -193,14 +198,15 @@ stated. Any theorems that could solve the lemma directly are listed underneath the goal. -* New command find_consts searches for constants based on type and name -patterns, e.g. +* New command find_consts searches for constants based on type and +name patterns, e.g. find_consts "_ => bool" -By default, matching is against subtypes, but it may be restricted to the -whole type. Searching by name is possible. Multiple queries are conjunctive -and queries may be negated by prefixing them with a hyphen: +By default, matching is against subtypes, but it may be restricted to +the whole type. Searching by name is possible. Multiple queries are +conjunctive and queries may be negated by prefixing them with a +hyphen: find_consts strict: "_ => bool" name: "Int" -"int => int" @@ -240,7 +246,7 @@ src/HOL/Library/Order_Relation.thy ~> src/HOL/ src/HOL/Library/Parity.thy ~> src/HOL/ src/HOL/Library/Univ_Poly.thy ~> src/HOL/ - src/HOL/Real/ContNotDenum.thy ~> src/HOL/ + src/HOL/Real/ContNotDenum.thy ~> src/HOL/Library/ src/HOL/Real/Lubs.thy ~> src/HOL/ src/HOL/Real/PReal.thy ~> src/HOL/ src/HOL/Real/Rational.thy ~> src/HOL/ @@ -250,8 +256,8 @@ src/HOL/Real/Real.thy ~> src/HOL/ src/HOL/Complex/Complex_Main.thy ~> src/HOL/ src/HOL/Complex/Complex.thy ~> src/HOL/ - src/HOL/Complex/FrechetDeriv.thy ~> src/HOL/ - src/HOL/Complex/Fundamental_Theorem_Algebra.thy ~> src/HOL/ + src/HOL/Complex/FrechetDeriv.thy ~> src/HOL/Library/ + src/HOL/Complex/Fundamental_Theorem_Algebra.thy ~> src/HOL/Library/ src/HOL/Hyperreal/Deriv.thy ~> src/HOL/ src/HOL/Hyperreal/Fact.thy ~> src/HOL/ src/HOL/Hyperreal/Integration.thy ~> src/HOL/ @@ -312,7 +318,7 @@ process. New thread-based implementation also works on non-Unix platforms (Cygwin). Provers are no longer hardwired, but defined within the theory via plain ML wrapper functions. Basic Sledgehammer -commands are covered in the isar-ref manual +commands are covered in the isar-ref manual. * Wrapper scripts for remote SystemOnTPTP service allows to use sledgehammer without local ATP installation (Vampire etc.). See also @@ -342,6 +348,9 @@ etc. slightly changed. Some theorems named order_class.* now named preorder_class.*. +* HOL/Relation: +Renamed "refl" to "refl_on", "reflexive" to "refl, "diag" to "Id_on". + * HOL/Finite_Set: added a new fold combinator of type ('a => 'b => 'b) => 'b => 'a set => 'b Occasionally this is more convenient than the old fold combinator which is @@ -367,10 +376,60 @@ mult_div ~> div_mult_self2_is_id mult_mod ~> mod_mult_self2_is_0 +* HOL/IntDiv: removed many lemmas that are instances of class-based +generalizations (from Divides and Ring_and_Field). +INCOMPATIBILITY. Rename old lemmas as follows: + +dvd_diff -> nat_dvd_diff +dvd_zminus_iff -> dvd_minus_iff +mod_add1_eq -> mod_add_eq +mod_mult1_eq -> mod_mult_right_eq +mod_mult1_eq' -> mod_mult_left_eq +mod_mult_distrib_mod -> mod_mult_eq +nat_mod_add_left_eq -> mod_add_left_eq +nat_mod_add_right_eq -> mod_add_right_eq +nat_mod_div_trivial -> mod_div_trivial +nat_mod_mod_trivial -> mod_mod_trivial +zdiv_zadd_self1 -> div_add_self1 +zdiv_zadd_self2 -> div_add_self2 +zdiv_zmult_self1 -> div_mult_self2_is_id +zdiv_zmult_self2 -> div_mult_self1_is_id +zdvd_triv_left -> dvd_triv_left +zdvd_triv_right -> dvd_triv_right +zdvd_zmult_cancel_disj -> dvd_mult_cancel_left +zmod_eq0_zdvd_iff -> dvd_eq_mod_eq_0[symmetric] +zmod_zadd_left_eq -> mod_add_left_eq +zmod_zadd_right_eq -> mod_add_right_eq +zmod_zadd_self1 -> mod_add_self1 +zmod_zadd_self2 -> mod_add_self2 +zmod_zadd1_eq -> mod_add_eq +zmod_zdiff1_eq -> mod_diff_eq +zmod_zdvd_zmod -> mod_mod_cancel +zmod_zmod_cancel -> mod_mod_cancel +zmod_zmult_self1 -> mod_mult_self2_is_0 +zmod_zmult_self2 -> mod_mult_self1_is_0 +zmod_1 -> mod_by_1 +zdiv_1 -> div_by_1 +zdvd_abs1 -> abs_dvd_iff +zdvd_abs2 -> dvd_abs_iff +zdvd_refl -> dvd_refl +zdvd_trans -> dvd_trans +zdvd_zadd -> dvd_add +zdvd_zdiff -> dvd_diff +zdvd_zminus_iff -> dvd_minus_iff +zdvd_zminus2_iff -> minus_dvd_iff +zdvd_zmultD -> dvd_mult_right +zdvd_zmultD2 -> dvd_mult_left +zdvd_zmult_mono -> mult_dvd_mono +zdvd_0_right -> dvd_0_right +zdvd_0_left -> dvd_0_left_iff +zdvd_1_left -> one_dvd +zminus_dvd_iff -> minus_dvd_iff + * HOL/Library/GCD: Curried operations gcd, lcm (for nat) and zgcd, zlcm (for int); carried together from various gcd/lcm developements in the HOL Distribution. zgcd and zlcm replace former igcd and ilcm; -corresponding theorems renamed accordingly. INCOMPATIBILY. To +corresponding theorems renamed accordingly. INCOMPATIBILITY. To recover tupled syntax, use syntax declarations like: hide (open) const gcd @@ -384,7 +443,7 @@ * HOL/Real/Rational: 'Fract k 0' now equals '0'. INCOMPATIBILITY. * The real numbers offer decimal input syntax: 12.34 is translated into - 1234/10^4. This translation is not reversed upon output. + 1234/10^2. This translation is not reversed upon output. * New ML antiquotation @{code}: takes constant as argument, generates corresponding code in background and inserts name of the corresponding @@ -441,6 +500,9 @@ Suc_Suc_eq ~> nat.inject Suc_not_Zero Zero_not_Suc ~> nat.distinct +* The option datatype has been moved to a new theory HOL/Option.thy. +Renamed option_map to Option.map. + * Library/Nat_Infinity: added addition, numeral syntax and more instantiations for algebraic structures. Removed some duplicate theorems. Changes in simp rules. INCOMPATIBILITY. @@ -452,9 +514,8 @@ *** HOL-Algebra *** * New locales for orders and lattices where the equivalence relation - is not restricted to equality. INCOMPATIBILITY: all order and - lattice locales use a record structure with field eq for the - equivalence. +is not restricted to equality. INCOMPATIBILITY: all order and lattice +locales use a record structure with field eq for the equivalence. * New theory of factorial domains. @@ -485,6 +546,23 @@ * Proof of Zorn's Lemma for partial orders. +*** HOLCF *** + +* Reimplemented the simplification procedure for proving continuity +subgoals. The new simproc is extensible; users can declare additional +continuity introduction rules with the attribute [cont2cont]. + +* The continuity simproc now uses a different introduction rule for +solving continuity subgoals on terms with lambda abstractions. In +some rare cases the new simproc may fail to solve subgoals that the +old one could solve, and "simp add: cont2cont_LAM" may be necessary. +Potential INCOMPATIBILITY. + +* The syntax of the fixrec package has changed. The specification +syntax now conforms in style to definition, primrec, function, etc. +See HOLCF/ex/Fixrec_ex.thy for examples. INCOMPATIBILITY. + + *** ML *** * High-level support for concurrent ML programming, see diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 README_REPOSITORY --- a/README_REPOSITORY Wed Mar 04 11:05:02 2009 +0100 +++ b/README_REPOSITORY Wed Mar 04 11:05:29 2009 +0100 @@ -32,9 +32,9 @@ Initial configuration --------------------- -Always use Mercurial version 1.0 or later, such as 1.0.1 or 1.0.2. +Always use Mercurial versions from the 1.0 or 1.1 branch, or later. The old 0.9.x versions do not work in a multi-user environment with -shared file spaces. +shared file spaces! The official Isabelle repository can be cloned like this: @@ -62,7 +62,8 @@ In principle, user names can be chosen freely, but for longterm committers of the Isabelle repository the obvious choice is to keep -with the old CVS naming scheme. +with the old CVS naming scheme. Others should use their regular "full +name"; including an email address is optional. There are other useful configuration to go into $HOME/.hgrc, @@ -135,6 +136,29 @@ hg clone ssh://wenzelm@atbroy100//home/isabelle-repository/repos/isabelle +Simplified merges +----------------- + +The main idea of Mercurial is to let individual users produce +independent branches of development first, but merge with others +frequently. The basic hg merge operation is more general than +required for the mode of operation with a shared pull/push area. The +hg fetch extension accommodates this case nicely, automating trivial +merges and requiring manual intervention for actual conflicts only. + +The fetch extension can be configured via the user's ~/.hgrc like +this: + + [extensions] + hgext.fetch = + + [defaults] + fetch = -m "merged" + +Note that the potential for merge conflicts can be greatly reduced by +doing "hg fetch" before any starting local changes! + + Content discipline ------------------ @@ -172,7 +196,9 @@ Mercurial provides nice web presentation of incoming changes with a digest of log entries; this also includes RSS/Atom news feeds. Users should be aware that others will actually read what is - written into log messages. + written into log messages. There are also add-on browsers, + notably hgtk that is part of the TortoiseHg distribution and works + for generic Python/GTk platforms. The usual changelog presentation style for the Isabelle repository admits log entries that consist of several lines, but without the @@ -194,6 +220,3 @@ Needless to say, the results from the build process must not be added to the repository! - - -Makarius 30-Nov-2008 diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/Group/Group.thy --- a/doc-src/AxClass/Group/Group.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,322 +0,0 @@ - -header {* Basic group theory *} - -theory Group imports Main begin - -text {* - \medskip\noindent The meta-level type system of Isabelle supports - \emph{intersections} and \emph{inclusions} of type classes. These - directly correspond to intersections and inclusions of type - predicates in a purely set theoretic sense. This is sufficient as a - means to describe simple hierarchies of structures. As an - illustration, we use the well-known example of semigroups, monoids, - general groups and Abelian groups. -*} - -subsection {* Monoids and Groups *} - -text {* - First we declare some polymorphic constants required later for the - signature parts of our structures. -*} - -consts - times :: "'a \ 'a \ 'a" (infixl "\" 70) - invers :: "'a \ 'a" ("(_\)" [1000] 999) - one :: 'a ("\") - -text {* - \noindent Next we define class @{text monoid} of monoids with - operations @{text \} and @{text \}. Note that multiple class - axioms are allowed for user convenience --- they simply represent - the conjunction of their respective universal closures. -*} - -axclass monoid \ type - assoc: "(x \ y) \ z = x \ (y \ z)" - left_unit: "\ \ x = x" - right_unit: "x \ \ = x" - -text {* - \noindent So class @{text monoid} contains exactly those types - @{text \} where @{text "\ \ \ \ \ \ \"} and @{text "\ \ \"} - are specified appropriately, such that @{text \} is associative and - @{text \} is a left and right unit element for the @{text \} - operation. -*} - -text {* - \medskip Independently of @{text monoid}, we now define a linear - hierarchy of semigroups, general groups and Abelian groups. Note - that the names of class axioms are automatically qualified with each - class name, so we may re-use common names such as @{text assoc}. -*} - -axclass semigroup \ type - assoc: "(x \ y) \ z = x \ (y \ z)" - -axclass group \ semigroup - left_unit: "\ \ x = x" - left_inverse: "x\ \ x = \" - -axclass agroup \ group - commute: "x \ y = y \ x" - -text {* - \noindent Class @{text group} inherits associativity of @{text \} - from @{text semigroup} and adds two further group axioms. Similarly, - @{text agroup} is defined as the subset of @{text group} such that - for all of its elements @{text \}, the operation @{text "\ \ \ \ \ \ - \"} is even commutative. -*} - - -subsection {* Abstract reasoning *} - -text {* - In a sense, axiomatic type classes may be viewed as \emph{abstract - theories}. Above class definitions gives rise to abstract axioms - @{text assoc}, @{text left_unit}, @{text left_inverse}, @{text - commute}, where any of these contain a type variable @{text "'a \ - c"} that is restricted to types of the corresponding class @{text - c}. \emph{Sort constraints} like this express a logical - precondition for the whole formula. For example, @{text assoc} - states that for all @{text \}, provided that @{text "\ \ - semigroup"}, the operation @{text "\ \ \ \ \ \ \"} is associative. - - \medskip From a technical point of view, abstract axioms are just - ordinary Isabelle theorems, which may be used in proofs without - special treatment. Such ``abstract proofs'' usually yield new - ``abstract theorems''. For example, we may now derive the following - well-known laws of general groups. -*} - -theorem group_right_inverse: "x \ x\ = (\\'a\group)" -proof - - have "x \ x\ = \ \ (x \ x\)" - by (simp only: group_class.left_unit) - also have "... = \ \ x \ x\" - by (simp only: semigroup_class.assoc) - also have "... = (x\)\ \ x\ \ x \ x\" - by (simp only: group_class.left_inverse) - also have "... = (x\)\ \ (x\ \ x) \ x\" - by (simp only: semigroup_class.assoc) - also have "... = (x\)\ \ \ \ x\" - by (simp only: group_class.left_inverse) - also have "... = (x\)\ \ (\ \ x\)" - by (simp only: semigroup_class.assoc) - also have "... = (x\)\ \ x\" - by (simp only: group_class.left_unit) - also have "... = \" - by (simp only: group_class.left_inverse) - finally show ?thesis . -qed - -text {* - \noindent With @{text group_right_inverse} already available, @{text - group_right_unit}\label{thm:group-right-unit} is now established - much easier. -*} - -theorem group_right_unit: "x \ \ = (x\'a\group)" -proof - - have "x \ \ = x \ (x\ \ x)" - by (simp only: group_class.left_inverse) - also have "... = x \ x\ \ x" - by (simp only: semigroup_class.assoc) - also have "... = \ \ x" - by (simp only: group_right_inverse) - also have "... = x" - by (simp only: group_class.left_unit) - finally show ?thesis . -qed - -text {* - \medskip Abstract theorems may be instantiated to only those types - @{text \} where the appropriate class membership @{text "\ \ c"} is - known at Isabelle's type signature level. Since we have @{text - "agroup \ group \ semigroup"} by definition, all theorems of @{text - semigroup} and @{text group} are automatically inherited by @{text - group} and @{text agroup}. -*} - - -subsection {* Abstract instantiation *} - -text {* - From the definition, the @{text monoid} and @{text group} classes - have been independent. Note that for monoids, @{text right_unit} - had to be included as an axiom, but for groups both @{text - right_unit} and @{text right_inverse} are derivable from the other - axioms. With @{text group_right_unit} derived as a theorem of group - theory (see page~\pageref{thm:group-right-unit}), we may now - instantiate @{text "monoid \ semigroup"} and @{text "group \ - monoid"} properly as follows (cf.\ \figref{fig:monoid-group}). - - \begin{figure}[htbp] - \begin{center} - \small - \unitlength 0.6mm - \begin{picture}(65,90)(0,-10) - \put(15,10){\line(0,1){10}} \put(15,30){\line(0,1){10}} - \put(15,50){\line(1,1){10}} \put(35,60){\line(1,-1){10}} - \put(15,5){\makebox(0,0){@{text agroup}}} - \put(15,25){\makebox(0,0){@{text group}}} - \put(15,45){\makebox(0,0){@{text semigroup}}} - \put(30,65){\makebox(0,0){@{text type}}} \put(50,45){\makebox(0,0){@{text monoid}}} - \end{picture} - \hspace{4em} - \begin{picture}(30,90)(0,0) - \put(15,10){\line(0,1){10}} \put(15,30){\line(0,1){10}} - \put(15,50){\line(0,1){10}} \put(15,70){\line(0,1){10}} - \put(15,5){\makebox(0,0){@{text agroup}}} - \put(15,25){\makebox(0,0){@{text group}}} - \put(15,45){\makebox(0,0){@{text monoid}}} - \put(15,65){\makebox(0,0){@{text semigroup}}} - \put(15,85){\makebox(0,0){@{text type}}} - \end{picture} - \caption{Monoids and groups: according to definition, and by proof} - \label{fig:monoid-group} - \end{center} - \end{figure} -*} - -instance monoid \ semigroup -proof - fix x y z :: "'a\monoid" - show "x \ y \ z = x \ (y \ z)" - by (rule monoid_class.assoc) -qed - -instance group \ monoid -proof - fix x y z :: "'a\group" - show "x \ y \ z = x \ (y \ z)" - by (rule semigroup_class.assoc) - show "\ \ x = x" - by (rule group_class.left_unit) - show "x \ \ = x" - by (rule group_right_unit) -qed - -text {* - \medskip The \isakeyword{instance} command sets up an appropriate - goal that represents the class inclusion (or type arity, see - \secref{sec:inst-arity}) to be proven (see also - \cite{isabelle-isar-ref}). The initial proof step causes - back-chaining of class membership statements wrt.\ the hierarchy of - any classes defined in the current theory; the effect is to reduce - to the initial statement to a number of goals that directly - correspond to any class axioms encountered on the path upwards - through the class hierarchy. -*} - - -subsection {* Concrete instantiation \label{sec:inst-arity} *} - -text {* - So far we have covered the case of the form - \isakeyword{instance}~@{text "c\<^sub>1 \ c\<^sub>2"}, namely - \emph{abstract instantiation} --- $c@1$ is more special than @{text - "c\<^sub>1"} and thus an instance of @{text "c\<^sub>2"}. Even more - interesting for practical applications are \emph{concrete - instantiations} of axiomatic type classes. That is, certain simple - schemes @{text "(\\<^sub>1, \, \\<^sub>n) t \ c"} of class - membership may be established at the logical level and then - transferred to Isabelle's type signature level. - - \medskip As a typical example, we show that type @{typ bool} with - exclusive-or as @{text \} operation, identity as @{text \}, and - @{term False} as @{text \} forms an Abelian group. -*} - -defs (overloaded) - times_bool_def: "x \ y \ x \ (y\bool)" - inverse_bool_def: "x\ \ x\bool" - unit_bool_def: "\ \ False" - -text {* - \medskip It is important to note that above \isakeyword{defs} are - just overloaded meta-level constant definitions, where type classes - are not yet involved at all. This form of constant definition with - overloading (and optional recursion over the syntactic structure of - simple types) are admissible as definitional extensions of plain HOL - \cite{Wenzel:1997:TPHOL}. The Haskell-style type system is not - required for overloading. Nevertheless, overloaded definitions are - best applied in the context of type classes. - - \medskip Since we have chosen above \isakeyword{defs} of the generic - group operations on type @{typ bool} appropriately, the class - membership @{text "bool \ agroup"} may be now derived as follows. -*} - -instance bool :: agroup -proof (intro_classes, - unfold times_bool_def inverse_bool_def unit_bool_def) - fix x y z - show "((x \ y) \ z) = (x \ (y \ z))" by blast - show "(False \ x) = x" by blast - show "(x \ x) = False" by blast - show "(x \ y) = (y \ x)" by blast -qed - -text {* - The result of an \isakeyword{instance} statement is both expressed - as a theorem of Isabelle's meta-logic, and as a type arity of the - type signature. The latter enables type-inference system to take - care of this new instance automatically. - - \medskip We could now also instantiate our group theory classes to - many other concrete types. For example, @{text "int \ agroup"} - (e.g.\ by defining @{text \} as addition, @{text \} as negation - and @{text \} as zero) or @{text "list \ (type) semigroup"} - (e.g.\ if @{text \} is defined as list append). Thus, the - characteristic constants @{text \}, @{text \}, @{text \} - really become overloaded, i.e.\ have different meanings on different - types. -*} - - -subsection {* Lifting and Functors *} - -text {* - As already mentioned above, overloading in the simply-typed HOL - systems may include recursion over the syntactic structure of types. - That is, definitional equations @{text "c\<^sup>\ \ t"} may also - contain constants of name @{text c} on the right-hand side --- if - these have types that are structurally simpler than @{text \}. - - This feature enables us to \emph{lift operations}, say to Cartesian - products, direct sums or function spaces. Subsequently we lift - @{text \} component-wise to binary products @{typ "'a \ 'b"}. -*} - -defs (overloaded) - times_prod_def: "p \ q \ (fst p \ fst q, snd p \ snd q)" - -text {* - It is very easy to see that associativity of @{text \} on @{typ 'a} - and @{text \} on @{typ 'b} transfers to @{text \} on @{typ "'a \ - 'b"}. Hence the binary type constructor @{text \} maps semigroups - to semigroups. This may be established formally as follows. -*} - -instance * :: (semigroup, semigroup) semigroup -proof (intro_classes, unfold times_prod_def) - fix p q r :: "'a\semigroup \ 'b\semigroup" - show - "(fst (fst p \ fst q, snd p \ snd q) \ fst r, - snd (fst p \ fst q, snd p \ snd q) \ snd r) = - (fst p \ fst (fst q \ fst r, snd q \ snd r), - snd p \ snd (fst q \ fst r, snd q \ snd r))" - by (simp add: semigroup_class.assoc) -qed - -text {* - Thus, if we view class instances as ``structures'', then overloaded - constant definitions with recursion over types indirectly provide - some kind of ``functors'' --- i.e.\ mappings between abstract - theories. -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/Group/Product.thy --- a/doc-src/AxClass/Group/Product.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,85 +0,0 @@ - -header {* Syntactic classes *} - -theory Product imports Main begin - -text {* - \medskip\noindent There is still a feature of Isabelle's type system - left that we have not yet discussed. When declaring polymorphic - constants @{text "c \ \"}, the type variables occurring in @{text \} - may be constrained by type classes (or even general sorts) in an - arbitrary way. Note that by default, in Isabelle/HOL the - declaration @{text "\ \ 'a \ 'a \ 'a"} is actually an abbreviation - for @{text "\ \ 'a\type \ 'a \ 'a"} Since class @{text type} is the - universal class of HOL, this is not really a constraint at all. - - The @{text product} class below provides a less degenerate example of - syntactic type classes. -*} - -axclass - product \ type -consts - product :: "'a\product \ 'a \ 'a" (infixl "\" 70) - -text {* - Here class @{text product} is defined as subclass of @{text type} - without any additional axioms. This effects in logical equivalence - of @{text product} and @{text type}, as is reflected by the trivial - introduction rule generated for this definition. - - \medskip So what is the difference of declaring @{text "\ \ - 'a\product \ 'a \ 'a"} vs.\ declaring @{text "\ \ 'a\type \ 'a \ - 'a"} anyway? In this particular case where @{text "product \ - type"}, it should be obvious that both declarations are the same - from the logic's point of view. It even makes the most sense to - remove sort constraints from constant declarations, as far as the - purely logical meaning is concerned \cite{Wenzel:1997:TPHOL}. - - On the other hand there are syntactic differences, of course. - Constants @{text \} on some type @{text \} are rejected by the - type-checker, unless the arity @{text "\ \ product"} is part of the - type signature. In our example, this arity may be always added when - required by means of an \isakeyword{instance} with the default proof - (double-dot). - - \medskip Thus, we may observe the following discipline of using - syntactic classes. Overloaded polymorphic constants have their type - arguments restricted to an associated (logically trivial) class - @{text c}. Only immediately before \emph{specifying} these - constants on a certain type @{text \} do we instantiate @{text "\ \ - c"}. - - This is done for class @{text product} and type @{typ bool} as - follows. -*} - -instance bool :: product .. -defs (overloaded) - product_bool_def: "x \ y \ x \ y" - -text {* - The definition @{text prod_bool_def} becomes syntactically - well-formed only after the arity @{text "bool \ product"} is made - known to the type checker. - - \medskip It is very important to see that above \isakeyword{defs} are - not directly connected with \isakeyword{instance} at all! We were - just following our convention to specify @{text \} on @{typ bool} - after having instantiated @{text "bool \ product"}. Isabelle does - not require these definitions, which is in contrast to programming - languages like Haskell \cite{haskell-report}. - - \medskip While Isabelle type classes and those of Haskell are almost - the same as far as type-checking and type inference are concerned, - there are important semantic differences. Haskell classes require - their instances to \emph{provide operations} of certain \emph{names}. - Therefore, its \texttt{instance} has a \texttt{where} part that tells - the system what these ``member functions'' should be. - - This style of \texttt{instance} would not make much sense in - Isabelle's meta-logic, because there is no internal notion of - ``providing operations'' or even ``names of functions''. -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/Group/ROOT.ML --- a/doc-src/AxClass/Group/ROOT.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,4 +0,0 @@ - -use_thy "Semigroups"; -use_thy "Group"; -use_thy "Product"; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/Group/Semigroups.thy --- a/doc-src/AxClass/Group/Semigroups.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,54 +0,0 @@ - -header {* Semigroups *} - -theory Semigroups imports Main begin - -text {* - \medskip\noindent An axiomatic type class is simply a class of types - that all meet certain properties, which are also called \emph{class - axioms}. Thus, type classes may be also understood as type - predicates --- i.e.\ abstractions over a single type argument @{typ - 'a}. Class axioms typically contain polymorphic constants that - depend on this type @{typ 'a}. These \emph{characteristic - constants} behave like operations associated with the ``carrier'' - type @{typ 'a}. - - We illustrate these basic concepts by the following formulation of - semigroups. -*} - -consts - times :: "'a \ 'a \ 'a" (infixl "\" 70) -axclass semigroup \ type - assoc: "(x \ y) \ z = x \ (y \ z)" - -text {* - \noindent Above we have first declared a polymorphic constant @{text - "\ \ 'a \ 'a \ 'a"} and then defined the class @{text semigroup} of - all types @{text \} such that @{text "\ \ \ \ \ \ \"} is indeed an - associative operator. The @{text assoc} axiom contains exactly one - type variable, which is invisible in the above presentation, though. - Also note that free term variables (like @{term x}, @{term y}, - @{term z}) are allowed for user convenience --- conceptually all of - these are bound by outermost universal quantifiers. - - \medskip In general, type classes may be used to describe - \emph{structures} with exactly one carrier @{typ 'a} and a fixed - \emph{signature}. Different signatures require different classes. - Below, class @{text plus_semigroup} represents semigroups @{text - "(\, \\<^sup>\)"}, while the original @{text semigroup} would - correspond to semigroups of the form @{text "(\, \\<^sup>\)"}. -*} - -consts - plus :: "'a \ 'a \ 'a" (infixl "\" 70) -axclass plus_semigroup \ type - assoc: "(x \ y) \ z = x \ (y \ z)" - -text {* - \noindent Even if classes @{text plus_semigroup} and @{text - semigroup} both represent semigroups in a sense, they are certainly - not quite the same. -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/Group/document/Group.tex --- a/doc-src/AxClass/Group/document/Group.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,512 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Group}% -% -\isamarkupheader{Basic group theory% -} -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Group\ \isakeyword{imports}\ Main\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\begin{isamarkuptext}% -\medskip\noindent The meta-level type system of Isabelle supports - \emph{intersections} and \emph{inclusions} of type classes. These - directly correspond to intersections and inclusions of type - predicates in a purely set theoretic sense. This is sufficient as a - means to describe simple hierarchies of structures. As an - illustration, we use the well-known example of semigroups, monoids, - general groups and Abelian groups.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Monoids and Groups% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -First we declare some polymorphic constants required later for the - signature parts of our structures.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{consts}\isamarkupfalse% -\isanewline -\ \ times\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymodot}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ invers\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}{\isachardoublequoteopen}{\isacharparenleft}{\isacharunderscore}{\isasyminv}{\isacharparenright}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline -\ \ one\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ \ \ \ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}{\isacharparenright}% -\begin{isamarkuptext}% -\noindent Next we define class \isa{monoid} of monoids with - operations \isa{{\isasymodot}} and \isa{{\isasymone}}. Note that multiple class - axioms are allowed for user convenience --- they simply represent - the conjunction of their respective universal closures.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{axclass}\isamarkupfalse% -\ monoid\ {\isasymsubseteq}\ type\isanewline -\ \ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymodot}\ y{\isacharparenright}\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ left{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymodot}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -\ \ right{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymodot}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent So class \isa{monoid} contains exactly those types - \isa{{\isasymtau}} where \isa{{\isasymodot}\ {\isasymColon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}} and \isa{{\isasymone}\ {\isasymColon}\ {\isasymtau}} - are specified appropriately, such that \isa{{\isasymodot}} is associative and - \isa{{\isasymone}} is a left and right unit element for the \isa{{\isasymodot}} - operation.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\medskip Independently of \isa{monoid}, we now define a linear - hierarchy of semigroups, general groups and Abelian groups. Note - that the names of class axioms are automatically qualified with each - class name, so we may re-use common names such as \isa{assoc}.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{axclass}\isamarkupfalse% -\ semigroup\ {\isasymsubseteq}\ type\isanewline -\ \ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymodot}\ y{\isacharparenright}\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{axclass}\isamarkupfalse% -\ group\ {\isasymsubseteq}\ semigroup\isanewline -\ \ left{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymodot}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -\ \ left{\isacharunderscore}inverse{\isacharcolon}\ {\isachardoublequoteopen}x{\isasyminv}\ {\isasymodot}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{axclass}\isamarkupfalse% -\ agroup\ {\isasymsubseteq}\ group\isanewline -\ \ commute{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymodot}\ y\ {\isacharequal}\ y\ {\isasymodot}\ x{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent Class \isa{group} inherits associativity of \isa{{\isasymodot}} - from \isa{semigroup} and adds two further group axioms. Similarly, - \isa{agroup} is defined as the subset of \isa{group} such that - for all of its elements \isa{{\isasymtau}}, the operation \isa{{\isasymodot}\ {\isasymColon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}} is even commutative.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Abstract reasoning% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In a sense, axiomatic type classes may be viewed as \emph{abstract - theories}. Above class definitions gives rise to abstract axioms - \isa{assoc}, \isa{left{\isacharunderscore}unit}, \isa{left{\isacharunderscore}inverse}, \isa{commute}, where any of these contain a type variable \isa{{\isacharprime}a\ {\isasymColon}\ c} that is restricted to types of the corresponding class \isa{c}. \emph{Sort constraints} like this express a logical - precondition for the whole formula. For example, \isa{assoc} - states that for all \isa{{\isasymtau}}, provided that \isa{{\isasymtau}\ {\isasymColon}\ semigroup}, the operation \isa{{\isasymodot}\ {\isasymColon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}} is associative. - - \medskip From a technical point of view, abstract axioms are just - ordinary Isabelle theorems, which may be used in proofs without - special treatment. Such ``abstract proofs'' usually yield new - ``abstract theorems''. For example, we may now derive the following - well-known laws of general groups.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{theorem}\isamarkupfalse% -\ group{\isacharunderscore}right{\isacharunderscore}inverse{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymodot}\ x{\isasyminv}\ {\isacharequal}\ {\isacharparenleft}{\isasymone}{\isasymColon}{\isacharprime}a{\isasymColon}group{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isacharminus}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymodot}\ x{\isasyminv}\ {\isacharequal}\ {\isasymone}\ {\isasymodot}\ {\isacharparenleft}x\ {\isasymodot}\ x{\isasyminv}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isacharunderscore}class{\isachardot}left{\isacharunderscore}unit{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isasymone}\ {\isasymodot}\ x\ {\isasymodot}\ x{\isasyminv}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ semigroup{\isacharunderscore}class{\isachardot}assoc{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ x{\isasyminv}\ {\isasymodot}\ x\ {\isasymodot}\ x{\isasyminv}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isacharunderscore}class{\isachardot}left{\isacharunderscore}inverse{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ {\isacharparenleft}x{\isasyminv}\ {\isasymodot}\ x{\isacharparenright}\ {\isasymodot}\ x{\isasyminv}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ semigroup{\isacharunderscore}class{\isachardot}assoc{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ {\isasymone}\ {\isasymodot}\ x{\isasyminv}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isacharunderscore}class{\isachardot}left{\isacharunderscore}inverse{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ {\isacharparenleft}{\isasymone}\ {\isasymodot}\ x{\isasyminv}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ semigroup{\isacharunderscore}class{\isachardot}assoc{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ x{\isasyminv}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isacharunderscore}class{\isachardot}left{\isacharunderscore}unit{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isacharunderscore}class{\isachardot}left{\isacharunderscore}inverse{\isacharparenright}\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent With \isa{group{\isacharunderscore}right{\isacharunderscore}inverse} already available, \isa{group{\isacharunderscore}right{\isacharunderscore}unit}\label{thm:group-right-unit} is now established - much easier.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{theorem}\isamarkupfalse% -\ group{\isacharunderscore}right{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymodot}\ {\isasymone}\ {\isacharequal}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isasymColon}group{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isacharminus}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymodot}\ {\isasymone}\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}x{\isasyminv}\ {\isasymodot}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isacharunderscore}class{\isachardot}left{\isacharunderscore}inverse{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ x\ {\isasymodot}\ x{\isasyminv}\ {\isasymodot}\ x{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ semigroup{\isacharunderscore}class{\isachardot}assoc{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isasymone}\ {\isasymodot}\ x{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isacharunderscore}right{\isacharunderscore}inverse{\isacharparenright}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isacharunderscore}class{\isachardot}left{\isacharunderscore}unit{\isacharparenright}\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isacharquery}thesis\ \isacommand{{\isachardot}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\medskip Abstract theorems may be instantiated to only those types - \isa{{\isasymtau}} where the appropriate class membership \isa{{\isasymtau}\ {\isasymColon}\ c} is - known at Isabelle's type signature level. Since we have \isa{agroup\ {\isasymsubseteq}\ group\ {\isasymsubseteq}\ semigroup} by definition, all theorems of \isa{semigroup} and \isa{group} are automatically inherited by \isa{group} and \isa{agroup}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Abstract instantiation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -From the definition, the \isa{monoid} and \isa{group} classes - have been independent. Note that for monoids, \isa{right{\isacharunderscore}unit} - had to be included as an axiom, but for groups both \isa{right{\isacharunderscore}unit} and \isa{right{\isacharunderscore}inverse} are derivable from the other - axioms. With \isa{group{\isacharunderscore}right{\isacharunderscore}unit} derived as a theorem of group - theory (see page~\pageref{thm:group-right-unit}), we may now - instantiate \isa{monoid\ {\isasymsubseteq}\ semigroup} and \isa{group\ {\isasymsubseteq}\ monoid} properly as follows (cf.\ \figref{fig:monoid-group}). - - \begin{figure}[htbp] - \begin{center} - \small - \unitlength 0.6mm - \begin{picture}(65,90)(0,-10) - \put(15,10){\line(0,1){10}} \put(15,30){\line(0,1){10}} - \put(15,50){\line(1,1){10}} \put(35,60){\line(1,-1){10}} - \put(15,5){\makebox(0,0){\isa{agroup}}} - \put(15,25){\makebox(0,0){\isa{group}}} - \put(15,45){\makebox(0,0){\isa{semigroup}}} - \put(30,65){\makebox(0,0){\isa{type}}} \put(50,45){\makebox(0,0){\isa{monoid}}} - \end{picture} - \hspace{4em} - \begin{picture}(30,90)(0,0) - \put(15,10){\line(0,1){10}} \put(15,30){\line(0,1){10}} - \put(15,50){\line(0,1){10}} \put(15,70){\line(0,1){10}} - \put(15,5){\makebox(0,0){\isa{agroup}}} - \put(15,25){\makebox(0,0){\isa{group}}} - \put(15,45){\makebox(0,0){\isa{monoid}}} - \put(15,65){\makebox(0,0){\isa{semigroup}}} - \put(15,85){\makebox(0,0){\isa{type}}} - \end{picture} - \caption{Monoids and groups: according to definition, and by proof} - \label{fig:monoid-group} - \end{center} - \end{figure}% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{instance}\isamarkupfalse% -\ monoid\ {\isasymsubseteq}\ semigroup\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ x\ y\ z\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}monoid{\isachardoublequoteclose}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymodot}\ y\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ monoid{\isacharunderscore}class{\isachardot}assoc{\isacharparenright}\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{instance}\isamarkupfalse% -\ group\ {\isasymsubseteq}\ monoid\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ x\ y\ z\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}group{\isachardoublequoteclose}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymodot}\ y\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ semigroup{\isacharunderscore}class{\isachardot}assoc{\isacharparenright}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymone}\ {\isasymodot}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ group{\isacharunderscore}class{\isachardot}left{\isacharunderscore}unit{\isacharparenright}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymodot}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ group{\isacharunderscore}right{\isacharunderscore}unit{\isacharparenright}\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\medskip The \isakeyword{instance} command sets up an appropriate - goal that represents the class inclusion (or type arity, see - \secref{sec:inst-arity}) to be proven (see also - \cite{isabelle-isar-ref}). The initial proof step causes - back-chaining of class membership statements wrt.\ the hierarchy of - any classes defined in the current theory; the effect is to reduce - to the initial statement to a number of goals that directly - correspond to any class axioms encountered on the path upwards - through the class hierarchy.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Concrete instantiation \label{sec:inst-arity}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -So far we have covered the case of the form - \isakeyword{instance}~\isa{c\isactrlsub {\isadigit{1}}\ {\isasymsubseteq}\ c\isactrlsub {\isadigit{2}}}, namely - \emph{abstract instantiation} --- $c@1$ is more special than \isa{c\isactrlsub {\isadigit{1}}} and thus an instance of \isa{c\isactrlsub {\isadigit{2}}}. Even more - interesting for practical applications are \emph{concrete - instantiations} of axiomatic type classes. That is, certain simple - schemes \isa{{\isacharparenleft}{\isasymalpha}\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlsub n{\isacharparenright}\ t\ {\isasymColon}\ c} of class - membership may be established at the logical level and then - transferred to Isabelle's type signature level. - - \medskip As a typical example, we show that type \isa{bool} with - exclusive-or as \isa{{\isasymodot}} operation, identity as \isa{{\isasyminv}}, and - \isa{False} as \isa{{\isasymone}} forms an Abelian group.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{defs}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline -\ \ times{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymodot}\ y\ {\isasymequiv}\ x\ {\isasymnoteq}\ {\isacharparenleft}y{\isasymColon}bool{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ inverse{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}x{\isasyminv}\ {\isasymequiv}\ x{\isasymColon}bool{\isachardoublequoteclose}\isanewline -\ \ unit{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymequiv}\ False{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\medskip It is important to note that above \isakeyword{defs} are - just overloaded meta-level constant definitions, where type classes - are not yet involved at all. This form of constant definition with - overloading (and optional recursion over the syntactic structure of - simple types) are admissible as definitional extensions of plain HOL - \cite{Wenzel:1997:TPHOL}. The Haskell-style type system is not - required for overloading. Nevertheless, overloaded definitions are - best applied in the context of type classes. - - \medskip Since we have chosen above \isakeyword{defs} of the generic - group operations on type \isa{bool} appropriately, the class - membership \isa{bool\ {\isasymColon}\ agroup} may be now derived as follows.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{instance}\isamarkupfalse% -\ bool\ {\isacharcolon}{\isacharcolon}\ agroup\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isacharparenleft}intro{\isacharunderscore}classes{\isacharcomma}\isanewline -\ \ \ \ unfold\ times{\isacharunderscore}bool{\isacharunderscore}def\ inverse{\isacharunderscore}bool{\isacharunderscore}def\ unit{\isacharunderscore}bool{\isacharunderscore}def{\isacharparenright}\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ x\ y\ z\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharparenleft}x\ {\isasymnoteq}\ y{\isacharparenright}\ {\isasymnoteq}\ z{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymnoteq}\ {\isacharparenleft}y\ {\isasymnoteq}\ z{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ blast\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}False\ {\isasymnoteq}\ x{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ blast\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymnoteq}\ x{\isacharparenright}\ {\isacharequal}\ False{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ blast\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymnoteq}\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y\ {\isasymnoteq}\ x{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ blast\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -The result of an \isakeyword{instance} statement is both expressed - as a theorem of Isabelle's meta-logic, and as a type arity of the - type signature. The latter enables type-inference system to take - care of this new instance automatically. - - \medskip We could now also instantiate our group theory classes to - many other concrete types. For example, \isa{int\ {\isasymColon}\ agroup} - (e.g.\ by defining \isa{{\isasymodot}} as addition, \isa{{\isasyminv}} as negation - and \isa{{\isasymone}} as zero) or \isa{list\ {\isasymColon}\ {\isacharparenleft}type{\isacharparenright}\ semigroup} - (e.g.\ if \isa{{\isasymodot}} is defined as list append). Thus, the - characteristic constants \isa{{\isasymodot}}, \isa{{\isasyminv}}, \isa{{\isasymone}} - really become overloaded, i.e.\ have different meanings on different - types.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Lifting and Functors% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -As already mentioned above, overloading in the simply-typed HOL - systems may include recursion over the syntactic structure of types. - That is, definitional equations \isa{c\isactrlsup {\isasymtau}\ {\isasymequiv}\ t} may also - contain constants of name \isa{c} on the right-hand side --- if - these have types that are structurally simpler than \isa{{\isasymtau}}. - - This feature enables us to \emph{lift operations}, say to Cartesian - products, direct sums or function spaces. Subsequently we lift - \isa{{\isasymodot}} component-wise to binary products \isa{{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b}.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{defs}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline -\ \ times{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}p\ {\isasymodot}\ q\ {\isasymequiv}\ {\isacharparenleft}fst\ p\ {\isasymodot}\ fst\ q{\isacharcomma}\ snd\ p\ {\isasymodot}\ snd\ q{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -It is very easy to see that associativity of \isa{{\isasymodot}} on \isa{{\isacharprime}a} - and \isa{{\isasymodot}} on \isa{{\isacharprime}b} transfers to \isa{{\isasymodot}} on \isa{{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b}. Hence the binary type constructor \isa{{\isasymodot}} maps semigroups - to semigroups. This may be established formally as follows.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{instance}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}semigroup{\isacharcomma}\ semigroup{\isacharparenright}\ semigroup\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isacharparenleft}intro{\isacharunderscore}classes{\isacharcomma}\ unfold\ times{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ p\ q\ r\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}semigroup\ {\isasymtimes}\ {\isacharprime}b{\isasymColon}semigroup{\isachardoublequoteclose}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\isanewline -\ \ \ \ {\isachardoublequoteopen}{\isacharparenleft}fst\ {\isacharparenleft}fst\ p\ {\isasymodot}\ fst\ q{\isacharcomma}\ snd\ p\ {\isasymodot}\ snd\ q{\isacharparenright}\ {\isasymodot}\ fst\ r{\isacharcomma}\isanewline -\ \ \ \ \ \ snd\ {\isacharparenleft}fst\ p\ {\isasymodot}\ fst\ q{\isacharcomma}\ snd\ p\ {\isasymodot}\ snd\ q{\isacharparenright}\ {\isasymodot}\ snd\ r{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ \ \ {\isacharparenleft}fst\ p\ {\isasymodot}\ fst\ {\isacharparenleft}fst\ q\ {\isasymodot}\ fst\ r{\isacharcomma}\ snd\ q\ {\isasymodot}\ snd\ r{\isacharparenright}{\isacharcomma}\isanewline -\ \ \ \ \ \ \ \ snd\ p\ {\isasymodot}\ snd\ {\isacharparenleft}fst\ q\ {\isasymodot}\ fst\ r{\isacharcomma}\ snd\ q\ {\isasymodot}\ snd\ r{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ semigroup{\isacharunderscore}class{\isachardot}assoc{\isacharparenright}\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Thus, if we view class instances as ``structures'', then overloaded - constant definitions with recursion over types indirectly provide - some kind of ``functors'' --- i.e.\ mappings between abstract - theories.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/Group/document/Product.tex --- a/doc-src/AxClass/Group/document/Product.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,133 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Product}% -% -\isamarkupheader{Syntactic classes% -} -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Product\ \isakeyword{imports}\ Main\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\begin{isamarkuptext}% -\medskip\noindent There is still a feature of Isabelle's type system - left that we have not yet discussed. When declaring polymorphic - constants \isa{c\ {\isasymColon}\ {\isasymsigma}}, the type variables occurring in \isa{{\isasymsigma}} - may be constrained by type classes (or even general sorts) in an - arbitrary way. Note that by default, in Isabelle/HOL the - declaration \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} is actually an abbreviation - for \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}type\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} Since class \isa{type} is the - universal class of HOL, this is not really a constraint at all. - - The \isa{product} class below provides a less degenerate example of - syntactic type classes.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{axclass}\isamarkupfalse% -\isanewline -\ \ product\ {\isasymsubseteq}\ type\isanewline -\isacommand{consts}\isamarkupfalse% -\isanewline -\ \ product\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}product\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymodot}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}% -\begin{isamarkuptext}% -Here class \isa{product} is defined as subclass of \isa{type} - without any additional axioms. This effects in logical equivalence - of \isa{product} and \isa{type}, as is reflected by the trivial - introduction rule generated for this definition. - - \medskip So what is the difference of declaring \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}product\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} vs.\ declaring \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}type\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} anyway? In this particular case where \isa{product\ {\isasymequiv}\ type}, it should be obvious that both declarations are the same - from the logic's point of view. It even makes the most sense to - remove sort constraints from constant declarations, as far as the - purely logical meaning is concerned \cite{Wenzel:1997:TPHOL}. - - On the other hand there are syntactic differences, of course. - Constants \isa{{\isasymodot}} on some type \isa{{\isasymtau}} are rejected by the - type-checker, unless the arity \isa{{\isasymtau}\ {\isasymColon}\ product} is part of the - type signature. In our example, this arity may be always added when - required by means of an \isakeyword{instance} with the default proof - (double-dot). - - \medskip Thus, we may observe the following discipline of using - syntactic classes. Overloaded polymorphic constants have their type - arguments restricted to an associated (logically trivial) class - \isa{c}. Only immediately before \emph{specifying} these - constants on a certain type \isa{{\isasymtau}} do we instantiate \isa{{\isasymtau}\ {\isasymColon}\ c}. - - This is done for class \isa{product} and type \isa{bool} as - follows.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{instance}\isamarkupfalse% -\ bool\ {\isacharcolon}{\isacharcolon}\ product% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{defs}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline -\ \ product{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymodot}\ y\ {\isasymequiv}\ x\ {\isasymand}\ y{\isachardoublequoteclose}% -\begin{isamarkuptext}% -The definition \isa{prod{\isacharunderscore}bool{\isacharunderscore}def} becomes syntactically - well-formed only after the arity \isa{bool\ {\isasymColon}\ product} is made - known to the type checker. - - \medskip It is very important to see that above \isakeyword{defs} are - not directly connected with \isakeyword{instance} at all! We were - just following our convention to specify \isa{{\isasymodot}} on \isa{bool} - after having instantiated \isa{bool\ {\isasymColon}\ product}. Isabelle does - not require these definitions, which is in contrast to programming - languages like Haskell \cite{haskell-report}. - - \medskip While Isabelle type classes and those of Haskell are almost - the same as far as type-checking and type inference are concerned, - there are important semantic differences. Haskell classes require - their instances to \emph{provide operations} of certain \emph{names}. - Therefore, its \texttt{instance} has a \texttt{where} part that tells - the system what these ``member functions'' should be. - - This style of \texttt{instance} would not make much sense in - Isabelle's meta-logic, because there is no internal notion of - ``providing operations'' or even ``names of functions''.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/Group/document/Semigroups.tex --- a/doc-src/AxClass/Group/document/Semigroups.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,88 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Semigroups}% -% -\isamarkupheader{Semigroups% -} -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Semigroups\ \isakeyword{imports}\ Main\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\begin{isamarkuptext}% -\medskip\noindent An axiomatic type class is simply a class of types - that all meet certain properties, which are also called \emph{class - axioms}. Thus, type classes may be also understood as type - predicates --- i.e.\ abstractions over a single type argument \isa{{\isacharprime}a}. Class axioms typically contain polymorphic constants that - depend on this type \isa{{\isacharprime}a}. These \emph{characteristic - constants} behave like operations associated with the ``carrier'' - type \isa{{\isacharprime}a}. - - We illustrate these basic concepts by the following formulation of - semigroups.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{consts}\isamarkupfalse% -\isanewline -\ \ times\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymodot}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\isacommand{axclass}\isamarkupfalse% -\ semigroup\ {\isasymsubseteq}\ type\isanewline -\ \ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymodot}\ y{\isacharparenright}\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent Above we have first declared a polymorphic constant \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} and then defined the class \isa{semigroup} of - all types \isa{{\isasymtau}} such that \isa{{\isasymodot}\ {\isasymColon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}} is indeed an - associative operator. The \isa{assoc} axiom contains exactly one - type variable, which is invisible in the above presentation, though. - Also note that free term variables (like \isa{x}, \isa{y}, - \isa{z}) are allowed for user convenience --- conceptually all of - these are bound by outermost universal quantifiers. - - \medskip In general, type classes may be used to describe - \emph{structures} with exactly one carrier \isa{{\isacharprime}a} and a fixed - \emph{signature}. Different signatures require different classes. - Below, class \isa{plus{\isacharunderscore}semigroup} represents semigroups \isa{{\isacharparenleft}{\isasymtau}{\isacharcomma}\ {\isasymoplus}\isactrlsup {\isasymtau}{\isacharparenright}}, while the original \isa{semigroup} would - correspond to semigroups of the form \isa{{\isacharparenleft}{\isasymtau}{\isacharcomma}\ {\isasymodot}\isactrlsup {\isasymtau}{\isacharparenright}}.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{consts}\isamarkupfalse% -\isanewline -\ \ plus\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymoplus}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\isacommand{axclass}\isamarkupfalse% -\ plus{\isacharunderscore}semigroup\ {\isasymsubseteq}\ type\isanewline -\ \ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymoplus}\ y{\isacharparenright}\ {\isasymoplus}\ z\ {\isacharequal}\ x\ {\isasymoplus}\ {\isacharparenleft}y\ {\isasymoplus}\ z{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent Even if classes \isa{plus{\isacharunderscore}semigroup} and \isa{semigroup} both represent semigroups in a sense, they are certainly - not quite the same.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/IsaMakefile --- a/doc-src/AxClass/IsaMakefile Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,47 +0,0 @@ - -## targets - -default: Group Nat -images: -test: Group Nat - -all: images test - - -## global settings - -SRC = $(ISABELLE_HOME)/src -OUT = $(ISABELLE_OUTPUT) -LOG = $(OUT)/log -USEDIR = $(ISABELLE_TOOL) usedir -d false -D document - - -## Group - -Group: HOL $(LOG)/HOL-Group.gz - -HOL: - @cd $(SRC)/HOL; $(ISABELLE_TOOL) make HOL - -$(LOG)/HOL-Group.gz: $(OUT)/HOL Group/ROOT.ML Group/Group.thy \ - Group/Product.thy Group/Semigroups.thy - @$(USEDIR) $(OUT)/HOL Group - @rm -f Group/document/pdfsetup.sty Group/document/session.tex - - -## Nat - -Nat: FOL $(LOG)/FOL-Nat.gz - -FOL: - @cd $(SRC)/FOL; $(ISABELLE_TOOL) make FOL - -$(LOG)/FOL-Nat.gz: $(OUT)/FOL Nat/ROOT.ML Nat/NatClass.thy - @$(USEDIR) $(OUT)/FOL Nat - @rm -f Nat/document/*.sty Nat/document/session.tex - - -## clean - -clean: - @rm -f $(LOG)/HOL-Group.gz $(LOG)/FOL-Nat.gz diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/Makefile --- a/doc-src/AxClass/Makefile Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,36 +0,0 @@ -# -# $Id$ -# - -## targets - -default: dvi - - -## dependencies - -include ../Makefile.in - -NAME = axclass - -FILES = axclass.tex body.tex ../iman.sty ../extra.sty ../isar.sty \ - ../isabelle.sty ../isabellesym.sty ../pdfsetup.sty \ - Group/document/Group.tex Nat/document/NatClass.tex \ - Group/document/Product.tex Group/document/Semigroups.tex - -dvi: $(NAME).dvi - -$(NAME).dvi: $(FILES) isabelle_isar.eps - $(LATEX) $(NAME) - $(BIBTEX) $(NAME) - $(LATEX) $(NAME) - $(LATEX) $(NAME) - -pdf: $(NAME).pdf - -$(NAME).pdf: $(FILES) isabelle_isar.pdf - $(PDFLATEX) $(NAME) - $(FIXBOOKMARKS) $(NAME).out - $(BIBTEX) $(NAME) - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/Nat/NatClass.thy --- a/doc-src/AxClass/Nat/NatClass.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,117 +0,0 @@ - -header {* Defining natural numbers in FOL \label{sec:ex-natclass} *} - -theory NatClass imports FOL begin - -text {* - \medskip\noindent Axiomatic type classes abstract over exactly one - type argument. Thus, any \emph{axiomatic} theory extension where each - axiom refers to at most one type variable, may be trivially turned - into a \emph{definitional} one. - - We illustrate this with the natural numbers in - Isabelle/FOL.\footnote{See also - \url{http://isabelle.in.tum.de/library/FOL/ex/NatClass.html}} -*} - -consts - zero :: 'a ("\") - Suc :: "'a \ 'a" - rec :: "'a \ 'a \ ('a \ 'a \ 'a) \ 'a" - -axclass nat \ "term" - induct: "P(\) \ (\x. P(x) \ P(Suc(x))) \ P(n)" - Suc_inject: "Suc(m) = Suc(n) \ m = n" - Suc_neq_0: "Suc(m) = \ \ R" - rec_0: "rec(\, a, f) = a" - rec_Suc: "rec(Suc(m), a, f) = f(m, rec(m, a, f))" - -constdefs - add :: "'a::nat \ 'a \ 'a" (infixl "+" 60) - "m + n \ rec(m, n, \x y. Suc(y))" - -text {* - This is an abstract version of the plain @{text Nat} theory in - FOL.\footnote{See - \url{http://isabelle.in.tum.de/library/FOL/ex/Nat.html}} Basically, - we have just replaced all occurrences of type @{text nat} by @{typ - 'a} and used the natural number axioms to define class @{text nat}. - There is only a minor snag, that the original recursion operator - @{term rec} had to be made monomorphic. - - Thus class @{text nat} contains exactly those types @{text \} that - are isomorphic to ``the'' natural numbers (with signature @{term - \}, @{term Suc}, @{term rec}). - - \medskip What we have done here can be also viewed as \emph{type - specification}. Of course, it still remains open if there is some - type at all that meets the class axioms. Now a very nice property of - axiomatic type classes is that abstract reasoning is always possible - --- independent of satisfiability. The meta-logic won't break, even - if some classes (or general sorts) turns out to be empty later --- - ``inconsistent'' class definitions may be useless, but do not cause - any harm. - - Theorems of the abstract natural numbers may be derived in the same - way as for the concrete version. The original proof scripts may be - re-used with some trivial changes only (mostly adding some type - constraints). -*} - -(*<*) -lemma Suc_n_not_n: "Suc(k) ~= (k::'a::nat)" -apply (rule_tac n = k in induct) -apply (rule notI) -apply (erule Suc_neq_0) -apply (rule notI) -apply (erule notE) -apply (erule Suc_inject) -done - -lemma "(k+m)+n = k+(m+n)" -apply (rule induct) -back -back -back -back -back -back -oops - -lemma add_0 [simp]: "\+n = n" -apply (unfold add_def) -apply (rule rec_0) -done - -lemma add_Suc [simp]: "Suc(m)+n = Suc(m+n)" -apply (unfold add_def) -apply (rule rec_Suc) -done - -lemma add_assoc: "(k+m)+n = k+(m+n)" -apply (rule_tac n = k in induct) -apply simp -apply simp -done - -lemma add_0_right: "m+\ = m" -apply (rule_tac n = m in induct) -apply simp -apply simp -done - -lemma add_Suc_right: "m+Suc(n) = Suc(m+n)" -apply (rule_tac n = m in induct) -apply simp_all -done - -lemma - assumes prem: "!!n. f(Suc(n)) = Suc(f(n))" - shows "f(i+j) = i+f(j)" -apply (rule_tac n = i in induct) -apply simp -apply (simp add: prem) -done -(*>*) - -end \ No newline at end of file diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/Nat/ROOT.ML --- a/doc-src/AxClass/Nat/ROOT.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2 +0,0 @@ - -use_thy "NatClass"; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/Nat/document/NatClass.tex --- a/doc-src/AxClass/Nat/document/NatClass.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,201 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{NatClass}% -% -\isamarkupheader{Defining natural numbers in FOL \label{sec:ex-natclass}% -} -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ NatClass\ \isakeyword{imports}\ FOL\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\begin{isamarkuptext}% -\medskip\noindent Axiomatic type classes abstract over exactly one - type argument. Thus, any \emph{axiomatic} theory extension where each - axiom refers to at most one type variable, may be trivially turned - into a \emph{definitional} one. - - We illustrate this with the natural numbers in - Isabelle/FOL.\footnote{See also - \url{http://isabelle.in.tum.de/library/FOL/ex/NatClass.html}}% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{consts}\isamarkupfalse% -\isanewline -\ \ zero\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ \ \ \ {\isacharparenleft}{\isachardoublequoteopen}{\isasymzero}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\ \ Suc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\isanewline -\ \ rec\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{axclass}\isamarkupfalse% -\ nat\ {\isasymsubseteq}\ {\isachardoublequoteopen}term{\isachardoublequoteclose}\isanewline -\ \ induct{\isacharcolon}\ {\isachardoublequoteopen}P{\isacharparenleft}{\isasymzero}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ P{\isacharparenleft}x{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharparenleft}Suc{\isacharparenleft}x{\isacharparenright}{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharparenleft}n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ Suc{\isacharunderscore}inject{\isacharcolon}\ {\isachardoublequoteopen}Suc{\isacharparenleft}m{\isacharparenright}\ {\isacharequal}\ Suc{\isacharparenleft}n{\isacharparenright}\ {\isasymLongrightarrow}\ m\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ Suc{\isacharunderscore}neq{\isacharunderscore}{\isadigit{0}}{\isacharcolon}\ {\isachardoublequoteopen}Suc{\isacharparenleft}m{\isacharparenright}\ {\isacharequal}\ {\isasymzero}\ {\isasymLongrightarrow}\ R{\isachardoublequoteclose}\isanewline -\ \ rec{\isacharunderscore}{\isadigit{0}}{\isacharcolon}\ {\isachardoublequoteopen}rec{\isacharparenleft}{\isasymzero}{\isacharcomma}\ a{\isacharcomma}\ f{\isacharparenright}\ {\isacharequal}\ a{\isachardoublequoteclose}\isanewline -\ \ rec{\isacharunderscore}Suc{\isacharcolon}\ {\isachardoublequoteopen}rec{\isacharparenleft}Suc{\isacharparenleft}m{\isacharparenright}{\isacharcomma}\ a{\isacharcomma}\ f{\isacharparenright}\ {\isacharequal}\ f{\isacharparenleft}m{\isacharcomma}\ rec{\isacharparenleft}m{\isacharcomma}\ a{\isacharcomma}\ f{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{constdefs}\isamarkupfalse% -\isanewline -\ \ add\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isacharcolon}{\isacharcolon}nat\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isacharplus}{\isachardoublequoteclose}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ {\isachardoublequoteopen}m\ {\isacharplus}\ n\ {\isasymequiv}\ rec{\isacharparenleft}m{\isacharcomma}\ n{\isacharcomma}\ {\isasymlambda}x\ y{\isachardot}\ Suc{\isacharparenleft}y{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -This is an abstract version of the plain \isa{Nat} theory in - FOL.\footnote{See - \url{http://isabelle.in.tum.de/library/FOL/ex/Nat.html}} Basically, - we have just replaced all occurrences of type \isa{nat} by \isa{{\isacharprime}a} and used the natural number axioms to define class \isa{nat}. - There is only a minor snag, that the original recursion operator - \isa{rec} had to be made monomorphic. - - Thus class \isa{nat} contains exactly those types \isa{{\isasymtau}} that - are isomorphic to ``the'' natural numbers (with signature \isa{{\isasymzero}}, \isa{Suc}, \isa{rec}). - - \medskip What we have done here can be also viewed as \emph{type - specification}. Of course, it still remains open if there is some - type at all that meets the class axioms. Now a very nice property of - axiomatic type classes is that abstract reasoning is always possible - --- independent of satisfiability. The meta-logic won't break, even - if some classes (or general sorts) turns out to be empty later --- - ``inconsistent'' class definitions may be useless, but do not cause - any harm. - - Theorems of the abstract natural numbers may be derived in the same - way as for the concrete version. The original proof scripts may be - re-used with some trivial changes only (mostly adding some type - constraints).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/axclass.tex --- a/doc-src/AxClass/axclass.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,80 +0,0 @@ - -\documentclass[12pt,a4paper,fleqn]{report} -\usepackage{graphicx,../iman,../extra,../isar} -\usepackage{../isabelle,../isabellesym} -\usepackage{../pdfsetup} % last one! - -\isabellestyle{it} -\newcommand{\isasyminv}{\isamath{{}^{-1}}} -\renewcommand{\isasymzero}{\isamath{0}} -\renewcommand{\isasymone}{\isamath{1}} - -\newcommand{\secref}[1]{\S\ref{#1}} -\newcommand{\figref}[1]{figure~\ref{#1}} - -\hyphenation{Isabelle} -\hyphenation{Isar} -\hyphenation{Haskell} - -\title{\includegraphics[scale=0.5]{isabelle_isar} - \\[4ex] Using Axiomatic Type Classes in Isabelle} -\author{\emph{Markus Wenzel} \\ TU M\"unchen} - - -\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} - -\pagestyle{headings} -\sloppy -\binperiod %%%treat . like a binary operator - - -\begin{document} - -\underscoreoff - -\maketitle - -\begin{abstract} - Isabelle offers order-sorted type classes on top of the simple types of - plain Higher-Order Logic. The resulting type system is similar to that of - the programming language Haskell. Its interpretation within the logic - enables further application, though, apart from restricting polymorphism - syntactically. In particular, the concept of \emph{Axiomatic Type Classes} - provides a useful light-weight mechanism for hierarchically-structured - abstract theories. Subsequently, we demonstrate typical uses of Isabelle's - axiomatic type classes to model basic algebraic structures. - - This document describes axiomatic type classes using Isabelle/Isar theories, - with proofs expressed via Isar proof language elements. The new theory - format greatly simplifies the arrangement of the overall development, since - definitions and proofs may be freely intermixed. Users who prefer tactic - scripts over structured proofs do not need to fall back on separate ML - scripts, though, but may refer to Isar's tactic emulation commands. -\end{abstract} - - -\pagenumbering{roman} \tableofcontents \clearfirst - -\include{body} - -%FIXME -\nocite{nipkow-types93} -\nocite{nipkow-sorts93} -\nocite{Wenzel:1997:TPHOL} -\nocite{paulson-isa-book} -\nocite{isabelle-isar-ref} -\nocite{Wenzel:1999:TPHOL} - -\begingroup - \bibliographystyle{plain} \small\raggedright\frenchspacing - \bibliography{../manual} -\endgroup - -\end{document} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: t -%%% End: -% LocalWords: Isabelle FIXME diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/AxClass/body.tex --- a/doc-src/AxClass/body.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,166 +0,0 @@ - -\chapter{Introduction} - -A Haskell-style type-system \cite{haskell-report} with ordered type-classes -has been present in Isabelle since 1991 already \cite{nipkow-sorts93}. -Initially, classes have mainly served as a \emph{purely syntactic} tool to -formulate polymorphic object-logics in a clean way, such as the standard -Isabelle formulation of many-sorted FOL \cite{paulson-isa-book}. - -Applying classes at the \emph{logical level} to provide a simple notion of -abstract theories and instantiations to concrete ones, has been long proposed -as well \cite{nipkow-types93,nipkow-sorts93}. At that time, Isabelle still -lacked built-in support for these \emph{axiomatic type classes}. More -importantly, their semantics was not yet fully fleshed out (and unnecessarily -complicated, too). - -Since Isabelle94, actual axiomatic type classes have been an integral part of -Isabelle's meta-logic. This very simple implementation is based on a -straight-forward extension of traditional simply-typed Higher-Order Logic, by -including types qualified by logical predicates and overloaded constant -definitions (see \cite{Wenzel:1997:TPHOL} for further details). - -Yet even until Isabelle99, there used to be still a fundamental methodological -problem in using axiomatic type classes conveniently, due to the traditional -distinction of Isabelle theory files vs.\ ML proof scripts. This has been -finally overcome with the advent of Isabelle/Isar theories -\cite{isabelle-isar-ref}: now definitions and proofs may be freely intermixed. -This nicely accommodates the usual procedure of defining axiomatic type -classes, proving abstract properties, defining operations on concrete types, -proving concrete properties for instantiation of classes etc. - -\medskip - -So to cut a long story short, the present version of axiomatic type classes -now provides an even more useful and convenient mechanism for light-weight -abstract theories, without any special technical provisions to be observed by -the user. - - -\chapter{Examples}\label{sec:ex} - -Axiomatic type classes are a concept of Isabelle's meta-logic -\cite{paulson-isa-book,Wenzel:1997:TPHOL}. They may be applied to any -object-logic that directly uses the meta type system, such as Isabelle/HOL -\cite{isabelle-HOL}. Subsequently, we present various examples that are all -formulated within HOL, except the one of \secref{sec:ex-natclass} which is in -FOL. See also \url{http://isabelle.in.tum.de/library/HOL/AxClasses/} and -\url{http://isabelle.in.tum.de/library/FOL/ex/NatClass.html}. - -\input{Group/document/Semigroups} - -\input{Group/document/Group} - -\input{Group/document/Product} - -\input{Nat/document/NatClass} - - -%% FIXME move some parts to ref or isar-ref manual (!?); - -% \chapter{The user interface of Isabelle's axclass package} - -% The actual axiomatic type class package of Isabelle/Pure mainly consists -% of two new theory sections: \texttt{axclass} and \texttt{instance}. Some -% typical applications of these have already been demonstrated in -% \secref{sec:ex}, below their syntax and semantics are presented more -% completely. - - -% \section{The axclass section} - -% Within theory files, \texttt{axclass} introduces an axiomatic type class -% definition. Its concrete syntax is: - -% \begin{matharray}{l} -% \texttt{axclass} \\ -% \ \ c \texttt{ < } c@1\texttt, \ldots\texttt, c@n \\ -% \ \ id@1\ axm@1 \\ -% \ \ \vdots \\ -% \ \ id@m\ axm@m -% \emphnd{matharray} - -% Where $c, c@1, \ldots, c@n$ are classes (category $id$ or -% $string$) and $axm@1, \ldots, axm@m$ (with $m \geq -% 0$) are formulas (category $string$). - -% Class $c$ has to be new, and sort $\{c@1, \ldots, c@n\}$ a subsort of -% \texttt{logic}. Each class axiom $axm@j$ may contain any term -% variables, but at most one type variable (which need not be the same -% for all axioms). The sort of this type variable has to be a supersort -% of $\{c@1, \ldots, c@n\}$. - -% \medskip - -% The \texttt{axclass} section declares $c$ as subclass of $c@1, \ldots, -% c@n$ to the type signature. - -% Furthermore, $axm@1, \ldots, axm@m$ are turned into the -% ``abstract axioms'' of $c$ with names $id@1, \ldots, -% id@m$. This is done by replacing all occurring type variables -% by $\alpha :: c$. Original axioms that do not contain any type -% variable will be prefixed by the logical precondition -% $\texttt{OFCLASS}(\alpha :: \texttt{logic}, c\texttt{_class})$. - -% Another axiom of name $c\texttt{I}$ --- the ``class $c$ introduction -% rule'' --- is built from the respective universal closures of -% $axm@1, \ldots, axm@m$ appropriately. - - -% \section{The instance section} - -% Section \texttt{instance} proves class inclusions or type arities at the -% logical level and then transfers these into the type signature. - -% Its concrete syntax is: - -% \begin{matharray}{l} -% \texttt{instance} \\ -% \ \ [\ c@1 \texttt{ < } c@2 \ |\ -% t \texttt{ ::\ (}sort@1\texttt, \ldots \texttt, sort@n\texttt) sort\ ] \\ -% \ \ [\ \texttt(name@1 \texttt, \ldots\texttt, name@m\texttt)\ ] \\ -% \ \ [\ \texttt{\{|} text \texttt{|\}}\ ] -% \emphnd{matharray} - -% Where $c@1, c@2$ are classes and $t$ is an $n$-place type constructor -% (all of category $id$ or $string)$. Furthermore, -% $sort@i$ are sorts in the usual Isabelle-syntax. - -% \medskip - -% Internally, \texttt{instance} first sets up an appropriate goal that -% expresses the class inclusion or type arity as a meta-proposition. -% Then tactic \texttt{AxClass.axclass_tac} is applied with all preceding -% meta-definitions of the current theory file and the user-supplied -% witnesses. The latter are $name@1, \ldots, name@m$, where -% $id$ refers to an \ML-name of a theorem, and $string$ to an -% axiom of the current theory node\footnote{Thus, the user may reference -% axioms from above this \texttt{instance} in the theory file. Note -% that new axioms appear at the \ML-toplevel only after the file is -% processed completely.}. - -% Tactic \texttt{AxClass.axclass_tac} first unfolds the class definition by -% resolving with rule $c\texttt\texttt{I}$, and then applies the witnesses -% according to their form: Meta-definitions are unfolded, all other -% formulas are repeatedly resolved\footnote{This is done in a way that -% enables proper object-\emph{rules} to be used as witnesses for -% corresponding class axioms.} with. - -% The final optional argument $text$ is \ML-code of an arbitrary -% user tactic which is applied last to any remaining goals. - -% \medskip - -% Because of the complexity of \texttt{instance}'s witnessing mechanisms, -% new users of the axclass package are advised to only use the simple -% form $\texttt{instance}\ \ldots\ (id@1, \ldots, id@!m)$, where -% the identifiers refer to theorems that are appropriate type instances -% of the class axioms. This typically requires an auxiliary theory, -% though, which defines some constants and then proves these witnesses. - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "axclass" -%%% End: -% LocalWords: Isabelle FOL diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Classes/IsaMakefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/IsaMakefile Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,33 @@ + +## targets + +default: Thy +images: +test: Thy + +all: images test + + +## global settings + +SRC = $(ISABELLE_HOME)/src +OUT = $(ISABELLE_OUTPUT) +LOG = $(OUT)/log + +USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -d false -C false -D document + + +## Thy + +THY = $(LOG)/HOL-Thy.gz + +Thy: $(THY) + +$(THY): Thy/ROOT.ML Thy/Setup.thy Thy/Classes.thy ../antiquote_setup.ML ../more_antiquote.ML + @$(USEDIR) HOL Thy + + +## clean + +clean: + @rm -f $(THY) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Classes/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/Makefile Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,35 @@ + +## targets + +default: dvi + + +## dependencies + +include ../Makefile.in + +NAME = classes + +FILES = $(NAME).tex classes.tex Thy/document/Classes.tex \ + style.sty ../iman.sty ../extra.sty ../isar.sty \ + ../isabelle.sty ../isabellesym.sty ../pdfsetup.sty \ + ../manual.bib ../proof.sty + +dvi: $(NAME).dvi + +$(NAME).dvi: $(FILES) isabelle_isar.eps + $(LATEX) $(NAME) + $(BIBTEX) $(NAME) + $(LATEX) $(NAME) + $(LATEX) $(NAME) + +pdf: $(NAME).pdf + +$(NAME).pdf: $(FILES) isabelle_isar.pdf + $(PDFLATEX) $(NAME) + $(BIBTEX) $(NAME) + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) + $(FIXBOOKMARKS) $(NAME).out + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Classes/Thy/Classes.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/Thy/Classes.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,634 @@ +theory Classes +imports Main Setup +begin + +section {* Introduction *} + +text {* + Type classes were introduces by Wadler and Blott \cite{wadler89how} + into the Haskell language, to allow for a reasonable implementation + of overloading\footnote{throughout this tutorial, we are referring + to classical Haskell 1.0 type classes, not considering + later additions in expressiveness}. + As a canonical example, a polymorphic equality function + @{text "eq \ \ \ \ \ bool"} which is overloaded on different + types for @{text "\"}, which is achieved by splitting introduction + of the @{text eq} function from its overloaded definitions by means + of @{text class} and @{text instance} declarations: + \footnote{syntax here is a kind of isabellized Haskell} + + \begin{quote} + + \noindent@{text "class eq where"} \\ + \hspace*{2ex}@{text "eq \ \ \ \ \ bool"} + + \medskip\noindent@{text "instance nat \ eq where"} \\ + \hspace*{2ex}@{text "eq 0 0 = True"} \\ + \hspace*{2ex}@{text "eq 0 _ = False"} \\ + \hspace*{2ex}@{text "eq _ 0 = False"} \\ + \hspace*{2ex}@{text "eq (Suc n) (Suc m) = eq n m"} + + \medskip\noindent@{text "instance (\\eq, \\eq) pair \ eq where"} \\ + \hspace*{2ex}@{text "eq (x1, y1) (x2, y2) = eq x1 x2 \ eq y1 y2"} + + \medskip\noindent@{text "class ord extends eq where"} \\ + \hspace*{2ex}@{text "less_eq \ \ \ \ \ bool"} \\ + \hspace*{2ex}@{text "less \ \ \ \ \ bool"} + + \end{quote} + + \noindent Type variables are annotated with (finitely many) classes; + these annotations are assertions that a particular polymorphic type + provides definitions for overloaded functions. + + Indeed, type classes not only allow for simple overloading + but form a generic calculus, an instance of order-sorted + algebra \cite{Nipkow-Prehofer:1993,nipkow-sorts93,Wenzel:1997:TPHOL}. + + From a software engeneering point of view, type classes + roughly correspond to interfaces in object-oriented languages like Java; + so, it is naturally desirable that type classes do not only + provide functions (class parameters) but also state specifications + implementations must obey. For example, the @{text "class eq"} + above could be given the following specification, demanding that + @{text "class eq"} is an equivalence relation obeying reflexivity, + symmetry and transitivity: + + \begin{quote} + + \noindent@{text "class eq where"} \\ + \hspace*{2ex}@{text "eq \ \ \ \ \ bool"} \\ + @{text "satisfying"} \\ + \hspace*{2ex}@{text "refl: eq x x"} \\ + \hspace*{2ex}@{text "sym: eq x y \ eq x y"} \\ + \hspace*{2ex}@{text "trans: eq x y \ eq y z \ eq x z"} + + \end{quote} + + \noindent From a theoretic point of view, type classes are lightweight + modules; Haskell type classes may be emulated by + SML functors \cite{classes_modules}. + Isabelle/Isar offers a discipline of type classes which brings + all those aspects together: + + \begin{enumerate} + \item specifying abstract parameters together with + corresponding specifications, + \item instantiating those abstract parameters by a particular + type + \item in connection with a ``less ad-hoc'' approach to overloading, + \item with a direct link to the Isabelle module system + (aka locales \cite{kammueller-locales}). + \end{enumerate} + + \noindent Isar type classes also directly support code generation + in a Haskell like fashion. + + This tutorial demonstrates common elements of structured specifications + and abstract reasoning with type classes by the algebraic hierarchy of + semigroups, monoids and groups. Our background theory is that of + Isabelle/HOL \cite{isa-tutorial}, for which some + familiarity is assumed. + + Here we merely present the look-and-feel for end users. + Internally, those are mapped to more primitive Isabelle concepts. + See \cite{Haftmann-Wenzel:2006:classes} for more detail. +*} + +section {* A simple algebra example \label{sec:example} *} + +subsection {* Class definition *} + +text {* + Depending on an arbitrary type @{text "\"}, class @{text + "semigroup"} introduces a binary operator @{text "(\)"} that is + assumed to be associative: +*} + +class %quote semigroup = + fixes mult :: "\ \ \ \ \" (infixl "\" 70) + assumes assoc: "(x \ y) \ z = x \ (y \ z)" + +text {* + \noindent This @{command class} specification consists of two + parts: the \qn{operational} part names the class parameter + (@{element "fixes"}), the \qn{logical} part specifies properties on them + (@{element "assumes"}). The local @{element "fixes"} and + @{element "assumes"} are lifted to the theory toplevel, + yielding the global + parameter @{term [source] "mult \ \\semigroup \ \ \ \"} and the + global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\x y + z \ \\semigroup. (x \ y) \ z = x \ (y \ z)"}. +*} + + +subsection {* Class instantiation \label{sec:class_inst} *} + +text {* + The concrete type @{typ int} is made a @{class semigroup} + instance by providing a suitable definition for the class parameter + @{text "(\)"} and a proof for the specification of @{fact assoc}. + This is accomplished by the @{command instantiation} target: +*} + +instantiation %quote int :: semigroup +begin + +definition %quote + mult_int_def: "i \ j = i + (j\int)" + +instance %quote proof + fix i j k :: int have "(i + j) + k = i + (j + k)" by simp + then show "(i \ j) \ k = i \ (j \ k)" + unfolding mult_int_def . +qed + +end %quote + +text {* + \noindent @{command instantiation} allows to define class parameters + at a particular instance using common specification tools (here, + @{command definition}). The concluding @{command instance} + opens a proof that the given parameters actually conform + to the class specification. Note that the first proof step + is the @{method default} method, + which for such instance proofs maps to the @{method intro_classes} method. + This boils down an instance judgement to the relevant primitive + proof goals and should conveniently always be the first method applied + in an instantiation proof. + + From now on, the type-checker will consider @{typ int} + as a @{class semigroup} automatically, i.e.\ any general results + are immediately available on concrete instances. + + \medskip Another instance of @{class semigroup} are the natural numbers: +*} + +instantiation %quote nat :: semigroup +begin + +primrec %quote mult_nat where + "(0\nat) \ n = n" + | "Suc m \ n = Suc (m \ n)" + +instance %quote proof + fix m n q :: nat + show "m \ n \ q = m \ (n \ q)" + by (induct m) auto +qed + +end %quote + +text {* + \noindent Note the occurence of the name @{text mult_nat} + in the primrec declaration; by default, the local name of + a class operation @{text f} to instantiate on type constructor + @{text \} are mangled as @{text f_\}. In case of uncertainty, + these names may be inspected using the @{command "print_context"} command + or the corresponding ProofGeneral button. +*} + +subsection {* Lifting and parametric types *} + +text {* + Overloaded definitions giving on class instantiation + may include recursion over the syntactic structure of types. + As a canonical example, we model product semigroups + using our simple algebra: +*} + +instantiation %quote * :: (semigroup, semigroup) semigroup +begin + +definition %quote + mult_prod_def: "p\<^isub>1 \ p\<^isub>2 = (fst p\<^isub>1 \ fst p\<^isub>2, snd p\<^isub>1 \ snd p\<^isub>2)" + +instance %quote proof + fix p\<^isub>1 p\<^isub>2 p\<^isub>3 :: "\\semigroup \ \\semigroup" + show "p\<^isub>1 \ p\<^isub>2 \ p\<^isub>3 = p\<^isub>1 \ (p\<^isub>2 \ p\<^isub>3)" + unfolding mult_prod_def by (simp add: assoc) +qed + +end %quote + +text {* + \noindent Associativity from product semigroups is + established using + the definition of @{text "(\)"} on products and the hypothetical + associativity of the type components; these hypotheses + are facts due to the @{class semigroup} constraints imposed + on the type components by the @{command instance} proposition. + Indeed, this pattern often occurs with parametric types + and type classes. +*} + + +subsection {* Subclassing *} + +text {* + We define a subclass @{text monoidl} (a semigroup with a left-hand neutral) + by extending @{class semigroup} + with one additional parameter @{text neutral} together + with its property: +*} + +class %quote monoidl = semigroup + + fixes neutral :: "\" ("\") + assumes neutl: "\ \ x = x" + +text {* + \noindent Again, we prove some instances, by + providing suitable parameter definitions and proofs for the + additional specifications. Observe that instantiations + for types with the same arity may be simultaneous: +*} + +instantiation %quote nat and int :: monoidl +begin + +definition %quote + neutral_nat_def: "\ = (0\nat)" + +definition %quote + neutral_int_def: "\ = (0\int)" + +instance %quote proof + fix n :: nat + show "\ \ n = n" + unfolding neutral_nat_def by simp +next + fix k :: int + show "\ \ k = k" + unfolding neutral_int_def mult_int_def by simp +qed + +end %quote + +instantiation %quote * :: (monoidl, monoidl) monoidl +begin + +definition %quote + neutral_prod_def: "\ = (\, \)" + +instance %quote proof + fix p :: "\\monoidl \ \\monoidl" + show "\ \ p = p" + unfolding neutral_prod_def mult_prod_def by (simp add: neutl) +qed + +end %quote + +text {* + \noindent Fully-fledged monoids are modelled by another subclass + which does not add new parameters but tightens the specification: +*} + +class %quote monoid = monoidl + + assumes neutr: "x \ \ = x" + +instantiation %quote nat and int :: monoid +begin + +instance %quote proof + fix n :: nat + show "n \ \ = n" + unfolding neutral_nat_def by (induct n) simp_all +next + fix k :: int + show "k \ \ = k" + unfolding neutral_int_def mult_int_def by simp +qed + +end %quote + +instantiation %quote * :: (monoid, monoid) monoid +begin + +instance %quote proof + fix p :: "\\monoid \ \\monoid" + show "p \ \ = p" + unfolding neutral_prod_def mult_prod_def by (simp add: neutr) +qed + +end %quote + +text {* + \noindent To finish our small algebra example, we add a @{text group} class + with a corresponding instance: +*} + +class %quote group = monoidl + + fixes inverse :: "\ \ \" ("(_\
)" [1000] 999) + assumes invl: "x\
\ x = \" + +instantiation %quote int :: group +begin + +definition %quote + inverse_int_def: "i\
= - (i\int)" + +instance %quote proof + fix i :: int + have "-i + i = 0" by simp + then show "i\
\ i = \" + unfolding mult_int_def neutral_int_def inverse_int_def . +qed + +end %quote + + +section {* Type classes as locales *} + +subsection {* A look behind the scene *} + +text {* + The example above gives an impression how Isar type classes work + in practice. As stated in the introduction, classes also provide + a link to Isar's locale system. Indeed, the logical core of a class + is nothing else than a locale: +*} + +class %quote idem = + fixes f :: "\ \ \" + assumes idem: "f (f x) = f x" + +text {* + \noindent essentially introduces the locale +*} (*<*)setup %invisible {* Sign.add_path "foo" *} +(*>*) +locale %quote idem = + fixes f :: "\ \ \" + assumes idem: "f (f x) = f x" + +text {* \noindent together with corresponding constant(s): *} + +consts %quote f :: "\ \ \" + +text {* + \noindent The connection to the type system is done by means + of a primitive axclass +*} (*<*)setup %invisible {* Sign.add_path "foo" *} +(*>*) +axclass %quote idem < type + idem: "f (f x) = f x" (*<*)setup %invisible {* Sign.parent_path *}(*>*) + +text {* \noindent together with a corresponding interpretation: *} + +interpretation %quote idem_class: + idem "f \ (\\idem) \ \" +proof qed (rule idem) + +text {* + \noindent This gives you at hand the full power of the Isabelle module system; + conclusions in locale @{text idem} are implicitly propagated + to class @{text idem}. +*} (*<*)setup %invisible {* Sign.parent_path *} +(*>*) +subsection {* Abstract reasoning *} + +text {* + Isabelle locales enable reasoning at a general level, while results + are implicitly transferred to all instances. For example, we can + now establish the @{text "left_cancel"} lemma for groups, which + states that the function @{text "(x \)"} is injective: +*} + +lemma %quote (in group) left_cancel: "x \ y = x \ z \ y = z" +proof + assume "x \ y = x \ z" + then have "x\
\ (x \ y) = x\
\ (x \ z)" by simp + then have "(x\
\ x) \ y = (x\
\ x) \ z" using assoc by simp + then show "y = z" using neutl and invl by simp +next + assume "y = z" + then show "x \ y = x \ z" by simp +qed + +text {* + \noindent Here the \qt{@{keyword "in"} @{class group}} target specification + indicates that the result is recorded within that context for later + use. This local theorem is also lifted to the global one @{fact + "group.left_cancel:"} @{prop [source] "\x y z \ \\group. x \ y = x \ + z \ y = z"}. Since type @{text "int"} has been made an instance of + @{text "group"} before, we may refer to that fact as well: @{prop + [source] "\x y z \ int. x \ y = x \ z \ y = z"}. +*} + + +subsection {* Derived definitions *} + +text {* + Isabelle locales support a concept of local definitions + in locales: +*} + +primrec %quote (in monoid) pow_nat :: "nat \ \ \ \" where + "pow_nat 0 x = \" + | "pow_nat (Suc n) x = x \ pow_nat n x" + +text {* + \noindent If the locale @{text group} is also a class, this local + definition is propagated onto a global definition of + @{term [source] "pow_nat \ nat \ \\monoid \ \\monoid"} + with corresponding theorems + + @{thm pow_nat.simps [no_vars]}. + + \noindent As you can see from this example, for local + definitions you may use any specification tool + which works together with locales (e.g. \cite{krauss2006}). +*} + + +subsection {* A functor analogy *} + +text {* + We introduced Isar classes by analogy to type classes + functional programming; if we reconsider this in the + context of what has been said about type classes and locales, + we can drive this analogy further by stating that type + classes essentially correspond to functors which have + a canonical interpretation as type classes. + Anyway, there is also the possibility of other interpretations. + For example, also @{text list}s form a monoid with + @{text append} and @{term "[]"} as operations, but it + seems inappropriate to apply to lists + the same operations as for genuinely algebraic types. + In such a case, we simply can do a particular interpretation + of monoids for lists: +*} + +interpretation %quote list_monoid!: monoid append "[]" + proof qed auto + +text {* + \noindent This enables us to apply facts on monoids + to lists, e.g. @{thm list_monoid.neutl [no_vars]}. + + When using this interpretation pattern, it may also + be appropriate to map derived definitions accordingly: +*} + +primrec %quote replicate :: "nat \ \ list \ \ list" where + "replicate 0 _ = []" + | "replicate (Suc n) xs = xs @ replicate n xs" + +interpretation %quote list_monoid!: monoid append "[]" where + "monoid.pow_nat append [] = replicate" +proof - + interpret monoid append "[]" .. + show "monoid.pow_nat append [] = replicate" + proof + fix n + show "monoid.pow_nat append [] n = replicate n" + by (induct n) auto + qed +qed intro_locales + + +subsection {* Additional subclass relations *} + +text {* + Any @{text "group"} is also a @{text "monoid"}; this + can be made explicit by claiming an additional + subclass relation, + together with a proof of the logical difference: +*} + +subclass %quote (in group) monoid +proof + fix x + from invl have "x\
\ x = \" by simp + with assoc [symmetric] neutl invl have "x\
\ (x \ \) = x\
\ x" by simp + with left_cancel show "x \ \ = x" by simp +qed + +text {* + The logical proof is carried out on the locale level. + Afterwards it is propagated + to the type system, making @{text group} an instance of + @{text monoid} by adding an additional edge + to the graph of subclass relations + (cf.\ \figref{fig:subclass}). + + \begin{figure}[htbp] + \begin{center} + \small + \unitlength 0.6mm + \begin{picture}(40,60)(0,0) + \put(20,60){\makebox(0,0){@{text semigroup}}} + \put(20,40){\makebox(0,0){@{text monoidl}}} + \put(00,20){\makebox(0,0){@{text monoid}}} + \put(40,00){\makebox(0,0){@{text group}}} + \put(20,55){\vector(0,-1){10}} + \put(15,35){\vector(-1,-1){10}} + \put(25,35){\vector(1,-3){10}} + \end{picture} + \hspace{8em} + \begin{picture}(40,60)(0,0) + \put(20,60){\makebox(0,0){@{text semigroup}}} + \put(20,40){\makebox(0,0){@{text monoidl}}} + \put(00,20){\makebox(0,0){@{text monoid}}} + \put(40,00){\makebox(0,0){@{text group}}} + \put(20,55){\vector(0,-1){10}} + \put(15,35){\vector(-1,-1){10}} + \put(05,15){\vector(3,-1){30}} + \end{picture} + \caption{Subclass relationship of monoids and groups: + before and after establishing the relationship + @{text "group \ monoid"}; transitive edges are left out.} + \label{fig:subclass} + \end{center} + \end{figure} + + For illustration, a derived definition + in @{text group} which uses @{text pow_nat}: +*} + +definition %quote (in group) pow_int :: "int \ \ \ \" where + "pow_int k x = (if k >= 0 + then pow_nat (nat k) x + else (pow_nat (nat (- k)) x)\
)" + +text {* + \noindent yields the global definition of + @{term [source] "pow_int \ int \ \\group \ \\group"} + with the corresponding theorem @{thm pow_int_def [no_vars]}. +*} + +subsection {* A note on syntax *} + +text {* + As a commodity, class context syntax allows to refer + to local class operations and their global counterparts + uniformly; type inference resolves ambiguities. For example: +*} + +context %quote semigroup +begin + +term %quote "x \ y" -- {* example 1 *} +term %quote "(x\nat) \ y" -- {* example 2 *} + +end %quote + +term %quote "x \ y" -- {* example 3 *} + +text {* + \noindent Here in example 1, the term refers to the local class operation + @{text "mult [\]"}, whereas in example 2 the type constraint + enforces the global class operation @{text "mult [nat]"}. + In the global context in example 3, the reference is + to the polymorphic global class operation @{text "mult [?\ \ semigroup]"}. +*} + +section {* Further issues *} + +subsection {* Type classes and code generation *} + +text {* + Turning back to the first motivation for type classes, + namely overloading, it is obvious that overloading + stemming from @{command class} statements and + @{command instantiation} + targets naturally maps to Haskell type classes. + The code generator framework \cite{isabelle-codegen} + takes this into account. Concerning target languages + lacking type classes (e.g.~SML), type classes + are implemented by explicit dictionary construction. + As example, let's go back to the power function: +*} + +definition %quote example :: int where + "example = pow_int 10 (-2)" + +text {* + \noindent This maps to Haskell as: +*} + +text %quote {*@{code_stmts example (Haskell)}*} + +text {* + \noindent The whole code in SML with explicit dictionary passing: +*} + +text %quote {*@{code_stmts example (SML)}*} + +subsection {* Inspecting the type class universe *} + +text {* + To facilitate orientation in complex subclass structures, + two diagnostics commands are provided: + + \begin{description} + + \item[@{command "print_classes"}] print a list of all classes + together with associated operations etc. + + \item[@{command "class_deps"}] visualizes the subclass relation + between all classes as a Hasse diagram. + + \end{description} +*} + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Classes/Thy/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/Thy/ROOT.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,6 @@ + +(* $Id$ *) + +no_document use_thy "Setup"; + +use_thy "Classes"; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Classes/Thy/Setup.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/Thy/Setup.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,34 @@ +theory Setup +imports Main Code_Integer +uses + "../../antiquote_setup" + "../../more_antiquote" +begin + +ML {* Code_Target.code_width := 74 *} + +syntax + "_alpha" :: "type" ("\") + "_alpha_ofsort" :: "sort \ type" ("\()\_" [0] 1000) + "_beta" :: "type" ("\") + "_beta_ofsort" :: "sort \ type" ("\()\_" [0] 1000) + +parse_ast_translation {* + let + fun alpha_ast_tr [] = Syntax.Variable "'a" + | alpha_ast_tr asts = raise Syntax.AST ("alpha_ast_tr", asts); + fun alpha_ofsort_ast_tr [ast] = + Syntax.Appl [Syntax.Constant "_ofsort", Syntax.Variable "'a", ast] + | alpha_ofsort_ast_tr asts = raise Syntax.AST ("alpha_ast_tr", asts); + fun beta_ast_tr [] = Syntax.Variable "'b" + | beta_ast_tr asts = raise Syntax.AST ("beta_ast_tr", asts); + fun beta_ofsort_ast_tr [ast] = + Syntax.Appl [Syntax.Constant "_ofsort", Syntax.Variable "'b", ast] + | beta_ofsort_ast_tr asts = raise Syntax.AST ("beta_ast_tr", asts); + in [ + ("_alpha", alpha_ast_tr), ("_alpha_ofsort", alpha_ofsort_ast_tr), + ("_beta", beta_ast_tr), ("_beta_ofsort", beta_ofsort_ast_tr) + ] end +*} + +end \ No newline at end of file diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Classes/Thy/document/Classes.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/Thy/document/Classes.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,1335 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Classes}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Classes\isanewline +\isakeyword{imports}\ Main\ Setup\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsection{Introduction% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Type classes were introduces by Wadler and Blott \cite{wadler89how} + into the Haskell language, to allow for a reasonable implementation + of overloading\footnote{throughout this tutorial, we are referring + to classical Haskell 1.0 type classes, not considering + later additions in expressiveness}. + As a canonical example, a polymorphic equality function + \isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} which is overloaded on different + types for \isa{{\isasymalpha}}, which is achieved by splitting introduction + of the \isa{eq} function from its overloaded definitions by means + of \isa{class} and \isa{instance} declarations: + \footnote{syntax here is a kind of isabellized Haskell} + + \begin{quote} + + \noindent\isa{class\ eq\ where} \\ + \hspace*{2ex}\isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} + + \medskip\noindent\isa{instance\ nat\ {\isasymColon}\ eq\ where} \\ + \hspace*{2ex}\isa{eq\ {\isadigit{0}}\ {\isadigit{0}}\ {\isacharequal}\ True} \\ + \hspace*{2ex}\isa{eq\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ False} \\ + \hspace*{2ex}\isa{eq\ {\isacharunderscore}\ {\isadigit{0}}\ {\isacharequal}\ False} \\ + \hspace*{2ex}\isa{eq\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharparenleft}Suc\ m{\isacharparenright}\ {\isacharequal}\ eq\ n\ m} + + \medskip\noindent\isa{instance\ {\isacharparenleft}{\isasymalpha}{\isasymColon}eq{\isacharcomma}\ {\isasymbeta}{\isasymColon}eq{\isacharparenright}\ pair\ {\isasymColon}\ eq\ where} \\ + \hspace*{2ex}\isa{eq\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ eq\ x{\isadigit{1}}\ x{\isadigit{2}}\ {\isasymand}\ eq\ y{\isadigit{1}}\ y{\isadigit{2}}} + + \medskip\noindent\isa{class\ ord\ extends\ eq\ where} \\ + \hspace*{2ex}\isa{less{\isacharunderscore}eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \\ + \hspace*{2ex}\isa{less\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} + + \end{quote} + + \noindent Type variables are annotated with (finitely many) classes; + these annotations are assertions that a particular polymorphic type + provides definitions for overloaded functions. + + Indeed, type classes not only allow for simple overloading + but form a generic calculus, an instance of order-sorted + algebra \cite{Nipkow-Prehofer:1993,nipkow-sorts93,Wenzel:1997:TPHOL}. + + From a software engeneering point of view, type classes + roughly correspond to interfaces in object-oriented languages like Java; + so, it is naturally desirable that type classes do not only + provide functions (class parameters) but also state specifications + implementations must obey. For example, the \isa{class\ eq} + above could be given the following specification, demanding that + \isa{class\ eq} is an equivalence relation obeying reflexivity, + symmetry and transitivity: + + \begin{quote} + + \noindent\isa{class\ eq\ where} \\ + \hspace*{2ex}\isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \\ + \isa{satisfying} \\ + \hspace*{2ex}\isa{refl{\isacharcolon}\ eq\ x\ x} \\ + \hspace*{2ex}\isa{sym{\isacharcolon}\ eq\ x\ y\ {\isasymlongleftrightarrow}\ eq\ x\ y} \\ + \hspace*{2ex}\isa{trans{\isacharcolon}\ eq\ x\ y\ {\isasymand}\ eq\ y\ z\ {\isasymlongrightarrow}\ eq\ x\ z} + + \end{quote} + + \noindent From a theoretic point of view, type classes are lightweight + modules; Haskell type classes may be emulated by + SML functors \cite{classes_modules}. + Isabelle/Isar offers a discipline of type classes which brings + all those aspects together: + + \begin{enumerate} + \item specifying abstract parameters together with + corresponding specifications, + \item instantiating those abstract parameters by a particular + type + \item in connection with a ``less ad-hoc'' approach to overloading, + \item with a direct link to the Isabelle module system + (aka locales \cite{kammueller-locales}). + \end{enumerate} + + \noindent Isar type classes also directly support code generation + in a Haskell like fashion. + + This tutorial demonstrates common elements of structured specifications + and abstract reasoning with type classes by the algebraic hierarchy of + semigroups, monoids and groups. Our background theory is that of + Isabelle/HOL \cite{isa-tutorial}, for which some + familiarity is assumed. + + Here we merely present the look-and-feel for end users. + Internally, those are mapped to more primitive Isabelle concepts. + See \cite{Haftmann-Wenzel:2006:classes} for more detail.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{A simple algebra example \label{sec:example}% +} +\isamarkuptrue% +% +\isamarkupsubsection{Class definition% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Depending on an arbitrary type \isa{{\isasymalpha}}, class \isa{semigroup} introduces a binary operator \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} that is + assumed to be associative:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{class}\isamarkupfalse% +\ semigroup\ {\isacharequal}\isanewline +\ \ \isakeyword{fixes}\ mult\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymotimes}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline +\ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} specification consists of two + parts: the \qn{operational} part names the class parameter + (\hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}}), the \qn{logical} part specifies properties on them + (\hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}}). The local \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} and + \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} are lifted to the theory toplevel, + yielding the global + parameter \isa{{\isachardoublequote}mult\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequote}} and the + global theorem \hyperlink{fact.semigroup.assoc:}{\mbox{\isa{semigroup{\isachardot}assoc{\isacharcolon}}}}~\isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup{\isachardot}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequote}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Class instantiation \label{sec:class_inst}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The concrete type \isa{int} is made a \isa{semigroup} + instance by providing a suitable definition for the class parameter + \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} and a proof for the specification of \hyperlink{fact.assoc}{\mbox{\isa{assoc}}}. + This is accomplished by the \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} target:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{instantiation}\isamarkupfalse% +\ int\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ mult{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}i\ {\isasymotimes}\ j\ {\isacharequal}\ i\ {\isacharplus}\ {\isacharparenleft}j{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ i\ j\ k\ {\isacharcolon}{\isacharcolon}\ int\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isacharplus}\ j{\isacharparenright}\ {\isacharplus}\ k\ {\isacharequal}\ i\ {\isacharplus}\ {\isacharparenleft}j\ {\isacharplus}\ k{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isasymotimes}\ j{\isacharparenright}\ {\isasymotimes}\ k\ {\isacharequal}\ i\ {\isasymotimes}\ {\isacharparenleft}j\ {\isasymotimes}\ k{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} allows to define class parameters + at a particular instance using common specification tools (here, + \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}). The concluding \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} + opens a proof that the given parameters actually conform + to the class specification. Note that the first proof step + is the \hyperlink{method.default}{\mbox{\isa{default}}} method, + which for such instance proofs maps to the \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isacharunderscore}classes}}} method. + This boils down an instance judgement to the relevant primitive + proof goals and should conveniently always be the first method applied + in an instantiation proof. + + From now on, the type-checker will consider \isa{int} + as a \isa{semigroup} automatically, i.e.\ any general results + are immediately available on concrete instances. + + \medskip Another instance of \isa{semigroup} are the natural numbers:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{instantiation}\isamarkupfalse% +\ nat\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{primrec}\isamarkupfalse% +\ mult{\isacharunderscore}nat\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}Suc\ m\ {\isasymotimes}\ n\ {\isacharequal}\ Suc\ {\isacharparenleft}m\ {\isasymotimes}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ m\ n\ q\ {\isacharcolon}{\isacharcolon}\ nat\ \isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymotimes}\ q\ {\isacharequal}\ m\ {\isasymotimes}\ {\isacharparenleft}n\ {\isasymotimes}\ q{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ m{\isacharparenright}\ auto\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Note the occurence of the name \isa{mult{\isacharunderscore}nat} + in the primrec declaration; by default, the local name of + a class operation \isa{f} to instantiate on type constructor + \isa{{\isasymkappa}} are mangled as \isa{f{\isacharunderscore}{\isasymkappa}}. In case of uncertainty, + these names may be inspected using the \hyperlink{command.print-context}{\mbox{\isa{\isacommand{print{\isacharunderscore}context}}}} command + or the corresponding ProofGeneral button.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Lifting and parametric types% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Overloaded definitions giving on class instantiation + may include recursion over the syntactic structure of types. + As a canonical example, we model product semigroups + using our simple algebra:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{instantiation}\isamarkupfalse% +\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}semigroup{\isacharcomma}\ semigroup{\isacharparenright}\ semigroup\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ mult{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{2}}\ {\isacharequal}\ {\isacharparenleft}fst\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ fst\ p\isactrlisub {\isadigit{2}}{\isacharcomma}\ snd\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ snd\ p\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ p\isactrlisub {\isadigit{1}}\ p\isactrlisub {\isadigit{2}}\ p\isactrlisub {\isadigit{3}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}semigroup\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}semigroup{\isachardoublequoteclose}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{2}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{3}}\ {\isacharequal}\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ {\isacharparenleft}p\isactrlisub {\isadigit{2}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{3}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ assoc{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +\ \ \ \ \ \ \isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Associativity from product semigroups is + established using + the definition of \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} on products and the hypothetical + associativity of the type components; these hypotheses + are facts due to the \isa{semigroup} constraints imposed + on the type components by the \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} proposition. + Indeed, this pattern often occurs with parametric types + and type classes.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Subclassing% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We define a subclass \isa{monoidl} (a semigroup with a left-hand neutral) + by extending \isa{semigroup} + with one additional parameter \isa{neutral} together + with its property:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{class}\isamarkupfalse% +\ monoidl\ {\isacharequal}\ semigroup\ {\isacharplus}\isanewline +\ \ \isakeyword{fixes}\ neutral\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isachardoublequoteclose}\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\ \ \isakeyword{assumes}\ neutl{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Again, we prove some instances, by + providing suitable parameter definitions and proofs for the + additional specifications. Observe that instantiations + for types with the same arity may be simultaneous:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{instantiation}\isamarkupfalse% +\ nat\ \isakeyword{and}\ int\ {\isacharcolon}{\isacharcolon}\ monoidl\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ neutral{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ k\ {\isacharequal}\ k{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoidl{\isacharcomma}\ monoidl{\isacharparenright}\ monoidl\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ neutral{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isasymone}{\isacharcomma}\ {\isasymone}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ p\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}monoidl\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}monoidl{\isachardoublequoteclose}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}prod{\isacharunderscore}def\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ neutl{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Fully-fledged monoids are modelled by another subclass + which does not add new parameters but tightens the specification:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{class}\isamarkupfalse% +\ monoid\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline +\ \ \isakeyword{assumes}\ neutr{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ nat\ \isakeyword{and}\ int\ {\isacharcolon}{\isacharcolon}\ monoid\ \isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}n\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ n{\isacharparenright}\ simp{\isacharunderscore}all\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}k\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ k{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoid{\isacharcomma}\ monoid{\isacharparenright}\ monoid\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\ \isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ p\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}monoid\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}monoid{\isachardoublequoteclose}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}p\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}prod{\isacharunderscore}def\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ neutr{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent To finish our small algebra example, we add a \isa{group} class + with a corresponding instance:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{class}\isamarkupfalse% +\ group\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline +\ \ \isakeyword{fixes}\ inverse\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}{\isachardoublequoteopen}{\isacharparenleft}{\isacharunderscore}{\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline +\ \ \isakeyword{assumes}\ invl{\isacharcolon}\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ int\ {\isacharcolon}{\isacharcolon}\ group\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ inverse{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}i{\isasymdiv}\ {\isacharequal}\ {\isacharminus}\ {\isacharparenleft}i{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ i\ {\isacharcolon}{\isacharcolon}\ int\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharminus}i\ {\isacharplus}\ i\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}i{\isasymdiv}\ {\isasymotimes}\ i\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ mult{\isacharunderscore}int{\isacharunderscore}def\ neutral{\isacharunderscore}int{\isacharunderscore}def\ inverse{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isamarkupsection{Type classes as locales% +} +\isamarkuptrue% +% +\isamarkupsubsection{A look behind the scene% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The example above gives an impression how Isar type classes work + in practice. As stated in the introduction, classes also provide + a link to Isar's locale system. Indeed, the logical core of a class + is nothing else than a locale:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{class}\isamarkupfalse% +\ idem\ {\isacharequal}\isanewline +\ \ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{assumes}\ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent essentially introduces the locale% +\end{isamarkuptext}% +\isamarkuptrue% +\ % +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{locale}\isamarkupfalse% +\ idem\ {\isacharequal}\isanewline +\ \ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{assumes}\ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent together with corresponding constant(s):% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{consts}\isamarkupfalse% +\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The connection to the type system is done by means + of a primitive axclass% +\end{isamarkuptext}% +\isamarkuptrue% +\ % +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{axclass}\isamarkupfalse% +\ idem\ {\isacharless}\ type\isanewline +\ \ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}\ % +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\begin{isamarkuptext}% +\noindent together with a corresponding interpretation:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{interpretation}\isamarkupfalse% +\ idem{\isacharunderscore}class{\isacharcolon}\isanewline +\ \ idem\ {\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isasymalpha}{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline +\isacommand{proof}\isamarkupfalse% +\ \isacommand{qed}\isamarkupfalse% +\ {\isacharparenleft}rule\ idem{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This gives you at hand the full power of the Isabelle module system; + conclusions in locale \isa{idem} are implicitly propagated + to class \isa{idem}.% +\end{isamarkuptext}% +\isamarkuptrue% +\ % +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isamarkupsubsection{Abstract reasoning% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle locales enable reasoning at a general level, while results + are implicitly transferred to all instances. For example, we can + now establish the \isa{left{\isacharunderscore}cancel} lemma for groups, which + states that the function \isa{{\isacharparenleft}x\ {\isasymotimes}{\isacharparenright}} is injective:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ left{\isacharunderscore}cancel{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z{\isachardoublequoteclose}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isacharequal}\ x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymdiv}\ {\isasymotimes}\ x{\isacharparenright}\ {\isasymotimes}\ y\ {\isacharequal}\ {\isacharparenleft}x{\isasymdiv}\ {\isasymotimes}\ x{\isacharparenright}\ {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% +\ assoc\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% +\ neutl\ \isakeyword{and}\ invl\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Here the \qt{\hyperlink{keyword.in}{\mbox{\isa{\isakeyword{in}}}} \isa{group}} target specification + indicates that the result is recorded within that context for later + use. This local theorem is also lifted to the global one \hyperlink{fact.group.left-cancel:}{\mbox{\isa{group{\isachardot}left{\isacharunderscore}cancel{\isacharcolon}}}} \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}group{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}. Since type \isa{int} has been made an instance of + \isa{group} before, we may refer to that fact as well: \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ int{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Derived definitions% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle locales support a concept of local definitions + in locales:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{primrec}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ monoid{\isacharparenright}\ pow{\isacharunderscore}nat\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent If the locale \isa{group} is also a class, this local + definition is propagated onto a global definition of + \isa{{\isachardoublequote}pow{\isacharunderscore}nat\ {\isasymColon}\ nat\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid{\isachardoublequote}} + with corresponding theorems + + \isa{pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}\isasep\isanewline% +pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x}. + + \noindent As you can see from this example, for local + definitions you may use any specification tool + which works together with locales (e.g. \cite{krauss2006}).% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{A functor analogy% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We introduced Isar classes by analogy to type classes + functional programming; if we reconsider this in the + context of what has been said about type classes and locales, + we can drive this analogy further by stating that type + classes essentially correspond to functors which have + a canonical interpretation as type classes. + Anyway, there is also the possibility of other interpretations. + For example, also \isa{list}s form a monoid with + \isa{append} and \isa{{\isacharbrackleft}{\isacharbrackright}} as operations, but it + seems inappropriate to apply to lists + the same operations as for genuinely algebraic types. + In such a case, we simply can do a particular interpretation + of monoids for lists:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{interpretation}\isamarkupfalse% +\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ \isacommand{qed}\isamarkupfalse% +\ auto% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This enables us to apply facts on monoids + to lists, e.g. \isa{{\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ x\ {\isacharequal}\ x}. + + When using this interpretation pattern, it may also + be appropriate to map derived definitions accordingly:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{primrec}\isamarkupfalse% +\ replicate\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isasymalpha}\ list\ {\isasymRightarrow}\ {\isasymalpha}\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}replicate\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}replicate\ {\isacharparenleft}Suc\ n{\isacharparenright}\ xs\ {\isacharequal}\ xs\ {\isacharat}\ replicate\ n\ xs{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{interpretation}\isamarkupfalse% +\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{interpret}\isamarkupfalse% +\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ n\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ n\ {\isacharequal}\ replicate\ n{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ n{\isacharparenright}\ auto\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +\ intro{\isacharunderscore}locales% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isamarkupsubsection{Additional subclass relations% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Any \isa{group} is also a \isa{monoid}; this + can be made explicit by claiming an additional + subclass relation, + together with a proof of the logical difference:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{subclass}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ monoid\isanewline +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ x\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ invl\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{with}\isamarkupfalse% +\ assoc\ {\isacharbrackleft}symmetric{\isacharbrackright}\ neutl\ invl\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ {\isasymone}{\isacharparenright}\ {\isacharequal}\ x{\isasymdiv}\ {\isasymotimes}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{with}\isamarkupfalse% +\ left{\isacharunderscore}cancel\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +The logical proof is carried out on the locale level. + Afterwards it is propagated + to the type system, making \isa{group} an instance of + \isa{monoid} by adding an additional edge + to the graph of subclass relations + (cf.\ \figref{fig:subclass}). + + \begin{figure}[htbp] + \begin{center} + \small + \unitlength 0.6mm + \begin{picture}(40,60)(0,0) + \put(20,60){\makebox(0,0){\isa{semigroup}}} + \put(20,40){\makebox(0,0){\isa{monoidl}}} + \put(00,20){\makebox(0,0){\isa{monoid}}} + \put(40,00){\makebox(0,0){\isa{group}}} + \put(20,55){\vector(0,-1){10}} + \put(15,35){\vector(-1,-1){10}} + \put(25,35){\vector(1,-3){10}} + \end{picture} + \hspace{8em} + \begin{picture}(40,60)(0,0) + \put(20,60){\makebox(0,0){\isa{semigroup}}} + \put(20,40){\makebox(0,0){\isa{monoidl}}} + \put(00,20){\makebox(0,0){\isa{monoid}}} + \put(40,00){\makebox(0,0){\isa{group}}} + \put(20,55){\vector(0,-1){10}} + \put(15,35){\vector(-1,-1){10}} + \put(05,15){\vector(3,-1){30}} + \end{picture} + \caption{Subclass relationship of monoids and groups: + before and after establishing the relationship + \isa{group\ {\isasymsubseteq}\ monoid}; transitive edges are left out.} + \label{fig:subclass} + \end{center} + \end{figure} + + For illustration, a derived definition + in \isa{group} which uses \isa{pow{\isacharunderscore}nat}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{definition}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ pow{\isacharunderscore}int\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}int\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ k\ {\isachargreater}{\isacharequal}\ {\isadigit{0}}\isanewline +\ \ \ \ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\isanewline +\ \ \ \ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent yields the global definition of + \isa{{\isachardoublequote}pow{\isacharunderscore}int\ {\isasymColon}\ int\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group{\isachardoublequote}} + with the corresponding theorem \isa{pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{0}}\ {\isasymle}\ k\ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{A note on syntax% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +As a commodity, class context syntax allows to refer + to local class operations and their global counterparts + uniformly; type inference resolves ambiguities. For example:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{context}\isamarkupfalse% +\ semigroup\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{term}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % +\isamarkupcmt{example 1% +} +\isanewline +\isacommand{term}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymColon}nat{\isacharparenright}\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % +\isamarkupcmt{example 2% +} +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{term}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % +\isamarkupcmt{example 3% +} +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Here in example 1, the term refers to the local class operation + \isa{mult\ {\isacharbrackleft}{\isasymalpha}{\isacharbrackright}}, whereas in example 2 the type constraint + enforces the global class operation \isa{mult\ {\isacharbrackleft}nat{\isacharbrackright}}. + In the global context in example 3, the reference is + to the polymorphic global class operation \isa{mult\ {\isacharbrackleft}{\isacharquery}{\isasymalpha}\ {\isasymColon}\ semigroup{\isacharbrackright}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Further issues% +} +\isamarkuptrue% +% +\isamarkupsubsection{Type classes and code generation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Turning back to the first motivation for type classes, + namely overloading, it is obvious that overloading + stemming from \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} statements and + \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} + targets naturally maps to Haskell type classes. + The code generator framework \cite{isabelle-codegen} + takes this into account. Concerning target languages + lacking type classes (e.g.~SML), type classes + are implemented by explicit dictionary construction. + As example, let's go back to the power function:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{definition}\isamarkupfalse% +\ example\ {\isacharcolon}{\isacharcolon}\ int\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}example\ {\isacharequal}\ pow{\isacharunderscore}int\ {\isadigit{1}}{\isadigit{0}}\ {\isacharparenleft}{\isacharminus}{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This maps to Haskell as:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}module Example where {\char123}\\ +\hspace*{0pt}\\ +\hspace*{0pt}\\ +\hspace*{0pt}data Nat = Zero{\char95}nat | Suc Nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}nat{\char95}aux ::~Integer -> Nat -> Nat;\\ +\hspace*{0pt}nat{\char95}aux i n = (if i <= 0 then n else nat{\char95}aux (i - 1) (Suc n));\\ +\hspace*{0pt}\\ +\hspace*{0pt}nat ::~Integer -> Nat;\\ +\hspace*{0pt}nat i = nat{\char95}aux i Zero{\char95}nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}class Semigroup a where {\char123}\\ +\hspace*{0pt} ~mult ::~a -> a -> a;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}class (Semigroup a) => Monoidl a where {\char123}\\ +\hspace*{0pt} ~neutral ::~a;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}class (Monoidl a) => Monoid a where {\char123}\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}class (Monoid a) => Group a where {\char123}\\ +\hspace*{0pt} ~inverse ::~a -> a;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}inverse{\char95}int ::~Integer -> Integer;\\ +\hspace*{0pt}inverse{\char95}int i = negate i;\\ +\hspace*{0pt}\\ +\hspace*{0pt}neutral{\char95}int ::~Integer;\\ +\hspace*{0pt}neutral{\char95}int = 0;\\ +\hspace*{0pt}\\ +\hspace*{0pt}mult{\char95}int ::~Integer -> Integer -> Integer;\\ +\hspace*{0pt}mult{\char95}int i j = i + j;\\ +\hspace*{0pt}\\ +\hspace*{0pt}instance Semigroup Integer where {\char123}\\ +\hspace*{0pt} ~mult = mult{\char95}int;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}instance Monoidl Integer where {\char123}\\ +\hspace*{0pt} ~neutral = neutral{\char95}int;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}instance Monoid Integer where {\char123}\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}instance Group Integer where {\char123}\\ +\hspace*{0pt} ~inverse = inverse{\char95}int;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}pow{\char95}nat ::~forall a.~(Monoid a) => Nat -> a -> a;\\ +\hspace*{0pt}pow{\char95}nat Zero{\char95}nat x = neutral;\\ +\hspace*{0pt}pow{\char95}nat (Suc n) x = mult x (pow{\char95}nat n x);\\ +\hspace*{0pt}\\ +\hspace*{0pt}pow{\char95}int ::~forall a.~(Group a) => Integer -> a -> a;\\ +\hspace*{0pt}pow{\char95}int k x =\\ +\hspace*{0pt} ~(if 0 <= k then pow{\char95}nat (nat k) x\\ +\hspace*{0pt} ~~~else inverse (pow{\char95}nat (nat (negate k)) x));\\ +\hspace*{0pt}\\ +\hspace*{0pt}example ::~Integer;\\ +\hspace*{0pt}example = pow{\char95}int 10 (-2);\\ +\hspace*{0pt}\\ +\hspace*{0pt}{\char125}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The whole code in SML with explicit dictionary passing:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun nat{\char95}aux i n =\\ +\hspace*{0pt} ~(if IntInf.<= (i,~(0 :~IntInf.int)) then n\\ +\hspace*{0pt} ~~~else nat{\char95}aux (IntInf.- (i,~(1 :~IntInf.int))) (Suc n));\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun nat i = nat{\char95}aux i Zero{\char95}nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a semigroup = {\char123}mult :~'a -> 'a -> 'a{\char125};\\ +\hspace*{0pt}fun mult (A{\char95}:'a semigroup) = {\char35}mult A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a monoidl =\\ +\hspace*{0pt} ~{\char123}Classes{\char95}{\char95}semigroup{\char95}monoidl :~'a semigroup,~neutral :~'a{\char125};\\ +\hspace*{0pt}fun semigroup{\char95}monoidl (A{\char95}:'a monoidl) = {\char35}Classes{\char95}{\char95}semigroup{\char95}monoidl A{\char95};\\ +\hspace*{0pt}fun neutral (A{\char95}:'a monoidl) = {\char35}neutral A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a monoid = {\char123}Classes{\char95}{\char95}monoidl{\char95}monoid :~'a monoidl{\char125};\\ +\hspace*{0pt}fun monoidl{\char95}monoid (A{\char95}:'a monoid) = {\char35}Classes{\char95}{\char95}monoidl{\char95}monoid A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a group = {\char123}Classes{\char95}{\char95}monoid{\char95}group :~'a monoid,~inverse :~'a -> 'a{\char125};\\ +\hspace*{0pt}fun monoid{\char95}group (A{\char95}:'a group) = {\char35}Classes{\char95}{\char95}monoid{\char95}group A{\char95};\\ +\hspace*{0pt}fun inverse (A{\char95}:'a group) = {\char35}inverse A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun inverse{\char95}int i = IntInf.{\char126}~i;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val neutral{\char95}int :~IntInf.int = (0 :~IntInf.int)\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun mult{\char95}int i j = IntInf.+ (i,~j);\\ +\hspace*{0pt}\\ +\hspace*{0pt}val semigroup{\char95}int = {\char123}mult = mult{\char95}int{\char125}~:~IntInf.int semigroup;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val monoidl{\char95}int =\\ +\hspace*{0pt} ~{\char123}Classes{\char95}{\char95}semigroup{\char95}monoidl = semigroup{\char95}int,~neutral = neutral{\char95}int{\char125}~:\\ +\hspace*{0pt} ~IntInf.int monoidl;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val monoid{\char95}int = {\char123}Classes{\char95}{\char95}monoidl{\char95}monoid = monoidl{\char95}int{\char125}~:\\ +\hspace*{0pt} ~IntInf.int monoid;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val group{\char95}int =\\ +\hspace*{0pt} ~{\char123}Classes{\char95}{\char95}monoid{\char95}group = monoid{\char95}int,~inverse = inverse{\char95}int{\char125}~:\\ +\hspace*{0pt} ~IntInf.int group;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun pow{\char95}nat A{\char95}~Zero{\char95}nat x = neutral (monoidl{\char95}monoid A{\char95})\\ +\hspace*{0pt} ~| pow{\char95}nat A{\char95}~(Suc n) x =\\ +\hspace*{0pt} ~~~mult ((semigroup{\char95}monoidl o monoidl{\char95}monoid) A{\char95}) x (pow{\char95}nat A{\char95}~n x);\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun pow{\char95}int A{\char95}~k x =\\ +\hspace*{0pt} ~(if IntInf.<= ((0 :~IntInf.int),~k)\\ +\hspace*{0pt} ~~~then pow{\char95}nat (monoid{\char95}group A{\char95}) (nat k) x\\ +\hspace*{0pt} ~~~else inverse A{\char95}~(pow{\char95}nat (monoid{\char95}group A{\char95}) (nat (IntInf.{\char126}~k)) x));\\ +\hspace*{0pt}\\ +\hspace*{0pt}val example :~IntInf.int =\\ +\hspace*{0pt} ~pow{\char95}int group{\char95}int (10 :~IntInf.int) ({\char126}2 :~IntInf.int)\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isamarkupsubsection{Inspecting the type class universe% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +To facilitate orientation in complex subclass structures, + two diagnostics commands are provided: + + \begin{description} + + \item[\hyperlink{command.print-classes}{\mbox{\isa{\isacommand{print{\isacharunderscore}classes}}}}] print a list of all classes + together with associated operations etc. + + \item[\hyperlink{command.class-deps}{\mbox{\isa{\isacommand{class{\isacharunderscore}deps}}}}] visualizes the subclass relation + between all classes as a Hasse diagram. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Classes/classes.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/classes.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,50 @@ + +\documentclass[12pt,a4paper,fleqn]{article} +\usepackage{latexsym,graphicx} +\usepackage[refpage]{nomencl} +\usepackage{../iman,../extra,../isar,../proof} +\usepackage{../isabelle,../isabellesym} +\usepackage{style} +\usepackage{../pdfsetup} + + +\hyphenation{Isabelle} +\hyphenation{Isar} +\isadroptag{theory} + +\title{\includegraphics[scale=0.5]{isabelle_isar} + \\[4ex] Haskell-style type classes with Isabelle/Isar} +\author{\emph{Florian Haftmann}} + +\begin{document} + +\maketitle + +\begin{abstract} + \noindent This tutorial introduces the look-and-feel of Isar type classes + to the end-user; Isar type classes are a convenient mechanism + for organizing specifications, overcoming some drawbacks + of raw axiomatic type classes. Essentially, they combine + an operational aspect (in the manner of Haskell) with + a logical aspect, both managed uniformly. +\end{abstract} + +\thispagestyle{empty}\clearpage + +\pagenumbering{roman} +\clearfirst + +\input{Thy/document/Classes.tex} + +\begingroup +\bibliographystyle{plain} \small\raggedright\frenchspacing +\bibliography{../manual} +\endgroup + +\end{document} + + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Classes/style.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Classes/style.sty Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,48 @@ + +%% toc +\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} +\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} + +%% paragraphs +\setlength{\parindent}{1em} + +%% references +\newcommand{\secref}[1]{\S\ref{#1}} +\newcommand{\figref}[1]{figure~\ref{#1}} + +%% logical markup +\newcommand{\strong}[1]{{\bfseries {#1}}} +\newcommand{\qn}[1]{\emph{#1}} + +%% typographic conventions +\newcommand{\qt}[1]{``{#1}''} + +%% verbatim text +\newcommand{\isatypewriter}{\fontsize{9pt}{0pt}\tt\renewcommand{\baselinestretch}{1}\setlength{\baselineskip}{9pt}} + +%% quote environment +\isakeeptag{quote} +\renewenvironment{quote} + {\list{}{\leftmargin2em\rightmargin0pt}\parindent0pt\parskip0pt\item\relax} + {\endlist} +\renewcommand{\isatagquote}{\begin{quote}} +\renewcommand{\endisatagquote}{\end{quote}} +\newcommand{\quotebreak}{\\[1.2ex]} + +%% presentation +\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} + +%% character detail +\renewcommand{\isadigit}[1]{\isamath{#1}} +\binperiod +\underscoreoff + +%% format +\pagestyle{headings} +\isabellestyle{it} + + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "implementation" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/IsaMakefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/IsaMakefile Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,33 @@ + +## targets + +default: Thy +images: +test: Thy + +all: images test + + +## global settings + +SRC = $(ISABELLE_HOME)/src +OUT = $(ISABELLE_OUTPUT) +LOG = $(OUT)/log + +USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -d false -C false -D document + + +## Thy + +THY = $(LOG)/HOL-Thy.gz + +Thy: $(THY) + +$(THY): Thy/ROOT.ML Thy/*.thy ../antiquote_setup.ML ../more_antiquote.ML + @$(USEDIR) HOL Thy + + +## clean + +clean: + @rm -f $(THY) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Makefile Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,35 @@ + +## targets + +default: dvi + + +## dependencies + +include ../Makefile.in + +NAME = codegen + +FILES = $(NAME).tex Thy/document/*.tex \ + style.sty ../iman.sty ../extra.sty ../isar.sty \ + ../isabelle.sty ../isabellesym.sty ../pdfsetup.sty \ + ../manual.bib ../proof.sty + +dvi: $(NAME).dvi + +$(NAME).dvi: $(FILES) isabelle_isar.eps codegen_process.ps + $(LATEX) $(NAME) + $(BIBTEX) $(NAME) + $(LATEX) $(NAME) + $(LATEX) $(NAME) + +pdf: $(NAME).pdf + +$(NAME).pdf: $(FILES) isabelle_isar.pdf codegen_process.pdf + $(PDFLATEX) $(NAME) + $(BIBTEX) $(NAME) + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) + $(FIXBOOKMARKS) $(NAME).out + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/Adaption.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/Adaption.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,360 @@ +theory Adaption +imports Setup +begin + +setup %invisible {* Code_Target.extend_target ("\", ("SML", K I)) *} + +section {* Adaption to target languages \label{sec:adaption} *} + +subsection {* Adapting code generation *} + +text {* + The aspects of code generation introduced so far have two aspects + in common: + + \begin{itemize} + \item They act uniformly, without reference to a specific + target language. + \item They are \emph{safe} in the sense that as long as you trust + the code generator meta theory and implementation, you cannot + produce programs that yield results which are not derivable + in the logic. + \end{itemize} + + \noindent In this section we will introduce means to \emph{adapt} the serialiser + to a specific target language, i.e.~to print program fragments + in a way which accommodates \qt{already existing} ingredients of + a target language environment, for three reasons: + + \begin{itemize} + \item improving readability and aesthetics of generated code + \item gaining efficiency + \item interface with language parts which have no direct counterpart + in @{text "HOL"} (say, imperative data structures) + \end{itemize} + + \noindent Generally, you should avoid using those features yourself + \emph{at any cost}: + + \begin{itemize} + \item The safe configuration methods act uniformly on every target language, + whereas for adaption you have to treat each target language separate. + \item Application is extremely tedious since there is no abstraction + which would allow for a static check, making it easy to produce garbage. + \item More or less subtle errors can be introduced unconsciously. + \end{itemize} + + \noindent However, even if you ought refrain from setting up adaption + yourself, already the @{text "HOL"} comes with some reasonable default + adaptions (say, using target language list syntax). There also some + common adaption cases which you can setup by importing particular + library theories. In order to understand these, we provide some clues here; + these however are not supposed to replace a careful study of the sources. +*} + +subsection {* The adaption principle *} + +text {* + The following figure illustrates what \qt{adaption} is conceptually + supposed to be: + + \begin{figure}[here] + \begin{tikzpicture}[scale = 0.5] + \tikzstyle water=[color = blue, thick] + \tikzstyle ice=[color = black, very thick, cap = round, join = round, fill = white] + \tikzstyle process=[color = green, semithick, ->] + \tikzstyle adaption=[color = red, semithick, ->] + \tikzstyle target=[color = black] + \foreach \x in {0, ..., 24} + \draw[style=water] (\x, 0.25) sin + (0.25, 0.25) cos + (0.25, -0.25) sin + + (0.25, -0.25) cos + (0.25, 0.25); + \draw[style=ice] (1, 0) -- + (3, 6) node[above, fill=white] {logic} -- (5, 0) -- cycle; + \draw[style=ice] (9, 0) -- + (11, 6) node[above, fill=white] {intermediate language} -- (13, 0) -- cycle; + \draw[style=ice] (15, -6) -- + (19, 6) node[above, fill=white] {target language} -- (23, -6) -- cycle; + \draw[style=process] + (3.5, 3) .. controls (7, 5) .. node[fill=white] {translation} (10.5, 3); + \draw[style=process] + (11.5, 3) .. controls (15, 5) .. node[fill=white] (serialisation) {serialisation} (18.5, 3); + \node (adaption) at (11, -2) [style=adaption] {adaption}; + \node at (19, 3) [rotate=90] {generated}; + \node at (19.5, -5) {language}; + \node at (19.5, -3) {library}; + \node (includes) at (19.5, -1) {includes}; + \node (reserved) at (16.5, -3) [rotate=72] {reserved}; % proper 71.57 + \draw[style=process] + (includes) -- (serialisation); + \draw[style=process] + (reserved) -- (serialisation); + \draw[style=adaption] + (adaption) -- (serialisation); + \draw[style=adaption] + (adaption) -- (includes); + \draw[style=adaption] + (adaption) -- (reserved); + \end{tikzpicture} + \caption{The adaption principle} + \label{fig:adaption} + \end{figure} + + \noindent In the tame view, code generation acts as broker between + @{text logic}, @{text "intermediate language"} and + @{text "target language"} by means of @{text translation} and + @{text serialisation}; for the latter, the serialiser has to observe + the structure of the @{text language} itself plus some @{text reserved} + keywords which have to be avoided for generated code. + However, if you consider @{text adaption} mechanisms, the code generated + by the serializer is just the tip of the iceberg: + + \begin{itemize} + \item @{text serialisation} can be \emph{parametrised} such that + logical entities are mapped to target-specific ones + (e.g. target-specific list syntax, + see also \secref{sec:adaption_mechanisms}) + \item Such parametrisations can involve references to a + target-specific standard @{text library} (e.g. using + the @{text Haskell} @{verbatim Maybe} type instead + of the @{text HOL} @{type "option"} type); + if such are used, the corresponding identifiers + (in our example, @{verbatim Maybe}, @{verbatim Nothing} + and @{verbatim Just}) also have to be considered @{text reserved}. + \item Even more, the user can enrich the library of the + target-language by providing code snippets + (\qt{@{text "includes"}}) which are prepended to + any generated code (see \secref{sec:include}); this typically + also involves further @{text reserved} identifiers. + \end{itemize} + + \noindent As figure \ref{fig:adaption} illustrates, all these adaption mechanisms + have to act consistently; it is at the discretion of the user + to take care for this. +*} + +subsection {* Common adaption patterns *} + +text {* + The @{theory HOL} @{theory Main} theory already provides a code + generator setup + which should be suitable for most applications. Common extensions + and modifications are available by certain theories of the @{text HOL} + library; beside being useful in applications, they may serve + as a tutorial for customising the code generator setup (see below + \secref{sec:adaption_mechanisms}). + + \begin{description} + + \item[@{theory "Code_Integer"}] represents @{text HOL} integers by big + integer literals in target languages. + \item[@{theory "Code_Char"}] represents @{text HOL} characters by + character literals in target languages. + \item[@{theory "Code_Char_chr"}] like @{text "Code_Char"}, + but also offers treatment of character codes; includes + @{theory "Code_Char"}. + \item[@{theory "Efficient_Nat"}] \label{eff_nat} implements natural numbers by integers, + which in general will result in higher efficiency; pattern + matching with @{term "0\nat"} / @{const "Suc"} + is eliminated; includes @{theory "Code_Integer"} + and @{theory "Code_Index"}. + \item[@{theory "Code_Index"}] provides an additional datatype + @{typ index} which is mapped to target-language built-in integers. + Useful for code setups which involve e.g. indexing of + target-language arrays. + \item[@{theory "Code_Message"}] provides an additional datatype + @{typ message_string} which is isomorphic to strings; + @{typ message_string}s are mapped to target-language strings. + Useful for code setups which involve e.g. printing (error) messages. + + \end{description} + + \begin{warn} + When importing any of these theories, they should form the last + items in an import list. Since these theories adapt the + code generator setup in a non-conservative fashion, + strange effects may occur otherwise. + \end{warn} +*} + + +subsection {* Parametrising serialisation \label{sec:adaption_mechanisms} *} + +text {* + Consider the following function and its corresponding + SML code: +*} + +primrec %quote in_interval :: "nat \ nat \ nat \ bool" where + "in_interval (k, l) n \ k \ n \ n \ l" +(*<*) +code_type %invisible bool + (SML) +code_const %invisible True and False and "op \" and Not + (SML and and and) +(*>*) +text %quote {*@{code_stmts in_interval (SML)}*} + +text {* + \noindent Though this is correct code, it is a little bit unsatisfactory: + boolean values and operators are materialised as distinguished + entities with have nothing to do with the SML-built-in notion + of \qt{bool}. This results in less readable code; + additionally, eager evaluation may cause programs to + loop or break which would perfectly terminate when + the existing SML @{verbatim "bool"} would be used. To map + the HOL @{typ bool} on SML @{verbatim "bool"}, we may use + \qn{custom serialisations}: +*} + +code_type %quotett bool + (SML "bool") +code_const %quotett True and False and "op \" + (SML "true" and "false" and "_ andalso _") + +text {* + \noindent The @{command code_type} command takes a type constructor + as arguments together with a list of custom serialisations. + Each custom serialisation starts with a target language + identifier followed by an expression, which during + code serialisation is inserted whenever the type constructor + would occur. For constants, @{command code_const} implements + the corresponding mechanism. Each ``@{verbatim "_"}'' in + a serialisation expression is treated as a placeholder + for the type constructor's (the constant's) arguments. +*} + +text %quote {*@{code_stmts in_interval (SML)}*} + +text {* + \noindent This still is not perfect: the parentheses + around the \qt{andalso} expression are superfluous. + Though the serialiser + by no means attempts to imitate the rich Isabelle syntax + framework, it provides some common idioms, notably + associative infixes with precedences which may be used here: +*} + +code_const %quotett "op \" + (SML infixl 1 "andalso") + +text %quote {*@{code_stmts in_interval (SML)}*} + +text {* + \noindent The attentive reader may ask how we assert that no generated + code will accidentally overwrite. For this reason the serialiser has + an internal table of identifiers which have to be avoided to be used + for new declarations. Initially, this table typically contains the + keywords of the target language. It can be extended manually, thus avoiding + accidental overwrites, using the @{command "code_reserved"} command: +*} + +code_reserved %quote "\" bool true false andalso + +text {* + \noindent Next, we try to map HOL pairs to SML pairs, using the + infix ``@{verbatim "*"}'' type constructor and parentheses: +*} +(*<*) +code_type %invisible * + (SML) +code_const %invisible Pair + (SML) +(*>*) +code_type %quotett * + (SML infix 2 "*") +code_const %quotett Pair + (SML "!((_),/ (_))") + +text {* + \noindent The initial bang ``@{verbatim "!"}'' tells the serialiser + never to put + parentheses around the whole expression (they are already present), + while the parentheses around argument place holders + tell not to put parentheses around the arguments. + The slash ``@{verbatim "/"}'' (followed by arbitrary white space) + inserts a space which may be used as a break if necessary + during pretty printing. + + These examples give a glimpse what mechanisms + custom serialisations provide; however their usage + requires careful thinking in order not to introduce + inconsistencies -- or, in other words: + custom serialisations are completely axiomatic. + + A further noteworthy details is that any special + character in a custom serialisation may be quoted + using ``@{verbatim "'"}''; thus, in + ``@{verbatim "fn '_ => _"}'' the first + ``@{verbatim "_"}'' is a proper underscore while the + second ``@{verbatim "_"}'' is a placeholder. +*} + + +subsection {* @{text Haskell} serialisation *} + +text {* + For convenience, the default + @{text HOL} setup for @{text Haskell} maps the @{class eq} class to + its counterpart in @{text Haskell}, giving custom serialisations + for the class @{class eq} (by command @{command code_class}) and its operation + @{const HOL.eq} +*} + +code_class %quotett eq + (Haskell "Eq") + +code_const %quotett "op =" + (Haskell infixl 4 "==") + +text {* + \noindent A problem now occurs whenever a type which + is an instance of @{class eq} in @{text HOL} is mapped + on a @{text Haskell}-built-in type which is also an instance + of @{text Haskell} @{text Eq}: +*} + +typedecl %quote bar + +instantiation %quote bar :: eq +begin + +definition %quote "eq_class.eq (x\bar) y \ x = y" + +instance %quote by default (simp add: eq_bar_def) + +end %quote +code_type %quotett bar + (Haskell "Integer") + +text {* + \noindent The code generator would produce + an additional instance, which of course is rejected by the @{text Haskell} + compiler. + To suppress this additional instance, use + @{text "code_instance"}: +*} + +code_instance %quotett bar :: eq + (Haskell -) + + +subsection {* Enhancing the target language context \label{sec:include} *} + +text {* + In rare cases it is necessary to \emph{enrich} the context of a + target language; this is accomplished using the @{command "code_include"} + command: +*} + +code_include %quotett Haskell "Errno" +{*errno i = error ("Error number: " ++ show i)*} + +code_reserved %quotett Haskell Errno + +text {* + \noindent Such named @{text include}s are then prepended to every generated code. + Inspect such code in order to find out how @{command "code_include"} behaves + with respect to a particular target language. +*} + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/Further.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/Further.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,112 @@ +theory Further +imports Setup +begin + +section {* Further issues \label{sec:further} *} + +subsection {* Further reading *} + +text {* + Do dive deeper into the issue of code generation, you should visit + the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} which + contains exhaustive syntax diagrams. +*} + +subsection {* Modules *} + +text {* + When invoking the @{command export_code} command it is possible to leave + out the @{keyword "module_name"} part; then code is distributed over + different modules, where the module name space roughly is induced + by the @{text Isabelle} theory name space. + + Then sometimes the awkward situation occurs that dependencies between + definitions introduce cyclic dependencies between modules, which in the + @{text Haskell} world leaves you to the mercy of the @{text Haskell} implementation + you are using, while for @{text SML}/@{text OCaml} code generation is not possible. + + A solution is to declare module names explicitly. + Let use assume the three cyclically dependent + modules are named \emph{A}, \emph{B} and \emph{C}. + Then, by stating +*} + +code_modulename %quote SML + A ABC + B ABC + C ABC + +text {* + we explicitly map all those modules on \emph{ABC}, + resulting in an ad-hoc merge of this three modules + at serialisation time. +*} + +subsection {* Evaluation oracle *} + +text {* + Code generation may also be used to \emph{evaluate} expressions + (using @{text SML} as target language of course). + For instance, the @{command value} allows to reduce an expression to a + normal form with respect to the underlying code equations: +*} + +value %quote "42 / (12 :: rat)" + +text {* + \noindent will display @{term "7 / (2 :: rat)"}. + + The @{method eval} method tries to reduce a goal by code generation to @{term True} + and solves it in that case, but fails otherwise: +*} + +lemma %quote "42 / (12 :: rat) = 7 / 2" + by %quote eval + +text {* + \noindent The soundness of the @{method eval} method depends crucially + on the correctness of the code generator; this is one of the reasons + why you should not use adaption (see \secref{sec:adaption}) frivolously. +*} + +subsection {* Code antiquotation *} + +text {* + In scenarios involving techniques like reflection it is quite common + that code generated from a theory forms the basis for implementing + a proof procedure in @{text SML}. To facilitate interfacing of generated code + with system code, the code generator provides a @{text code} antiquotation: +*} + +datatype %quote form = T | F | And form form | Or form form +ML %quotett {* + fun eval_form @{code T} = true + | eval_form @{code F} = false + | eval_form (@{code And} (p, q)) = + eval_form p andalso eval_form q + | eval_form (@{code Or} (p, q)) = + eval_form p orelse eval_form q; +*} + +text {* + \noindent @{text code} takes as argument the name of a constant; after the + whole @{text SML} is read, the necessary code is generated transparently + and the corresponding constant names are inserted. This technique also + allows to use pattern matching on constructors stemming from compiled + @{text datatypes}. + + For a less simplistic example, theory @{theory Ferrack} is + a good reference. +*} + +subsection {* Imperative data structures *} + +text {* + If you consider imperative data structures as inevitable for a specific + application, you should consider + \emph{Imperative Functional Programming with Isabelle/HOL} + (\cite{bulwahn-et-al:2008:imperative}); + the framework described there is available in theory @{theory Imperative_HOL}. +*} + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/Introduction.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/Introduction.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,204 @@ +theory Introduction +imports Setup +begin + +section {* Introduction and Overview *} + +text {* + This tutorial introduces a generic code generator for the + @{text Isabelle} system. + Generic in the sense that the + \qn{target language} for which code shall ultimately be + generated is not fixed but may be an arbitrary state-of-the-art + functional programming language (currently, the implementation + supports @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml} and @{text Haskell} + \cite{haskell-revised-report}). + + Conceptually the code generator framework is part + of Isabelle's @{theory Pure} meta logic framework; the logic + @{theory HOL} which is an extension of @{theory Pure} + already comes with a reasonable framework setup and thus provides + a good working horse for raising code-generation-driven + applications. So, we assume some familiarity and experience + with the ingredients of the @{theory HOL} distribution theories. + (see also \cite{isa-tutorial}). + + The code generator aims to be usable with no further ado + in most cases while allowing for detailed customisation. + This manifests in the structure of this tutorial: after a short + conceptual introduction with an example (\secref{sec:intro}), + we discuss the generic customisation facilities (\secref{sec:program}). + A further section (\secref{sec:adaption}) is dedicated to the matter of + \qn{adaption} to specific target language environments. After some + further issues (\secref{sec:further}) we conclude with an overview + of some ML programming interfaces (\secref{sec:ml}). + + \begin{warn} + Ultimately, the code generator which this tutorial deals with + is supposed to replace the existing code generator + by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. + So, for the moment, there are two distinct code generators + in Isabelle. In case of ambiguity, we will refer to the framework + described here as @{text "generic code generator"}, to the + other as @{text "SML code generator"}. + Also note that while the framework itself is + object-logic independent, only @{theory HOL} provides a reasonable + framework setup. + \end{warn} + +*} + +subsection {* Code generation via shallow embedding \label{sec:intro} *} + +text {* + The key concept for understanding @{text Isabelle}'s code generation is + \emph{shallow embedding}, i.e.~logical entities like constants, types and + classes are identified with corresponding concepts in the target language. + + Inside @{theory HOL}, the @{command datatype} and + @{command definition}/@{command primrec}/@{command fun} declarations form + the core of a functional programming language. The default code generator setup + allows to turn those into functional programs immediately. + This means that \qt{naive} code generation can proceed without further ado. + For example, here a simple \qt{implementation} of amortised queues: +*} + +datatype %quote 'a queue = AQueue "'a list" "'a list" + +definition %quote empty :: "'a queue" where + "empty = AQueue [] []" + +primrec %quote enqueue :: "'a \ 'a queue \ 'a queue" where + "enqueue x (AQueue xs ys) = AQueue (x # xs) ys" + +fun %quote dequeue :: "'a queue \ 'a option \ 'a queue" where + "dequeue (AQueue [] []) = (None, AQueue [] [])" + | "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)" + | "dequeue (AQueue xs []) = + (case rev xs of y # ys \ (Some y, AQueue [] ys))" + +text {* \noindent Then we can generate code e.g.~for @{text SML} as follows: *} + +export_code %quote empty dequeue enqueue in SML + module_name Example file "examples/example.ML" + +text {* \noindent resulting in the following code: *} + +text %quote {*@{code_stmts empty enqueue dequeue (SML)}*} + +text {* + \noindent The @{command export_code} command takes a space-separated list of + constants for which code shall be generated; anything else needed for those + is added implicitly. Then follows a target language identifier + (@{text SML}, @{text OCaml} or @{text Haskell}) and a freely chosen module name. + A file name denotes the destination to store the generated code. Note that + the semantics of the destination depends on the target language: for + @{text SML} and @{text OCaml} it denotes a \emph{file}, for @{text Haskell} + it denotes a \emph{directory} where a file named as the module name + (with extension @{text ".hs"}) is written: +*} + +export_code %quote empty dequeue enqueue in Haskell + module_name Example file "examples/" + +text {* + \noindent This is how the corresponding code in @{text Haskell} looks like: +*} + +text %quote {*@{code_stmts empty enqueue dequeue (Haskell)}*} + +text {* + \noindent This demonstrates the basic usage of the @{command export_code} command; + for more details see \secref{sec:further}. +*} + +subsection {* Code generator architecture \label{sec:concept} *} + +text {* + What you have seen so far should be already enough in a lot of cases. If you + are content with this, you can quit reading here. Anyway, in order to customise + and adapt the code generator, it is inevitable to gain some understanding + how it works. + + \begin{figure}[h] + \begin{tikzpicture}[x = 4.2cm, y = 1cm] + \tikzstyle entity=[rounded corners, draw, thick, color = black, fill = white]; + \tikzstyle process=[ellipse, draw, thick, color = green, fill = white]; + \tikzstyle process_arrow=[->, semithick, color = green]; + \node (HOL) at (0, 4) [style=entity] {@{text "Isabelle/HOL"} theory}; + \node (eqn) at (2, 2) [style=entity] {code equations}; + \node (iml) at (2, 0) [style=entity] {intermediate language}; + \node (seri) at (1, 0) [style=process] {serialisation}; + \node (SML) at (0, 3) [style=entity] {@{text SML}}; + \node (OCaml) at (0, 2) [style=entity] {@{text OCaml}}; + \node (further) at (0, 1) [style=entity] {@{text "\"}}; + \node (Haskell) at (0, 0) [style=entity] {@{text Haskell}}; + \draw [style=process_arrow] (HOL) .. controls (2, 4) .. + node [style=process, near start] {selection} + node [style=process, near end] {preprocessing} + (eqn); + \draw [style=process_arrow] (eqn) -- node (transl) [style=process] {translation} (iml); + \draw [style=process_arrow] (iml) -- (seri); + \draw [style=process_arrow] (seri) -- (SML); + \draw [style=process_arrow] (seri) -- (OCaml); + \draw [style=process_arrow, dashed] (seri) -- (further); + \draw [style=process_arrow] (seri) -- (Haskell); + \end{tikzpicture} + \caption{Code generator architecture} + \label{fig:arch} + \end{figure} + + The code generator employs a notion of executability + for three foundational executable ingredients known + from functional programming: + \emph{code equations}, \emph{datatypes}, and + \emph{type classes}. A code equation as a first approximation + is a theorem of the form @{text "f t\<^isub>1 t\<^isub>2 \ t\<^isub>n \ t"} + (an equation headed by a constant @{text f} with arguments + @{text "t\<^isub>1 t\<^isub>2 \ t\<^isub>n"} and right hand side @{text t}). + Code generation aims to turn code equations + into a functional program. This is achieved by three major + components which operate sequentially, i.e. the result of one is + the input + of the next in the chain, see diagram \ref{fig:arch}: + + \begin{itemize} + + \item Out of the vast collection of theorems proven in a + \qn{theory}, a reasonable subset modelling + code equations is \qn{selected}. + + \item On those selected theorems, certain + transformations are carried out + (\qn{preprocessing}). Their purpose is to turn theorems + representing non- or badly executable + specifications into equivalent but executable counterparts. + The result is a structured collection of \qn{code theorems}. + + \item Before the selected code equations are continued with, + they can be \qn{preprocessed}, i.e. subjected to theorem + transformations. This \qn{preprocessor} is an interface which + allows to apply + the full expressiveness of ML-based theorem transformations + to code generation; motivating examples are shown below, see + \secref{sec:preproc}. + The result of the preprocessing step is a structured collection + of code equations. + + \item These code equations are \qn{translated} to a program + in an abstract intermediate language. Think of it as a kind + of \qt{Mini-Haskell} with four \qn{statements}: @{text data} + (for datatypes), @{text fun} (stemming from code equations), + also @{text class} and @{text inst} (for type classes). + + \item Finally, the abstract program is \qn{serialised} into concrete + source code of a target language. + + \end{itemize} + + \noindent From these steps, only the two last are carried out outside the logic; by + keeping this layer as thin as possible, the amount of code to trust is + kept to a minimum. +*} + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/ML.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/ML.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,177 @@ +theory "ML" +imports Setup +begin + +section {* ML system interfaces \label{sec:ml} *} + +text {* + Since the code generator framework not only aims to provide + a nice Isar interface but also to form a base for + code-generation-based applications, here a short + description of the most important ML interfaces. +*} + +subsection {* Executable theory content: @{text Code} *} + +text {* + This Pure module implements the core notions of + executable content of a theory. +*} + +subsubsection {* Managing executable content *} + +text %mlref {* + \begin{mldecls} + @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\ + @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\ + @{index_ML Code.add_eqnl: "string * (thm * bool) list lazy -> theory -> theory"} \\ + @{index_ML Code.map_pre: "(simpset -> simpset) -> theory -> theory"} \\ + @{index_ML Code.map_post: "(simpset -> simpset) -> theory -> theory"} \\ + @{index_ML Code.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option) + -> theory -> theory"} \\ + @{index_ML Code.del_functrans: "string -> theory -> theory"} \\ + @{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\ + @{index_ML Code.get_datatype: "theory -> string + -> (string * sort) list * (string * typ list) list"} \\ + @{index_ML Code.get_datatype_of_constr: "theory -> string -> string option"} + \end{mldecls} + + \begin{description} + + \item @{ML Code.add_eqn}~@{text "thm"}~@{text "thy"} adds function + theorem @{text "thm"} to executable content. + + \item @{ML Code.del_eqn}~@{text "thm"}~@{text "thy"} removes function + theorem @{text "thm"} from executable content, if present. + + \item @{ML Code.add_eqnl}~@{text "(const, lthms)"}~@{text "thy"} adds + suspended code equations @{text lthms} for constant + @{text const} to executable content. + + \item @{ML Code.map_pre}~@{text "f"}~@{text "thy"} changes + the preprocessor simpset. + + \item @{ML Code.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds + function transformer @{text f} (named @{text name}) to executable content; + @{text f} is a transformer of the code equations belonging + to a certain function definition, depending on the + current theory context. Returning @{text NONE} indicates that no + transformation took place; otherwise, the whole process will be iterated + with the new code equations. + + \item @{ML Code.del_functrans}~@{text "name"}~@{text "thy"} removes + function transformer named @{text name} from executable content. + + \item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds + a datatype to executable content, with generation + set @{text cs}. + + \item @{ML Code.get_datatype_of_constr}~@{text "thy"}~@{text "const"} + returns type constructor corresponding to + constructor @{text const}; returns @{text NONE} + if @{text const} is no constructor. + + \end{description} +*} + +subsection {* Auxiliary *} + +text %mlref {* + \begin{mldecls} + @{index_ML Code_Unit.read_const: "theory -> string -> string"} \\ + @{index_ML Code_Unit.head_eqn: "theory -> thm -> string * ((string * sort) list * typ)"} \\ + @{index_ML Code_Unit.rewrite_eqn: "simpset -> thm -> thm"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML Code_Unit.read_const}~@{text thy}~@{text s} + reads a constant as a concrete term expression @{text s}. + + \item @{ML Code_Unit.head_eqn}~@{text thy}~@{text thm} + extracts the constant and its type from a code equation @{text thm}. + + \item @{ML Code_Unit.rewrite_eqn}~@{text ss}~@{text thm} + rewrites a code equation @{text thm} with a simpset @{text ss}; + only arguments and right hand side are rewritten, + not the head of the code equation. + + \end{description} + +*} + +subsection {* Implementing code generator applications *} + +text {* + Implementing code generator applications on top + of the framework set out so far usually not only + involves using those primitive interfaces + but also storing code-dependent data and various + other things. +*} + +subsubsection {* Data depending on the theory's executable content *} + +text {* + Due to incrementality of code generation, changes in the + theory's executable content have to be propagated in a + certain fashion. Additionally, such changes may occur + not only during theory extension but also during theory + merge, which is a little bit nasty from an implementation + point of view. The framework provides a solution + to this technical challenge by providing a functorial + data slot @{ML_functor CodeDataFun}; on instantiation + of this functor, the following types and operations + are required: + + \medskip + \begin{tabular}{l} + @{text "type T"} \\ + @{text "val empty: T"} \\ + @{text "val purge: theory \ string list option \ T \ T"} + \end{tabular} + + \begin{description} + + \item @{text T} the type of data to store. + + \item @{text empty} initial (empty) data. + + \item @{text purge}~@{text thy}~@{text consts} propagates changes in executable content; + @{text consts} indicates the kind + of change: @{ML NONE} stands for a fundamental change + which invalidates any existing code, @{text "SOME consts"} + hints that executable content for constants @{text consts} + has changed. + + \end{description} + + \noindent An instance of @{ML_functor CodeDataFun} provides the following + interface: + + \medskip + \begin{tabular}{l} + @{text "get: theory \ T"} \\ + @{text "change: theory \ (T \ T) \ T"} \\ + @{text "change_yield: theory \ (T \ 'a * T) \ 'a * T"} + \end{tabular} + + \begin{description} + + \item @{text get} retrieval of the current data. + + \item @{text change} update of current data (cached!) + by giving a continuation. + + \item @{text change_yield} update with side result. + + \end{description} +*} + +text {* + \bigskip + + \emph{Happy proving, happy hacking!} +*} + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/Program.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/Program.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,526 @@ +theory Program +imports Introduction +begin + +section {* Turning Theories into Programs \label{sec:program} *} + +subsection {* The @{text "Isabelle/HOL"} default setup *} + +text {* + We have already seen how by default equations stemming from + @{command definition}/@{command primrec}/@{command fun} + statements are used for code generation. This default behaviour + can be changed, e.g. by providing different code equations. + All kinds of customisation shown in this section is \emph{safe} + in the sense that the user does not have to worry about + correctness -- all programs generatable that way are partially + correct. +*} + +subsection {* Selecting code equations *} + +text {* + Coming back to our introductory example, we + could provide an alternative code equations for @{const dequeue} + explicitly: +*} + +lemma %quote [code]: + "dequeue (AQueue xs []) = + (if xs = [] then (None, AQueue [] []) + else dequeue (AQueue [] (rev xs)))" + "dequeue (AQueue xs (y # ys)) = + (Some y, AQueue xs ys)" + by (cases xs, simp_all) (cases "rev xs", simp_all) + +text {* + \noindent The annotation @{text "[code]"} is an @{text Isar} + @{text attribute} which states that the given theorems should be + considered as code equations for a @{text fun} statement -- + the corresponding constant is determined syntactically. The resulting code: +*} + +text %quote {*@{code_stmts dequeue (consts) dequeue (Haskell)}*} + +text {* + \noindent You may note that the equality test @{term "xs = []"} has been + replaced by the predicate @{term "null xs"}. This is due to the default + setup in the \qn{preprocessor} to be discussed further below (\secref{sec:preproc}). + + Changing the default constructor set of datatypes is also + possible. See \secref{sec:datatypes} for an example. + + As told in \secref{sec:concept}, code generation is based + on a structured collection of code theorems. + For explorative purpose, this collection + may be inspected using the @{command code_thms} command: +*} + +code_thms %quote dequeue + +text {* + \noindent prints a table with \emph{all} code equations + for @{const dequeue}, including + \emph{all} code equations those equations depend + on recursively. + + Similarly, the @{command code_deps} command shows a graph + visualising dependencies between code equations. +*} + +subsection {* @{text class} and @{text instantiation} *} + +text {* + Concerning type classes and code generation, let us examine an example + from abstract algebra: +*} + +class %quote semigroup = + fixes mult :: "'a \ 'a \ 'a" (infixl "\" 70) + assumes assoc: "(x \ y) \ z = x \ (y \ z)" + +class %quote monoid = semigroup + + fixes neutral :: 'a ("\") + assumes neutl: "\ \ x = x" + and neutr: "x \ \ = x" + +instantiation %quote nat :: monoid +begin + +primrec %quote mult_nat where + "0 \ n = (0\nat)" + | "Suc m \ n = n + m \ n" + +definition %quote neutral_nat where + "\ = Suc 0" + +lemma %quote add_mult_distrib: + fixes n m q :: nat + shows "(n + m) \ q = n \ q + m \ q" + by (induct n) simp_all + +instance %quote proof + fix m n q :: nat + show "m \ n \ q = m \ (n \ q)" + by (induct m) (simp_all add: add_mult_distrib) + show "\ \ n = n" + by (simp add: neutral_nat_def) + show "m \ \ = m" + by (induct m) (simp_all add: neutral_nat_def) +qed + +end %quote + +text {* + \noindent We define the natural operation of the natural numbers + on monoids: +*} + +primrec %quote (in monoid) pow :: "nat \ 'a \ 'a" where + "pow 0 a = \" + | "pow (Suc n) a = a \ pow n a" + +text {* + \noindent This we use to define the discrete exponentiation function: +*} + +definition %quote bexp :: "nat \ nat" where + "bexp n = pow n (Suc (Suc 0))" + +text {* + \noindent The corresponding code: +*} + +text %quote {*@{code_stmts bexp (Haskell)}*} + +text {* + \noindent This is a convenient place to show how explicit dictionary construction + manifests in generated code (here, the same example in @{text SML}): +*} + +text %quote {*@{code_stmts bexp (SML)}*} + +text {* + \noindent Note the parameters with trailing underscore (@{verbatim "A_"}) + which are the dictionary parameters. +*} + +subsection {* The preprocessor \label{sec:preproc} *} + +text {* + Before selected function theorems are turned into abstract + code, a chain of definitional transformation steps is carried + out: \emph{preprocessing}. In essence, the preprocessor + consists of two components: a \emph{simpset} and \emph{function transformers}. + + The \emph{simpset} allows to employ the full generality of the Isabelle + simplifier. Due to the interpretation of theorems + as code equations, rewrites are applied to the right + hand side and the arguments of the left hand side of an + equation, but never to the constant heading the left hand side. + An important special case are \emph{inline theorems} which may be + declared and undeclared using the + \emph{code inline} or \emph{code inline del} attribute respectively. + + Some common applications: +*} + +text_raw {* + \begin{itemize} +*} + +text {* + \item replacing non-executable constructs by executable ones: +*} + +lemma %quote [code inline]: + "x \ set xs \ x mem xs" by (induct xs) simp_all + +text {* + \item eliminating superfluous constants: +*} + +lemma %quote [code inline]: + "1 = Suc 0" by simp + +text {* + \item replacing executable but inconvenient constructs: +*} + +lemma %quote [code inline]: + "xs = [] \ List.null xs" by (induct xs) simp_all + +text_raw {* + \end{itemize} +*} + +text {* + \noindent \emph{Function transformers} provide a very general interface, + transforming a list of function theorems to another + list of function theorems, provided that neither the heading + constant nor its type change. The @{term "0\nat"} / @{const Suc} + pattern elimination implemented in + theory @{text Efficient_Nat} (see \secref{eff_nat}) uses this + interface. + + \noindent The current setup of the preprocessor may be inspected using + the @{command print_codesetup} command. + @{command code_thms} provides a convenient + mechanism to inspect the impact of a preprocessor setup + on code equations. + + \begin{warn} + The attribute \emph{code unfold} + associated with the @{text "SML code generator"} also applies to + the @{text "generic code generator"}: + \emph{code unfold} implies \emph{code inline}. + \end{warn} +*} + +subsection {* Datatypes \label{sec:datatypes} *} + +text {* + Conceptually, any datatype is spanned by a set of + \emph{constructors} of type @{text "\ = \ \ \ \\<^isub>1 \ \\<^isub>n"} where @{text + "{\\<^isub>1, \, \\<^isub>n}"} is exactly the set of \emph{all} type variables in + @{text "\"}. The HOL datatype package by default registers any new + datatype in the table of datatypes, which may be inspected using the + @{command print_codesetup} command. + + In some cases, it is appropriate to alter or extend this table. As + an example, we will develop an alternative representation of the + queue example given in \secref{sec:intro}. The amortised + representation is convenient for generating code but exposes its + \qt{implementation} details, which may be cumbersome when proving + theorems about it. Therefore, here a simple, straightforward + representation of queues: +*} + +datatype %quote 'a queue = Queue "'a list" + +definition %quote empty :: "'a queue" where + "empty = Queue []" + +primrec %quote enqueue :: "'a \ 'a queue \ 'a queue" where + "enqueue x (Queue xs) = Queue (xs @ [x])" + +fun %quote dequeue :: "'a queue \ 'a option \ 'a queue" where + "dequeue (Queue []) = (None, Queue [])" + | "dequeue (Queue (x # xs)) = (Some x, Queue xs)" + +text {* + \noindent This we can use directly for proving; for executing, + we provide an alternative characterisation: +*} + +definition %quote AQueue :: "'a list \ 'a list \ 'a queue" where + "AQueue xs ys = Queue (ys @ rev xs)" + +code_datatype %quote AQueue + +text {* + \noindent Here we define a \qt{constructor} @{const "AQueue"} which + is defined in terms of @{text "Queue"} and interprets its arguments + according to what the \emph{content} of an amortised queue is supposed + to be. Equipped with this, we are able to prove the following equations + for our primitive queue operations which \qt{implement} the simple + queues in an amortised fashion: +*} + +lemma %quote empty_AQueue [code]: + "empty = AQueue [] []" + unfolding AQueue_def empty_def by simp + +lemma %quote enqueue_AQueue [code]: + "enqueue x (AQueue xs ys) = AQueue (x # xs) ys" + unfolding AQueue_def by simp + +lemma %quote dequeue_AQueue [code]: + "dequeue (AQueue xs []) = + (if xs = [] then (None, AQueue [] []) + else dequeue (AQueue [] (rev xs)))" + "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)" + unfolding AQueue_def by simp_all + +text {* + \noindent For completeness, we provide a substitute for the + @{text case} combinator on queues: +*} + +lemma %quote queue_case_AQueue [code]: + "queue_case f (AQueue xs ys) = f (ys @ rev xs)" + unfolding AQueue_def by simp + +text {* + \noindent The resulting code looks as expected: +*} + +text %quote {*@{code_stmts empty enqueue dequeue (SML)}*} + +text {* + \noindent From this example, it can be glimpsed that using own + constructor sets is a little delicate since it changes the set of + valid patterns for values of that type. Without going into much + detail, here some practical hints: + + \begin{itemize} + + \item When changing the constructor set for datatypes, take care + to provide alternative equations for the @{text case} combinator. + + \item Values in the target language need not to be normalised -- + different values in the target language may represent the same + value in the logic. + + \item Usually, a good methodology to deal with the subtleties of + pattern matching is to see the type as an abstract type: provide + a set of operations which operate on the concrete representation + of the type, and derive further operations by combinations of + these primitive ones, without relying on a particular + representation. + + \end{itemize} +*} + + +subsection {* Equality and wellsortedness *} + +text {* + Surely you have already noticed how equality is treated + by the code generator: +*} + +primrec %quote collect_duplicates :: "'a list \ 'a list \ 'a list \ 'a list" where + "collect_duplicates xs ys [] = xs" + | "collect_duplicates xs ys (z#zs) = (if z \ set xs + then if z \ set ys + then collect_duplicates xs ys zs + else collect_duplicates xs (z#ys) zs + else collect_duplicates (z#xs) (z#ys) zs)" + +text {* + \noindent The membership test during preprocessing is rewritten, + resulting in @{const List.member}, which itself + performs an explicit equality check. +*} + +text %quote {*@{code_stmts collect_duplicates (SML)}*} + +text {* + \noindent Obviously, polymorphic equality is implemented the Haskell + way using a type class. How is this achieved? HOL introduces + an explicit class @{class eq} with a corresponding operation + @{const eq_class.eq} such that @{thm eq [no_vars]}. + The preprocessing framework does the rest by propagating the + @{class eq} constraints through all dependent code equations. + For datatypes, instances of @{class eq} are implicitly derived + when possible. For other types, you may instantiate @{text eq} + manually like any other type class. + + Though this @{text eq} class is designed to get rarely in + the way, a subtlety + enters the stage when definitions of overloaded constants + are dependent on operational equality. For example, let + us define a lexicographic ordering on tuples + (also see theory @{theory Product_ord}): +*} + +instantiation %quote "*" :: (order, order) order +begin + +definition %quote [code del]: + "x \ y \ fst x < fst y \ fst x = fst y \ snd x \ snd y" + +definition %quote [code del]: + "x < y \ fst x < fst y \ fst x = fst y \ snd x < snd y" + +instance %quote proof +qed (auto simp: less_eq_prod_def less_prod_def intro: order_less_trans) + +end %quote + +lemma %quote order_prod [code]: + "(x1 \ 'a\order, y1 \ 'b\order) < (x2, y2) \ + x1 < x2 \ x1 = x2 \ y1 < y2" + "(x1 \ 'a\order, y1 \ 'b\order) \ (x2, y2) \ + x1 < x2 \ x1 = x2 \ y1 \ y2" + by (simp_all add: less_prod_def less_eq_prod_def) + +text {* + \noindent Then code generation will fail. Why? The definition + of @{term "op \"} depends on equality on both arguments, + which are polymorphic and impose an additional @{class eq} + class constraint, which the preprocessor does not propagate + (for technical reasons). + + The solution is to add @{class eq} explicitly to the first sort arguments in the + code theorems: +*} + +lemma %quote order_prod_code [code]: + "(x1 \ 'a\{order, eq}, y1 \ 'b\order) < (x2, y2) \ + x1 < x2 \ x1 = x2 \ y1 < y2" + "(x1 \ 'a\{order, eq}, y1 \ 'b\order) \ (x2, y2) \ + x1 < x2 \ x1 = x2 \ y1 \ y2" + by (simp_all add: less_prod_def less_eq_prod_def) + +text {* + \noindent Then code generation succeeds: +*} + +text %quote {*@{code_stmts "op \ \ _ \ _ \ _ \ _ \ bool" (SML)}*} + +text {* + In some cases, the automatically derived code equations + for equality on a particular type may not be appropriate. + As example, watch the following datatype representing + monomorphic parametric types (where type constructors + are referred to by natural numbers): +*} + +datatype %quote monotype = Mono nat "monotype list" +(*<*) +lemma monotype_eq: + "eq_class.eq (Mono tyco1 typargs1) (Mono tyco2 typargs2) \ + eq_class.eq tyco1 tyco2 \ eq_class.eq typargs1 typargs2" by (simp add: eq) +(*>*) + +text {* + \noindent Then code generation for SML would fail with a message + that the generated code contains illegal mutual dependencies: + the theorem @{thm monotype_eq [no_vars]} already requires the + instance @{text "monotype \ eq"}, which itself requires + @{thm monotype_eq [no_vars]}; Haskell has no problem with mutually + recursive @{text instance} and @{text function} definitions, + but the SML serialiser does not support this. + + In such cases, you have to provide your own equality equations + involving auxiliary constants. In our case, + @{const [show_types] list_all2} can do the job: +*} + +lemma %quote monotype_eq_list_all2 [code]: + "eq_class.eq (Mono tyco1 typargs1) (Mono tyco2 typargs2) \ + eq_class.eq tyco1 tyco2 \ list_all2 eq_class.eq typargs1 typargs2" + by (simp add: eq list_all2_eq [symmetric]) + +text {* + \noindent does not depend on instance @{text "monotype \ eq"}: +*} + +text %quote {*@{code_stmts "eq_class.eq :: monotype \ monotype \ bool" (SML)}*} + + +subsection {* Explicit partiality *} + +text {* + Partiality usually enters the game by partial patterns, as + in the following example, again for amortised queues: +*} + +definition %quote strict_dequeue :: "'a queue \ 'a \ 'a queue" where + "strict_dequeue q = (case dequeue q + of (Some x, q') \ (x, q'))" + +lemma %quote strict_dequeue_AQueue [code]: + "strict_dequeue (AQueue xs (y # ys)) = (y, AQueue xs ys)" + "strict_dequeue (AQueue xs []) = + (case rev xs of y # ys \ (y, AQueue [] ys))" + by (simp_all add: strict_dequeue_def dequeue_AQueue split: list.splits) + +text {* + \noindent In the corresponding code, there is no equation + for the pattern @{term "AQueue [] []"}: +*} + +text %quote {*@{code_stmts strict_dequeue (consts) strict_dequeue (Haskell)}*} + +text {* + \noindent In some cases it is desirable to have this + pseudo-\qt{partiality} more explicitly, e.g.~as follows: +*} + +axiomatization %quote empty_queue :: 'a + +definition %quote strict_dequeue' :: "'a queue \ 'a \ 'a queue" where + "strict_dequeue' q = (case dequeue q of (Some x, q') \ (x, q') | _ \ empty_queue)" + +lemma %quote strict_dequeue'_AQueue [code]: + "strict_dequeue' (AQueue xs []) = (if xs = [] then empty_queue + else strict_dequeue' (AQueue [] (rev xs)))" + "strict_dequeue' (AQueue xs (y # ys)) = + (y, AQueue xs ys)" + by (simp_all add: strict_dequeue'_def dequeue_AQueue split: list.splits) + +text {* + Observe that on the right hand side of the definition of @{const + "strict_dequeue'"} the constant @{const empty_queue} occurs + which is unspecified. + + Normally, if constants without any code equations occur in a + program, the code generator complains (since in most cases this is + not what the user expects). But such constants can also be thought + of as function definitions with no equations which always fail, + since there is never a successful pattern match on the left hand + side. In order to categorise a constant into that category + explicitly, use @{command "code_abort"}: +*} + +code_abort %quote empty_queue + +text {* + \noindent Then the code generator will just insert an error or + exception at the appropriate position: +*} + +text %quote {*@{code_stmts strict_dequeue' (consts) empty_queue strict_dequeue' (Haskell)}*} + +text {* + \noindent This feature however is rarely needed in practice. + Note also that the @{text HOL} default setup already declares + @{const undefined} as @{command "code_abort"}, which is most + likely to be used in such situations. +*} + +end + \ No newline at end of file diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/ROOT.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,11 @@ + +(* $Id$ *) + +no_document use_thy "Setup"; +no_document use_thys ["Efficient_Nat"]; + +use_thy "Introduction"; +use_thy "Program"; +use_thy "Adaption"; +use_thy "Further"; +use_thy "ML"; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/Setup.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/Setup.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,15 @@ +theory Setup +imports Complex_Main +uses + "../../antiquote_setup.ML" + "../../more_antiquote.ML" +begin + +ML {* no_document use_thys + ["Efficient_Nat", "Code_Char_chr", "Product_ord", "~~/src/HOL/Imperative_HOL/Imperative_HOL", + "~~/src/HOL/Decision_Procs/Ferrack"] *} + +ML_command {* Code_Target.code_width := 74 *} +ML_command {* reset unique_names *} + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/document/Adaption.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/document/Adaption.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,678 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Adaption}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Adaption\isanewline +\isakeyword{imports}\ Setup\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +\isanewline +% +\endisadelimtheory +% +\isadeliminvisible +\isanewline +% +\endisadeliminvisible +% +\isataginvisible +\isacommand{setup}\isamarkupfalse% +\ {\isacharverbatimopen}\ Code{\isacharunderscore}Target{\isachardot}extend{\isacharunderscore}target\ {\isacharparenleft}{\isachardoublequote}{\isasymSML}{\isachardoublequote}{\isacharcomma}\ {\isacharparenleft}{\isachardoublequote}SML{\isachardoublequote}{\isacharcomma}\ K\ I{\isacharparenright}{\isacharparenright}\ {\isacharverbatimclose}% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isamarkupsection{Adaption to target languages \label{sec:adaption}% +} +\isamarkuptrue% +% +\isamarkupsubsection{Adapting code generation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The aspects of code generation introduced so far have two aspects + in common: + + \begin{itemize} + \item They act uniformly, without reference to a specific + target language. + \item They are \emph{safe} in the sense that as long as you trust + the code generator meta theory and implementation, you cannot + produce programs that yield results which are not derivable + in the logic. + \end{itemize} + + \noindent In this section we will introduce means to \emph{adapt} the serialiser + to a specific target language, i.e.~to print program fragments + in a way which accommodates \qt{already existing} ingredients of + a target language environment, for three reasons: + + \begin{itemize} + \item improving readability and aesthetics of generated code + \item gaining efficiency + \item interface with language parts which have no direct counterpart + in \isa{HOL} (say, imperative data structures) + \end{itemize} + + \noindent Generally, you should avoid using those features yourself + \emph{at any cost}: + + \begin{itemize} + \item The safe configuration methods act uniformly on every target language, + whereas for adaption you have to treat each target language separate. + \item Application is extremely tedious since there is no abstraction + which would allow for a static check, making it easy to produce garbage. + \item More or less subtle errors can be introduced unconsciously. + \end{itemize} + + \noindent However, even if you ought refrain from setting up adaption + yourself, already the \isa{HOL} comes with some reasonable default + adaptions (say, using target language list syntax). There also some + common adaption cases which you can setup by importing particular + library theories. In order to understand these, we provide some clues here; + these however are not supposed to replace a careful study of the sources.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{The adaption principle% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The following figure illustrates what \qt{adaption} is conceptually + supposed to be: + + \begin{figure}[here] + \begin{tikzpicture}[scale = 0.5] + \tikzstyle water=[color = blue, thick] + \tikzstyle ice=[color = black, very thick, cap = round, join = round, fill = white] + \tikzstyle process=[color = green, semithick, ->] + \tikzstyle adaption=[color = red, semithick, ->] + \tikzstyle target=[color = black] + \foreach \x in {0, ..., 24} + \draw[style=water] (\x, 0.25) sin + (0.25, 0.25) cos + (0.25, -0.25) sin + + (0.25, -0.25) cos + (0.25, 0.25); + \draw[style=ice] (1, 0) -- + (3, 6) node[above, fill=white] {logic} -- (5, 0) -- cycle; + \draw[style=ice] (9, 0) -- + (11, 6) node[above, fill=white] {intermediate language} -- (13, 0) -- cycle; + \draw[style=ice] (15, -6) -- + (19, 6) node[above, fill=white] {target language} -- (23, -6) -- cycle; + \draw[style=process] + (3.5, 3) .. controls (7, 5) .. node[fill=white] {translation} (10.5, 3); + \draw[style=process] + (11.5, 3) .. controls (15, 5) .. node[fill=white] (serialisation) {serialisation} (18.5, 3); + \node (adaption) at (11, -2) [style=adaption] {adaption}; + \node at (19, 3) [rotate=90] {generated}; + \node at (19.5, -5) {language}; + \node at (19.5, -3) {library}; + \node (includes) at (19.5, -1) {includes}; + \node (reserved) at (16.5, -3) [rotate=72] {reserved}; % proper 71.57 + \draw[style=process] + (includes) -- (serialisation); + \draw[style=process] + (reserved) -- (serialisation); + \draw[style=adaption] + (adaption) -- (serialisation); + \draw[style=adaption] + (adaption) -- (includes); + \draw[style=adaption] + (adaption) -- (reserved); + \end{tikzpicture} + \caption{The adaption principle} + \label{fig:adaption} + \end{figure} + + \noindent In the tame view, code generation acts as broker between + \isa{logic}, \isa{intermediate\ language} and + \isa{target\ language} by means of \isa{translation} and + \isa{serialisation}; for the latter, the serialiser has to observe + the structure of the \isa{language} itself plus some \isa{reserved} + keywords which have to be avoided for generated code. + However, if you consider \isa{adaption} mechanisms, the code generated + by the serializer is just the tip of the iceberg: + + \begin{itemize} + \item \isa{serialisation} can be \emph{parametrised} such that + logical entities are mapped to target-specific ones + (e.g. target-specific list syntax, + see also \secref{sec:adaption_mechanisms}) + \item Such parametrisations can involve references to a + target-specific standard \isa{library} (e.g. using + the \isa{Haskell} \verb|Maybe| type instead + of the \isa{HOL} \isa{option} type); + if such are used, the corresponding identifiers + (in our example, \verb|Maybe|, \verb|Nothing| + and \verb|Just|) also have to be considered \isa{reserved}. + \item Even more, the user can enrich the library of the + target-language by providing code snippets + (\qt{\isa{includes}}) which are prepended to + any generated code (see \secref{sec:include}); this typically + also involves further \isa{reserved} identifiers. + \end{itemize} + + \noindent As figure \ref{fig:adaption} illustrates, all these adaption mechanisms + have to act consistently; it is at the discretion of the user + to take care for this.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Common adaption patterns% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} \hyperlink{theory.Main}{\mbox{\isa{Main}}} theory already provides a code + generator setup + which should be suitable for most applications. Common extensions + and modifications are available by certain theories of the \isa{HOL} + library; beside being useful in applications, they may serve + as a tutorial for customising the code generator setup (see below + \secref{sec:adaption_mechanisms}). + + \begin{description} + + \item[\hyperlink{theory.Code-Integer}{\mbox{\isa{Code{\isacharunderscore}Integer}}}] represents \isa{HOL} integers by big + integer literals in target languages. + \item[\hyperlink{theory.Code-Char}{\mbox{\isa{Code{\isacharunderscore}Char}}}] represents \isa{HOL} characters by + character literals in target languages. + \item[\hyperlink{theory.Code-Char-chr}{\mbox{\isa{Code{\isacharunderscore}Char{\isacharunderscore}chr}}}] like \isa{Code{\isacharunderscore}Char}, + but also offers treatment of character codes; includes + \hyperlink{theory.Code-Char}{\mbox{\isa{Code{\isacharunderscore}Char}}}. + \item[\hyperlink{theory.Efficient-Nat}{\mbox{\isa{Efficient{\isacharunderscore}Nat}}}] \label{eff_nat} implements natural numbers by integers, + which in general will result in higher efficiency; pattern + matching with \isa{{\isadigit{0}}} / \isa{Suc} + is eliminated; includes \hyperlink{theory.Code-Integer}{\mbox{\isa{Code{\isacharunderscore}Integer}}} + and \hyperlink{theory.Code-Index}{\mbox{\isa{Code{\isacharunderscore}Index}}}. + \item[\hyperlink{theory.Code-Index}{\mbox{\isa{Code{\isacharunderscore}Index}}}] provides an additional datatype + \isa{index} which is mapped to target-language built-in integers. + Useful for code setups which involve e.g. indexing of + target-language arrays. + \item[\hyperlink{theory.Code-Message}{\mbox{\isa{Code{\isacharunderscore}Message}}}] provides an additional datatype + \isa{message{\isacharunderscore}string} which is isomorphic to strings; + \isa{message{\isacharunderscore}string}s are mapped to target-language strings. + Useful for code setups which involve e.g. printing (error) messages. + + \end{description} + + \begin{warn} + When importing any of these theories, they should form the last + items in an import list. Since these theories adapt the + code generator setup in a non-conservative fashion, + strange effects may occur otherwise. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Parametrising serialisation \label{sec:adaption_mechanisms}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Consider the following function and its corresponding + SML code:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{primrec}\isamarkupfalse% +\ in{\isacharunderscore}interval\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymtimes}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}in{\isacharunderscore}interval\ {\isacharparenleft}k{\isacharcomma}\ l{\isacharparenright}\ n\ {\isasymlongleftrightarrow}\ k\ {\isasymle}\ n\ {\isasymand}\ n\ {\isasymle}\ l{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype boola = True | False;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun anda x True = x\\ +\hspace*{0pt} ~| anda x False = False\\ +\hspace*{0pt} ~| anda True x = x\\ +\hspace*{0pt} ~| anda False x = False;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\ +\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = False\\ +\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\ +\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = True;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun in{\char95}interval (k,~l) n = anda (less{\char95}eq{\char95}nat k n) (less{\char95}eq{\char95}nat n l);\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Though this is correct code, it is a little bit unsatisfactory: + boolean values and operators are materialised as distinguished + entities with have nothing to do with the SML-built-in notion + of \qt{bool}. This results in less readable code; + additionally, eager evaluation may cause programs to + loop or break which would perfectly terminate when + the existing SML \verb|bool| would be used. To map + the HOL \isa{bool} on SML \verb|bool|, we may use + \qn{custom serialisations}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}type}\isamarkupfalse% +\ bool\isanewline +\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}bool{\isachardoublequoteclose}{\isacharparenright}\isanewline +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ True\ \isakeyword{and}\ False\ \isakeyword{and}\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline +\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}true{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}false{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}{\isacharunderscore}\ andalso\ {\isacharunderscore}{\isachardoublequoteclose}{\isacharparenright}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\begin{isamarkuptext}% +\noindent The \hyperlink{command.code-type}{\mbox{\isa{\isacommand{code{\isacharunderscore}type}}}} command takes a type constructor + as arguments together with a list of custom serialisations. + Each custom serialisation starts with a target language + identifier followed by an expression, which during + code serialisation is inserted whenever the type constructor + would occur. For constants, \hyperlink{command.code-const}{\mbox{\isa{\isacommand{code{\isacharunderscore}const}}}} implements + the corresponding mechanism. Each ``\verb|_|'' in + a serialisation expression is treated as a placeholder + for the type constructor's (the constant's) arguments.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\ +\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = false\\ +\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\ +\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = true;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun in{\char95}interval (k,~l) n = (less{\char95}eq{\char95}nat k n) andalso (less{\char95}eq{\char95}nat n l);\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This still is not perfect: the parentheses + around the \qt{andalso} expression are superfluous. + Though the serialiser + by no means attempts to imitate the rich Isabelle syntax + framework, it provides some common idioms, notably + associative infixes with precedences which may be used here:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline +\ \ {\isacharparenleft}SML\ \isakeyword{infixl}\ {\isadigit{1}}\ {\isachardoublequoteopen}andalso{\isachardoublequoteclose}{\isacharparenright}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\ +\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = false\\ +\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\ +\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = true;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun in{\char95}interval (k,~l) n = less{\char95}eq{\char95}nat k n andalso less{\char95}eq{\char95}nat n l;\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The attentive reader may ask how we assert that no generated + code will accidentally overwrite. For this reason the serialiser has + an internal table of identifiers which have to be avoided to be used + for new declarations. Initially, this table typically contains the + keywords of the target language. It can be extended manually, thus avoiding + accidental overwrites, using the \hyperlink{command.code-reserved}{\mbox{\isa{\isacommand{code{\isacharunderscore}reserved}}}} command:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{code{\isacharunderscore}reserved}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymSML}{\isachardoublequoteclose}\ bool\ true\ false\ andalso% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Next, we try to map HOL pairs to SML pairs, using the + infix ``\verb|*|'' type constructor and parentheses:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}type}\isamarkupfalse% +\ {\isacharasterisk}\isanewline +\ \ {\isacharparenleft}SML\ \isakeyword{infix}\ {\isadigit{2}}\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ Pair\isanewline +\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}{\isacharbang}{\isacharparenleft}{\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharcomma}{\isacharslash}\ {\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\begin{isamarkuptext}% +\noindent The initial bang ``\verb|!|'' tells the serialiser + never to put + parentheses around the whole expression (they are already present), + while the parentheses around argument place holders + tell not to put parentheses around the arguments. + The slash ``\verb|/|'' (followed by arbitrary white space) + inserts a space which may be used as a break if necessary + during pretty printing. + + These examples give a glimpse what mechanisms + custom serialisations provide; however their usage + requires careful thinking in order not to introduce + inconsistencies -- or, in other words: + custom serialisations are completely axiomatic. + + A further noteworthy details is that any special + character in a custom serialisation may be quoted + using ``\verb|'|''; thus, in + ``\verb|fn '_ => _|'' the first + ``\verb|_|'' is a proper underscore while the + second ``\verb|_|'' is a placeholder.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{\isa{Haskell} serialisation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +For convenience, the default + \isa{HOL} setup for \isa{Haskell} maps the \isa{eq} class to + its counterpart in \isa{Haskell}, giving custom serialisations + for the class \isa{eq} (by command \hyperlink{command.code-class}{\mbox{\isa{\isacommand{code{\isacharunderscore}class}}}}) and its operation + \isa{eq{\isacharunderscore}class{\isachardot}eq}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}class}\isamarkupfalse% +\ eq\isanewline +\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Eq{\isachardoublequoteclose}{\isacharparenright}\isanewline +\isanewline +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\isanewline +\ \ {\isacharparenleft}Haskell\ \isakeyword{infixl}\ {\isadigit{4}}\ {\isachardoublequoteopen}{\isacharequal}{\isacharequal}{\isachardoublequoteclose}{\isacharparenright}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\begin{isamarkuptext}% +\noindent A problem now occurs whenever a type which + is an instance of \isa{eq} in \isa{HOL} is mapped + on a \isa{Haskell}-built-in type which is also an instance + of \isa{Haskell} \isa{Eq}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{typedecl}\isamarkupfalse% +\ bar\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}x{\isasymColon}bar{\isacharparenright}\ y\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{by}\isamarkupfalse% +\ default\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq{\isacharunderscore}bar{\isacharunderscore}def{\isacharparenright}\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +\isanewline +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}type}\isamarkupfalse% +\ bar\isanewline +\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Integer{\isachardoublequoteclose}{\isacharparenright}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\begin{isamarkuptext}% +\noindent The code generator would produce + an additional instance, which of course is rejected by the \isa{Haskell} + compiler. + To suppress this additional instance, use + \isa{code{\isacharunderscore}instance}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}instance}\isamarkupfalse% +\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline +\ \ {\isacharparenleft}Haskell\ {\isacharminus}{\isacharparenright}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\isamarkupsubsection{Enhancing the target language context \label{sec:include}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In rare cases it is necessary to \emph{enrich} the context of a + target language; this is accomplished using the \hyperlink{command.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}} + command:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{code{\isacharunderscore}include}\isamarkupfalse% +\ Haskell\ {\isachardoublequoteopen}Errno{\isachardoublequoteclose}\isanewline +{\isacharverbatimopen}errno\ i\ {\isacharequal}\ error\ {\isacharparenleft}{\isachardoublequote}Error\ number{\isacharcolon}\ {\isachardoublequote}\ {\isacharplus}{\isacharplus}\ show\ i{\isacharparenright}{\isacharverbatimclose}\isanewline +\isanewline +\isacommand{code{\isacharunderscore}reserved}\isamarkupfalse% +\ Haskell\ Errno% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\begin{isamarkuptext}% +\noindent Such named \isa{include}s are then prepended to every generated code. + Inspect such code in order to find out how \hyperlink{command.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}} behaves + with respect to a particular target language.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/document/Codegen.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/document/Codegen.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,1690 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Codegen}% +% +\isadelimtheory +\isanewline +\isanewline +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\isamarkupchapter{Code generation from Isabelle theories% +} +\isamarkuptrue% +% +\isamarkupsection{Introduction% +} +\isamarkuptrue% +% +\isamarkupsubsection{Motivation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Executing formal specifications as programs is a well-established + topic in the theorem proving community. With increasing + application of theorem proving systems in the area of + software development and verification, its relevance manifests + for running test cases and rapid prototyping. In logical + calculi like constructive type theory, + a notion of executability is implicit due to the nature + of the calculus. In contrast, specifications in Isabelle + can be highly non-executable. In order to bridge + the gap between logic and executable specifications, + an explicit non-trivial transformation has to be applied: + code generation. + + This tutorial introduces a generic code generator for the + Isabelle system \cite{isa-tutorial}. + Generic in the sense that the + \qn{target language} for which code shall ultimately be + generated is not fixed but may be an arbitrary state-of-the-art + functional programming language (currently, the implementation + supports SML \cite{SML}, OCaml \cite{OCaml} and Haskell + \cite{haskell-revised-report}). + We aim to provide a + versatile environment + suitable for software development and verification, + structuring the process + of code generation into a small set of orthogonal principles + while achieving a big coverage of application areas + with maximum flexibility. + + Conceptually the code generator framework is part + of Isabelle's \isa{Pure} meta logic; the object logic + \isa{HOL} which is an extension of \isa{Pure} + already comes with a reasonable framework setup and thus provides + a good working horse for raising code-generation-driven + applications. So, we assume some familiarity and experience + with the ingredients of the \isa{HOL} \emph{Main} theory + (see also \cite{isa-tutorial}).% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Overview% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The code generator aims to be usable with no further ado + in most cases while allowing for detailed customization. + This manifests in the structure of this tutorial: + we start with a generic example \secref{sec:example} + and introduce code generation concepts \secref{sec:concept}. + Section + \secref{sec:basics} explains how to use the framework naively, + presuming a reasonable default setup. Then, section + \secref{sec:advanced} deals with advanced topics, + introducing further aspects of the code generator framework + in a motivation-driven manner. Last, section \secref{sec:ml} + introduces the framework's internal programming interfaces. + + \begin{warn} + Ultimately, the code generator which this tutorial deals with + is supposed to replace the already established code generator + by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. + So, for the moment, there are two distinct code generators + in Isabelle. + Also note that while the framework itself is + object-logic independent, only \isa{HOL} provides a reasonable + framework setup. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{An example: a simple theory of search trees \label{sec:example}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +When writing executable specifications using \isa{HOL}, + it is convenient to use + three existing packages: the datatype package for defining + datatypes, the function package for (recursive) functions, + and the class package for overloaded definitions. + + We develope a small theory of search trees; trees are represented + as a datatype with key type \isa{{\isacharprime}a} and value type \isa{{\isacharprime}b}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isacharequal}\ Leaf\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}linorder{\isachardoublequoteclose}\ {\isacharprime}b\isanewline +\ \ {\isacharbar}\ Branch\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharprime}a{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent Note that we have constrained the type of keys + to the class of total orders, \isa{linorder}. + + We define \isa{find} and \isa{update} functions:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\isanewline +\ \ find\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isasymColon}linorder{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}find\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ it\ {\isacharequal}\ {\isacharparenleft}if\ it\ {\isacharequal}\ key\ then\ Some\ val\ else\ None{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}find\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ t{\isadigit{2}}{\isacharparenright}\ it\ {\isacharequal}\ {\isacharparenleft}if\ it\ {\isasymle}\ key\ then\ find\ t{\isadigit{1}}\ it\ else\ find\ t{\isadigit{2}}\ it{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{fun}\isamarkupfalse% +\isanewline +\ \ update\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}linorder\ {\isasymtimes}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}\isanewline +\ \ \ \ if\ it\ {\isacharequal}\ key\ then\ Leaf\ key\ entry\isanewline +\ \ \ \ \ \ else\ if\ it\ {\isasymle}\ key\isanewline +\ \ \ \ \ \ then\ Branch\ {\isacharparenleft}Leaf\ it\ entry{\isacharparenright}\ it\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\isanewline +\ \ \ \ \ \ else\ Branch\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ it\ {\isacharparenleft}Leaf\ it\ entry{\isacharparenright}\isanewline +\ \ \ {\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ t{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}\isanewline +\ \ \ \ if\ it\ {\isasymle}\ key\isanewline +\ \ \ \ \ \ then\ {\isacharparenleft}Branch\ {\isacharparenleft}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ t{\isadigit{1}}{\isacharparenright}\ key\ t{\isadigit{2}}{\isacharparenright}\isanewline +\ \ \ \ \ \ else\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ {\isacharparenleft}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ t{\isadigit{2}}{\isacharparenright}{\isacharparenright}\isanewline +\ \ \ {\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent For testing purpose, we define a small example + using natural numbers \isa{nat} (which are a \isa{linorder}) + as keys and list of nats as values:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ example\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat{\isacharcomma}\ nat\ list{\isacharparenright}\ searchtree{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}example\ {\isacharequal}\ update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\ {\isacharparenleft}update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \ \ {\isacharparenleft}update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\ {\isacharparenleft}Leaf\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent Then we generate code% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ example\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}tree{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent which looks like: + \lstsml{Thy/examples/tree.ML}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Code generation concepts and process \label{sec:concept}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{figure}[h] + \centering + \includegraphics[width=0.7\textwidth]{codegen_process} + \caption{code generator -- processing overview} + \label{fig:process} + \end{figure} + + The code generator employs a notion of executability + for three foundational executable ingredients known + from functional programming: + \emph{defining equations}, \emph{datatypes}, and + \emph{type classes}. A defining equation as a first approximation + is a theorem of the form \isa{f\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n\ {\isasymequiv}\ t} + (an equation headed by a constant \isa{f} with arguments + \isa{t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n} and right hand side \isa{t}). + Code generation aims to turn defining equations + into a functional program by running through + a process (see figure \ref{fig:process}): + + \begin{itemize} + + \item Out of the vast collection of theorems proven in a + \qn{theory}, a reasonable subset modeling + defining equations is \qn{selected}. + + \item On those selected theorems, certain + transformations are carried out + (\qn{preprocessing}). Their purpose is to turn theorems + representing non- or badly executable + specifications into equivalent but executable counterparts. + The result is a structured collection of \qn{code theorems}. + + \item These \qn{code theorems} then are \qn{translated} + into an Haskell-like intermediate + language. + + \item Finally, out of the intermediate language the final + code in the desired \qn{target language} is \qn{serialized}. + + \end{itemize} + + From these steps, only the two last are carried out + outside the logic; by keeping this layer as + thin as possible, the amount of code to trust is + kept to a minimum.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Basics \label{sec:basics}% +} +\isamarkuptrue% +% +\isamarkupsubsection{Invoking the code generator% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Thanks to a reasonable setup of the \isa{HOL} theories, in + most cases code generation proceeds without further ado:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\isanewline +\ \ fac\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ \ \ {\isachardoublequoteopen}fac\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}fac\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ Suc\ n\ {\isacharasterisk}\ fac\ n{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent This executable specification is now turned to SML code:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ fac\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}fac{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent The \isa{{\isasymEXPORTCODE}} command takes a space-separated list of + constants together with \qn{serialization directives} + These start with a \qn{target language} + identifier, followed by a file specification + where to write the generated code to. + + Internally, the defining equations for all selected + constants are taken, including any transitively required + constants, datatypes and classes, resulting in the following + code: + + \lstsml{Thy/examples/fac.ML} + + The code generator will complain when a required + ingredient does not provide a executable counterpart, + e.g.~generating code + for constants not yielding + a defining equation (e.g.~the Hilbert choice + operation \isa{SOME}):% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ pick{\isacharunderscore}some\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}pick{\isacharunderscore}some\ xs\ {\isacharequal}\ {\isacharparenleft}SOME\ x{\isachardot}\ x\ {\isasymin}\ set\ xs{\isacharparenright}{\isachardoublequoteclose}% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +\isanewline +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ pick{\isacharunderscore}some\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}fail{\isacharunderscore}const{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent will fail.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Theorem selection% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The list of all defining equations in a theory may be inspected + using the \isa{{\isasymPRINTCODESETUP}} command:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{print{\isacharunderscore}codesetup}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent which displays a table of constant with corresponding + defining equations (the additional stuff displayed + shall not bother us for the moment). + + The typical \isa{HOL} tools are already set up in a way that + function definitions introduced by \isa{{\isasymDEFINITION}}, + \isa{{\isasymPRIMREC}}, \isa{{\isasymFUN}}, + \isa{{\isasymFUNCTION}}, \isa{{\isasymCONSTDEFS}}, + \isa{{\isasymRECDEF}} are implicitly propagated + to this defining equation table. Specific theorems may be + selected using an attribute: \emph{code func}. As example, + a weight selector function:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\isanewline +\ \ pick\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ list\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}pick\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ n\ {\isacharequal}\ {\isacharparenleft}let\ {\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}\ {\isacharequal}\ x\ in\isanewline +\ \ \ \ if\ n\ {\isacharless}\ k\ then\ v\ else\ pick\ xs\ {\isacharparenleft}n\ {\isacharminus}\ k{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent We want to eliminate the explicit destruction + of \isa{x} to \isa{{\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}pick\ {\isacharparenleft}{\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}{\isacharhash}xs{\isacharparenright}\ n\ {\isacharequal}\ {\isacharparenleft}if\ n\ {\isacharless}\ k\ then\ v\ else\ pick\ xs\ {\isacharparenleft}n\ {\isacharminus}\ k{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ pick\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}pick{\isadigit{1}}{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent This theorem now is used for generating code: + + \lstsml{Thy/examples/pick1.ML} + + \noindent The policy is that \emph{default equations} stemming from + \isa{{\isasymDEFINITION}}, + \isa{{\isasymPRIMREC}}, \isa{{\isasymFUN}}, + \isa{{\isasymFUNCTION}}, \isa{{\isasymCONSTDEFS}}, + \isa{{\isasymRECDEF}} statements are discarded as soon as an + equation is explicitly selected by means of \emph{code func}. + Further applications of \emph{code func} add theorems incrementally, + but syntactic redundancies are implicitly dropped. For example, + using a modified version of the \isa{fac} function + as defining equation, the then redundant (since + syntactically subsumed) original defining equations + are dropped. + + \begin{warn} + The attributes \emph{code} and \emph{code del} + associated with the existing code generator also apply to + the new one: \emph{code} implies \emph{code func}, + and \emph{code del} implies \emph{code func del}. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Type classes% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Type classes enter the game via the Isar class package. + For a short introduction how to use it, see \cite{isabelle-classes}; + here we just illustrate its impact on code generation. + + In a target language, type classes may be represented + natively (as in the case of Haskell). For languages + like SML, they are implemented using \emph{dictionaries}. + Our following example specifies a class \qt{null}, + assigning to each of its inhabitants a \qt{null} value:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{class}\isamarkupfalse% +\ null\ {\isacharequal}\ type\ {\isacharplus}\isanewline +\ \ \isakeyword{fixes}\ null\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline +\isanewline +\isacommand{primrec}\isamarkupfalse% +\isanewline +\ \ head\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}null\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}head\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ null{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}head\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent We provide some instances for our \isa{null}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{instantiation}\isamarkupfalse% +\ option\ \isakeyword{and}\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}\ null\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ {\isachardoublequoteopen}null\ {\isacharequal}\ None{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ {\isachardoublequoteopen}null\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Constructing a dummy example:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ {\isachardoublequoteopen}dummy\ {\isacharequal}\ head\ {\isacharbrackleft}Some\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ None{\isacharbrackright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +Type classes offer a suitable occasion to introduce + the Haskell serializer. Its usage is almost the same + as SML, but, in accordance with conventions + some Haskell systems enforce, each module ends + up in a single file. The module hierarchy is reflected in + the file system, with root directory given as file specification.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ dummy\ \isakeyword{in}\ Haskell\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lsthaskell{Thy/examples/Codegen.hs} + \noindent (we have left out all other modules). + + \medskip + + The whole code in SML with explicit dictionary passing:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ dummy\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}class{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/class.ML} + + \medskip + + \noindent or in OCaml:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ dummy\ \isakeyword{in}\ OCaml\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}class{\isachardot}ocaml{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/class.ocaml} + + \medskip The explicit association of constants + to classes can be inspected using the \isa{{\isasymPRINTCLASSES}} + command.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Recipes and advanced topics \label{sec:advanced}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In this tutorial, we do not attempt to give an exhaustive + description of the code generator framework; instead, + we cast a light on advanced topics by introducing + them together with practically motivated examples. Concerning + further reading, see + + \begin{itemize} + + \item the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} + for exhaustive syntax diagrams. + \item or \cite{Haftmann-Nipkow:2007:codegen} which deals with foundational issues + of the code generator framework. + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Library theories \label{sec:library}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The \isa{HOL} \isa{Main} theory already provides a code + generator setup + which should be suitable for most applications. Common extensions + and modifications are available by certain theories of the \isa{HOL} + library; beside being useful in applications, they may serve + as a tutorial for customizing the code generator setup. + + \begin{description} + + \item[\isa{Code{\isacharunderscore}Integer}] represents \isa{HOL} integers by big + integer literals in target languages. + \item[\isa{Code{\isacharunderscore}Char}] represents \isa{HOL} characters by + character literals in target languages. + \item[\isa{Code{\isacharunderscore}Char{\isacharunderscore}chr}] like \isa{Code{\isacharunderscore}Char}, + but also offers treatment of character codes; includes + \isa{Code{\isacharunderscore}Integer}. + \item[\isa{Efficient{\isacharunderscore}Nat}] \label{eff_nat} implements natural numbers by integers, + which in general will result in higher efficency; pattern + matching with \isa{{\isadigit{0}}} / \isa{Suc} + is eliminated; includes \isa{Code{\isacharunderscore}Integer}. + \item[\isa{Code{\isacharunderscore}Index}] provides an additional datatype + \isa{index} which is mapped to target-language built-in integers. + Useful for code setups which involve e.g. indexing of + target-language arrays. + \item[\isa{Code{\isacharunderscore}Message}] provides an additional datatype + \isa{message{\isacharunderscore}string} which is isomorphic to strings; + \isa{message{\isacharunderscore}string}s are mapped to target-language strings. + Useful for code setups which involve e.g. printing (error) messages. + + \end{description} + + \begin{warn} + When importing any of these theories, they should form the last + items in an import list. Since these theories adapt the + code generator setup in a non-conservative fashion, + strange effects may occur otherwise. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Preprocessing% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Before selected function theorems are turned into abstract + code, a chain of definitional transformation steps is carried + out: \emph{preprocessing}. In essence, the preprocessor + consists of two components: a \emph{simpset} and \emph{function transformers}. + + The \emph{simpset} allows to employ the full generality of the Isabelle + simplifier. Due to the interpretation of theorems + as defining equations, rewrites are applied to the right + hand side and the arguments of the left hand side of an + equation, but never to the constant heading the left hand side. + An important special case are \emph{inline theorems} which may be + declared an undeclared using the + \emph{code inline} or \emph{code inline del} attribute respectively. + Some common applications:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{itemize} +% +\begin{isamarkuptext}% +\item replacing non-executable constructs by executable ones:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}x\ {\isasymin}\ set\ xs\ {\isasymlongleftrightarrow}\ x\ mem\ xs{\isachardoublequoteclose}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\item eliminating superfluous constants:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\item replacing executable but inconvenient constructs:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongleftrightarrow}\ List{\isachardot}null\ xs{\isachardoublequoteclose}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\end{itemize} +% +\begin{isamarkuptext}% +\emph{Function transformers} provide a very general interface, + transforming a list of function theorems to another + list of function theorems, provided that neither the heading + constant nor its type change. The \isa{{\isadigit{0}}} / \isa{Suc} + pattern elimination implemented in + theory \isa{Efficient{\isacharunderscore}Nat} (see \secref{eff_nat}) uses this + interface. + + \noindent The current setup of the preprocessor may be inspected using + the \isa{{\isasymPRINTCODESETUP}} command. + + \begin{warn} + The attribute \emph{code unfold} + associated with the existing code generator also applies to + the new one: \emph{code unfold} implies \emph{code inline}. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Concerning operational equality% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Surely you have already noticed how equality is treated + by the code generator:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\isanewline +\ \ collect{\isacharunderscore}duplicates\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ \ \ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharparenleft}z{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ z\ {\isasymin}\ set\ xs\isanewline +\ \ \ \ \ \ then\ if\ z\ {\isasymin}\ set\ ys\isanewline +\ \ \ \ \ \ \ \ then\ collect{\isacharunderscore}duplicates\ xs\ ys\ zs\isanewline +\ \ \ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ xs\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs\isanewline +\ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ {\isacharparenleft}z{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +The membership test during preprocessing is rewritten, + resulting in \isa{op\ mem}, which itself + performs an explicit equality check.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ collect{\isacharunderscore}duplicates\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}collect{\isacharunderscore}duplicates{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/collect_duplicates.ML}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +Obviously, polymorphic equality is implemented the Haskell + way using a type class. How is this achieved? HOL introduces + an explicit class \isa{eq} with a corresponding operation + \isa{eq{\isacharunderscore}class{\isachardot}eq} such that \isa{eq{\isacharunderscore}class{\isachardot}eq\ x\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ y{\isacharparenright}}. + The preprocessing framework does the rest. + For datatypes, instances of \isa{eq} are implicitly derived + when possible. For other types, you may instantiate \isa{eq} + manually like any other type class. + + Though this \isa{eq} class is designed to get rarely in + the way, a subtlety + enters the stage when definitions of overloaded constants + are dependent on operational equality. For example, let + us define a lexicographic ordering on tuples:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{instantiation}\isamarkupfalse% +\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}ord{\isacharcomma}\ ord{\isacharparenright}\ ord\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ {\isacharbrackleft}code\ func\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isacharless}\ p{\isadigit{2}}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}let\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline +\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ {\isacharbrackleft}code\ func\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isasymle}\ p{\isadigit{2}}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}let\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline +\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ ord{\isacharunderscore}prod\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}ord{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}ord{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp{\isacharunderscore}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Then code generation will fail. Why? The definition + of \isa{op\ {\isasymle}} depends on equality on both arguments, + which are polymorphic and impose an additional \isa{eq} + class constraint, thus violating the type discipline + for class operations. + + The solution is to add \isa{eq} explicitly to the first sort arguments in the + code theorems:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ ord{\isacharunderscore}prod{\isacharunderscore}code\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}ord{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}ord{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ ord{\isacharunderscore}prod\ \isacommand{by}\isamarkupfalse% +\ rule{\isacharplus}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Then code generation succeeds:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}eq{\isacharcomma}\ ord{\isacharbraceright}\ {\isasymtimes}\ {\isacharprime}b{\isasymColon}ord\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}lexicographic{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/lexicographic.ML}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +In general, code theorems for overloaded constants may have more + restrictive sort constraints than the underlying instance relation + between class and type constructor as long as the whole system of + constraints is coregular; code theorems violating coregularity + are rejected immediately. Consequently, it might be necessary + to delete disturbing theorems in the code theorem table, + as we have done here with the original definitions \isa{less{\isacharunderscore}prod{\isacharunderscore}def} + and \isa{less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def}. + + In some cases, the automatically derived defining equations + for equality on a particular type may not be appropriate. + As example, watch the following datatype representing + monomorphic parametric types (where type constructors + are referred to by natural numbers):% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ monotype\ {\isacharequal}\ Mono\ nat\ {\isachardoublequoteopen}monotype\ list{\isachardoublequoteclose}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Then code generation for SML would fail with a message + that the generated code conains illegal mutual dependencies: + the theorem \isa{Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymequiv}\ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ typargs{\isadigit{1}}\ {\isacharequal}\ typargs{\isadigit{2}}} already requires the + instance \isa{monotype\ {\isasymColon}\ eq}, which itself requires + \isa{Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymequiv}\ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ typargs{\isadigit{1}}\ {\isacharequal}\ typargs{\isadigit{2}}}; Haskell has no problem with mutually + recursive \isa{instance} and \isa{function} definitions, + but the SML serializer does not support this. + + In such cases, you have to provide you own equality equations + involving auxiliary constants. In our case, + \isa{list{\isacharunderscore}all{\isadigit{2}}} can do the job:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ monotype{\isacharunderscore}eq{\isacharunderscore}list{\isacharunderscore}all{\isadigit{2}}\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ \ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ list{\isacharunderscore}all{\isadigit{2}}\ {\isacharparenleft}op\ {\isacharequal}{\isacharparenright}\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ list{\isacharunderscore}all{\isadigit{2}}{\isacharunderscore}eq\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +does not depend on instance \isa{monotype\ {\isasymColon}\ eq}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isacharequal}\ {\isacharcolon}{\isacharcolon}\ monotype\ {\isasymRightarrow}\ monotype\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}monotype{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/monotype.ML}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Programs as sets of theorems% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +As told in \secref{sec:concept}, code generation is based + on a structured collection of code theorems. + For explorative purpose, this collection + may be inspected using the \isa{{\isasymCODETHMS}} command:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{code{\isacharunderscore}thms}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ mod\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent prints a table with \emph{all} defining equations + for \isa{op\ mod}, including + \emph{all} defining equations those equations depend + on recursivly. \isa{{\isasymCODETHMS}} provides a convenient + mechanism to inspect the impact of a preprocessor setup + on defining equations. + + Similarly, the \isa{{\isasymCODEDEPS}} command shows a graph + visualizing dependencies between defining equations.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Constructor sets for datatypes% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Conceptually, any datatype is spanned by a set of + \emph{constructors} of type \isa{{\isasymtau}\ {\isacharequal}\ {\isasymdots}\ {\isasymRightarrow}\ {\isasymkappa}\ {\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymdots}\ {\isasymalpha}\isactrlisub n} + where \isa{{\isacharbraceleft}{\isasymalpha}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlisub n{\isacharbraceright}} is excactly the set of \emph{all} + type variables in \isa{{\isasymtau}}. The HOL datatype package + by default registers any new datatype in the table + of datatypes, which may be inspected using + the \isa{{\isasymPRINTCODESETUP}} command. + + In some cases, it may be convenient to alter or + extend this table; as an example, we will develope an alternative + representation of natural numbers as binary digits, whose + size does increase logarithmically with its value, not linear + \footnote{Indeed, the \isa{Efficient{\isacharunderscore}Nat} theory (see \ref{eff_nat}) + does something similar}. First, the digit representation:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{definition}\isamarkupfalse% +\ Dig{\isadigit{0}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ n\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ n{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ Dig{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ n\ {\isacharequal}\ Suc\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent We will use these two ">digits"< to represent natural numbers + in binary digits, e.g.:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isadigit{4}}{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}Dig{\isadigit{1}}\ {\isacharparenleft}Dig{\isadigit{0}}\ {\isacharparenleft}Dig{\isadigit{1}}\ {\isacharparenleft}Dig{\isadigit{0}}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ Dig{\isadigit{0}}{\isacharunderscore}def\ Dig{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Of course we also have to provide proper code equations for + the operations, e.g. \isa{op\ {\isacharplus}}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ plus{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isadigit{0}}\ {\isacharplus}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}m\ {\isacharplus}\ {\isadigit{0}}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ n{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ {\isadigit{1}}\ {\isacharequal}\ Dig{\isadigit{1}}\ m{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ {\isadigit{1}}\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ Dig{\isadigit{0}}{\isacharunderscore}def\ Dig{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent We then instruct the code generator to view \isa{{\isadigit{0}}}, + \isa{{\isadigit{1}}}, \isa{Dig{\isadigit{0}}} and \isa{Dig{\isadigit{1}}} as + datatype constructors:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isadigit{0}}{\isasymColon}nat{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isadigit{1}}{\isasymColon}nat{\isachardoublequoteclose}\ Dig{\isadigit{0}}\ Dig{\isadigit{1}}% +\begin{isamarkuptext}% +\noindent For the former constructor \isa{Suc}, we provide a code + equation and remove some parts of the default code generator setup + which are an obstacle here:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ Suc{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}Suc\ n\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{declare}\isamarkupfalse% +\ One{\isacharunderscore}nat{\isacharunderscore}def\ {\isacharbrackleft}code\ inline\ del{\isacharbrackright}\isanewline +\isacommand{declare}\isamarkupfalse% +\ add{\isacharunderscore}Suc{\isacharunderscore}shift\ {\isacharbrackleft}code\ func\ del{\isacharbrackright}% +\begin{isamarkuptext}% +\noindent This yields the following code:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isacharplus}\ {\isasymColon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}nat{\isacharunderscore}binary{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/nat_binary.ML}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\medskip + + From this example, it can be easily glimpsed that using own constructor sets + is a little delicate since it changes the set of valid patterns for values + of that type. Without going into much detail, here some practical hints: + + \begin{itemize} + \item When changing the constuctor set for datatypes, take care to + provide an alternative for the \isa{case} combinator (e.g. by replacing + it using the preprocessor). + \item Values in the target language need not to be normalized -- different + values in the target language may represent the same value in the + logic (e.g. \isa{Dig{\isadigit{1}}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}}). + \item Usually, a good methodology to deal with the subleties of pattern + matching is to see the type as an abstract type: provide a set + of operations which operate on the concrete representation of the type, + and derive further operations by combinations of these primitive ones, + without relying on a particular representation. + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isataginvisible +\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isadigit{0}}{\isacharcolon}{\isacharcolon}nat{\isachardoublequoteclose}\ Suc\isanewline +\isacommand{declare}\isamarkupfalse% +\ plus{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func\ del{\isacharbrackright}\isanewline +\isacommand{declare}\isamarkupfalse% +\ One{\isacharunderscore}nat{\isacharunderscore}def\ {\isacharbrackleft}code\ inline{\isacharbrackright}\isanewline +\isacommand{declare}\isamarkupfalse% +\ add{\isacharunderscore}Suc{\isacharunderscore}shift\ {\isacharbrackleft}code\ func{\isacharbrackright}\ \isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isadigit{0}}\ {\isacharplus}\ n\ {\isacharequal}\ {\isacharparenleft}n\ {\isasymColon}\ nat{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp% +\endisataginvisible +{\isafoldinvisible}% +% +\isadeliminvisible +% +\endisadeliminvisible +% +\isamarkupsubsection{Customizing serialization% +} +\isamarkuptrue% +% +\isamarkupsubsubsection{Basics% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Consider the following function and its corresponding + SML code:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\isanewline +\ \ in{\isacharunderscore}interval\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymtimes}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}in{\isacharunderscore}interval\ {\isacharparenleft}k{\isacharcomma}\ l{\isacharparenright}\ n\ {\isasymlongleftrightarrow}\ k\ {\isasymle}\ n\ {\isasymand}\ n\ {\isasymle}\ l{\isachardoublequoteclose}% +\isadelimtt +% +\endisadelimtt +% +\isatagtt +% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ in{\isacharunderscore}interval\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}literal{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/bool_literal.ML} + + \noindent Though this is correct code, it is a little bit unsatisfactory: + boolean values and operators are materialized as distinguished + entities with have nothing to do with the SML-builtin notion + of \qt{bool}. This results in less readable code; + additionally, eager evaluation may cause programs to + loop or break which would perfectly terminate when + the existing SML \qt{bool} would be used. To map + the HOL \qt{bool} on SML \qt{bool}, we may use + \qn{custom serializations}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtt +% +\endisadelimtt +% +\isatagtt +\isacommand{code{\isacharunderscore}type}\isamarkupfalse% +\ bool\isanewline +\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}bool{\isachardoublequoteclose}{\isacharparenright}\isanewline +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ True\ \isakeyword{and}\ False\ \isakeyword{and}\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline +\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}true{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}false{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}{\isacharunderscore}\ andalso\ {\isacharunderscore}{\isachardoublequoteclose}{\isacharparenright}% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +% +\begin{isamarkuptext}% +The \isa{{\isasymCODETYPE}} commad takes a type constructor + as arguments together with a list of custom serializations. + Each custom serialization starts with a target language + identifier followed by an expression, which during + code serialization is inserted whenever the type constructor + would occur. For constants, \isa{{\isasymCODECONST}} implements + the corresponding mechanism. Each ``\verb|_|'' in + a serialization expression is treated as a placeholder + for the type constructor's (the constant's) arguments.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ in{\isacharunderscore}interval\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}mlbool{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/bool_mlbool.ML} + + \noindent This still is not perfect: the parentheses + around the \qt{andalso} expression are superfluous. + Though the serializer + by no means attempts to imitate the rich Isabelle syntax + framework, it provides some common idioms, notably + associative infixes with precedences which may be used here:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtt +% +\endisadelimtt +% +\isatagtt +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline +\ \ {\isacharparenleft}SML\ \isakeyword{infixl}\ {\isadigit{1}}\ {\isachardoublequoteopen}andalso{\isachardoublequoteclose}{\isacharparenright}% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +\isanewline +\isanewline +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ in{\isacharunderscore}interval\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}infix{\isachardot}ML{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\lstsml{Thy/examples/bool_infix.ML} + + \medskip + + Next, we try to map HOL pairs to SML pairs, using the + infix ``\verb|*|'' type constructor and parentheses:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtt +% +\endisadelimtt +% +\isatagtt +\isacommand{code{\isacharunderscore}type}\isamarkupfalse% +\ {\isacharasterisk}\isanewline +\ \ {\isacharparenleft}SML\ \isakeyword{infix}\ {\isadigit{2}}\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ Pair\isanewline +\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}{\isacharbang}{\isacharparenleft}{\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharcomma}{\isacharslash}\ {\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +% +\begin{isamarkuptext}% +The initial bang ``\verb|!|'' tells the serializer to never put + parentheses around the whole expression (they are already present), + while the parentheses around argument place holders + tell not to put parentheses around the arguments. + The slash ``\verb|/|'' (followed by arbitrary white space) + inserts a space which may be used as a break if necessary + during pretty printing. + + These examples give a glimpse what mechanisms + custom serializations provide; however their usage + requires careful thinking in order not to introduce + inconsistencies -- or, in other words: + custom serializations are completely axiomatic. + + A further noteworthy details is that any special + character in a custom serialization may be quoted + using ``\verb|'|''; thus, in + ``\verb|fn '_ => _|'' the first + ``\verb|_|'' is a proper underscore while the + second ``\verb|_|'' is a placeholder. + + The HOL theories provide further + examples for custom serializations.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Haskell serialization% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +For convenience, the default + HOL setup for Haskell maps the \isa{eq} class to + its counterpart in Haskell, giving custom serializations + for the class (\isa{{\isasymCODECLASS}}) and its operation:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtt +% +\endisadelimtt +% +\isatagtt +\isacommand{code{\isacharunderscore}class}\isamarkupfalse% +\ eq\isanewline +\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Eq{\isachardoublequoteclose}\ \isakeyword{where}\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\ {\isasymequiv}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharequal}{\isacharequal}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\isanewline +\isacommand{code{\isacharunderscore}const}\isamarkupfalse% +\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\isanewline +\ \ {\isacharparenleft}Haskell\ \isakeyword{infixl}\ {\isadigit{4}}\ {\isachardoublequoteopen}{\isacharequal}{\isacharequal}{\isachardoublequoteclose}{\isacharparenright}% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +% +\begin{isamarkuptext}% +A problem now occurs whenever a type which + is an instance of \isa{eq} in HOL is mapped + on a Haskell-builtin type which is also an instance + of Haskell \isa{Eq}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{typedecl}\isamarkupfalse% +\ bar\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}x{\isasymColon}bar{\isacharparenright}\ y\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ default\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq{\isacharunderscore}bar{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +% +\isadelimtt +\isanewline +% +\endisadelimtt +% +\isatagtt +\isacommand{code{\isacharunderscore}type}\isamarkupfalse% +\ bar\isanewline +\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Integer{\isachardoublequoteclose}{\isacharparenright}% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +% +\begin{isamarkuptext}% +The code generator would produce + an additional instance, which of course is rejected. + To suppress this additional instance, use + \isa{{\isasymCODEINSTANCE}}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtt +% +\endisadelimtt +% +\isatagtt +\isacommand{code{\isacharunderscore}instance}\isamarkupfalse% +\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline +\ \ {\isacharparenleft}Haskell\ {\isacharminus}{\isacharparenright}% +\endisatagtt +{\isafoldtt}% +% +\isadelimtt +% +\endisadelimtt +% +\isamarkupsubsubsection{Pretty printing% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The serializer provides ML interfaces to set up + pretty serializations for expressions like lists, numerals + and characters; these are + monolithic stubs and should only be used with the + theories introduced in \secref{sec:library}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Cyclic module dependencies% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Sometimes the awkward situation occurs that dependencies + between definitions introduce cyclic dependencies + between modules, which in the Haskell world leaves + you to the mercy of the Haskell implementation you are using, + while for SML code generation is not possible. + + A solution is to declare module names explicitly. + Let use assume the three cyclically dependent + modules are named \emph{A}, \emph{B} and \emph{C}. + Then, by stating% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{code{\isacharunderscore}modulename}\isamarkupfalse% +\ SML\isanewline +\ \ A\ ABC\isanewline +\ \ B\ ABC\isanewline +\ \ C\ ABC% +\begin{isamarkuptext}% +we explicitly map all those modules on \emph{ABC}, + resulting in an ad-hoc merge of this three modules + at serialization time.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Incremental code generation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Code generation is \emph{incremental}: theorems + and abstract intermediate code are cached and extended on demand. + The cache may be partially or fully dropped if the underlying + executable content of the theory changes. + Implementation of caching is supposed to transparently + hid away the details from the user. Anyway, caching + reaches the surface by using a slightly more general form + of the \isa{{\isasymCODETHMS}}, \isa{{\isasymCODEDEPS}} + and \isa{{\isasymEXPORTCODE}} commands: the list of constants + may be omitted. Then, all constants with code theorems + in the current cache are referred to.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{ML interfaces \label{sec:ml}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Since the code generator framework not only aims to provide + a nice Isar interface but also to form a base for + code-generation-based applications, here a short + description of the most important ML interfaces.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Executable theory content: \isa{Code}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +This Pure module implements the core notions of + executable content of a theory.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Managing executable content% +} +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexml{Code.add\_func}\verb|Code.add_func: thm -> theory -> theory| \\ + \indexml{Code.del\_func}\verb|Code.del_func: thm -> theory -> theory| \\ + \indexml{Code.add\_funcl}\verb|Code.add_funcl: string * thm list Susp.T -> theory -> theory| \\ + \indexml{Code.map\_pre}\verb|Code.map_pre: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\ + \indexml{Code.map\_post}\verb|Code.map_post: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\ + \indexml{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> thm list -> thm list option)|\isasep\isanewline% +\verb| -> theory -> theory| \\ + \indexml{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\ + \indexml{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\ + \indexml{Code.get\_datatype}\verb|Code.get_datatype: theory -> string|\isasep\isanewline% +\verb| -> (string * sort) list * (string * typ list) list| \\ + \indexml{Code.get\_datatype\_of\_constr}\verb|Code.get_datatype_of_constr: theory -> string -> string option| + \end{mldecls} + + \begin{description} + + \item \verb|Code.add_func|~\isa{thm}~\isa{thy} adds function + theorem \isa{thm} to executable content. + + \item \verb|Code.del_func|~\isa{thm}~\isa{thy} removes function + theorem \isa{thm} from executable content, if present. + + \item \verb|Code.add_funcl|~\isa{{\isacharparenleft}const{\isacharcomma}\ lthms{\isacharparenright}}~\isa{thy} adds + suspended defining equations \isa{lthms} for constant + \isa{const} to executable content. + + \item \verb|Code.map_pre|~\isa{f}~\isa{thy} changes + the preprocessor simpset. + + \item \verb|Code.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds + function transformer \isa{f} (named \isa{name}) to executable content; + \isa{f} is a transformer of the defining equations belonging + to a certain function definition, depending on the + current theory context. Returning \isa{NONE} indicates that no + transformation took place; otherwise, the whole process will be iterated + with the new defining equations. + + \item \verb|Code.del_functrans|~\isa{name}~\isa{thy} removes + function transformer named \isa{name} from executable content. + + \item \verb|Code.add_datatype|~\isa{cs}~\isa{thy} adds + a datatype to executable content, with generation + set \isa{cs}. + + \item \verb|Code.get_datatype_of_constr|~\isa{thy}~\isa{const} + returns type constructor corresponding to + constructor \isa{const}; returns \isa{NONE} + if \isa{const} is no constructor. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Auxiliary% +} +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexml{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\ + \indexml{Code\_Unit.head\_func}\verb|Code_Unit.head_func: thm -> string * ((string * sort) list * typ)| \\ + \indexml{Code\_Unit.rewrite\_func}\verb|Code_Unit.rewrite_func: MetaSimplifier.simpset -> thm -> thm| \\ + \end{mldecls} + + \begin{description} + + \item \verb|Code_Unit.read_const|~\isa{thy}~\isa{s} + reads a constant as a concrete term expression \isa{s}. + + \item \verb|Code_Unit.head_func|~\isa{thm} + extracts the constant and its type from a defining equation \isa{thm}. + + \item \verb|Code_Unit.rewrite_func|~\isa{ss}~\isa{thm} + rewrites a defining equation \isa{thm} with a simpset \isa{ss}; + only arguments and right hand side are rewritten, + not the head of the defining equation. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Implementing code generator applications% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Implementing code generator applications on top + of the framework set out so far usually not only + involves using those primitive interfaces + but also storing code-dependent data and various + other things. + + \begin{warn} + Some interfaces discussed here have not reached + a final state yet. + Changes likely to occur in future. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Data depending on the theory's executable content% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Due to incrementality of code generation, changes in the + theory's executable content have to be propagated in a + certain fashion. Additionally, such changes may occur + not only during theory extension but also during theory + merge, which is a little bit nasty from an implementation + point of view. The framework provides a solution + to this technical challenge by providing a functorial + data slot \verb|CodeDataFun|; on instantiation + of this functor, the following types and operations + are required: + + \medskip + \begin{tabular}{l} + \isa{type\ T} \\ + \isa{val\ empty{\isacharcolon}\ T} \\ + \isa{val\ merge{\isacharcolon}\ Pretty{\isachardot}pp\ {\isasymrightarrow}\ T\ {\isacharasterisk}\ T\ {\isasymrightarrow}\ T} \\ + \isa{val\ purge{\isacharcolon}\ theory\ option\ {\isasymrightarrow}\ CodeUnit{\isachardot}const\ list\ option\ {\isasymrightarrow}\ T\ {\isasymrightarrow}\ T} + \end{tabular} + + \begin{description} + + \item \isa{T} the type of data to store. + + \item \isa{empty} initial (empty) data. + + \item \isa{merge} merging two data slots. + + \item \isa{purge}~\isa{thy}~\isa{consts} propagates changes in executable content; + if possible, the current theory context is handed over + as argument \isa{thy} (if there is no current theory context (e.g.~during + theory merge, \verb|NONE|); \isa{consts} indicates the kind + of change: \verb|NONE| stands for a fundamental change + which invalidates any existing code, \isa{SOME\ consts} + hints that executable content for constants \isa{consts} + has changed. + + \end{description} + + An instance of \verb|CodeDataFun| provides the following + interface: + + \medskip + \begin{tabular}{l} + \isa{get{\isacharcolon}\ theory\ {\isasymrightarrow}\ T} \\ + \isa{change{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ T{\isacharparenright}\ {\isasymrightarrow}\ T} \\ + \isa{change{\isacharunderscore}yield{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T{\isacharparenright}\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T} + \end{tabular} + + \begin{description} + + \item \isa{get} retrieval of the current data. + + \item \isa{change} update of current data (cached!) + by giving a continuation. + + \item \isa{change{\isacharunderscore}yield} update with side result. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\emph{Happy proving, happy hacking!}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/document/Further.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/document/Further.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,246 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Further}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Further\isanewline +\isakeyword{imports}\ Setup\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsection{Further issues \label{sec:further}% +} +\isamarkuptrue% +% +\isamarkupsubsection{Further reading% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Do dive deeper into the issue of code generation, you should visit + the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} which + contains exhaustive syntax diagrams.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Modules% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +When invoking the \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command it is possible to leave + out the \hyperlink{keyword.module-name}{\mbox{\isa{\isakeyword{module{\isacharunderscore}name}}}} part; then code is distributed over + different modules, where the module name space roughly is induced + by the \isa{Isabelle} theory name space. + + Then sometimes the awkward situation occurs that dependencies between + definitions introduce cyclic dependencies between modules, which in the + \isa{Haskell} world leaves you to the mercy of the \isa{Haskell} implementation + you are using, while for \isa{SML}/\isa{OCaml} code generation is not possible. + + A solution is to declare module names explicitly. + Let use assume the three cyclically dependent + modules are named \emph{A}, \emph{B} and \emph{C}. + Then, by stating% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{code{\isacharunderscore}modulename}\isamarkupfalse% +\ SML\isanewline +\ \ A\ ABC\isanewline +\ \ B\ ABC\isanewline +\ \ C\ ABC% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +we explicitly map all those modules on \emph{ABC}, + resulting in an ad-hoc merge of this three modules + at serialisation time.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Evaluation oracle% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Code generation may also be used to \emph{evaluate} expressions + (using \isa{SML} as target language of course). + For instance, the \hyperlink{command.value}{\mbox{\isa{\isacommand{value}}}} allows to reduce an expression to a + normal form with respect to the underlying code equations:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{value}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharslash}\ {\isacharparenleft}{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ rat{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent will display \isa{{\isadigit{7}}\ {\isacharslash}\ {\isadigit{2}}}. + + The \hyperlink{method.eval}{\mbox{\isa{eval}}} method tries to reduce a goal by code generation to \isa{True} + and solves it in that case, but fails otherwise:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharslash}\ {\isacharparenleft}{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ rat{\isacharparenright}\ {\isacharequal}\ {\isadigit{7}}\ {\isacharslash}\ {\isadigit{2}}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ eval% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The soundness of the \hyperlink{method.eval}{\mbox{\isa{eval}}} method depends crucially + on the correctness of the code generator; this is one of the reasons + why you should not use adaption (see \secref{sec:adaption}) frivolously.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Code antiquotation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In scenarios involving techniques like reflection it is quite common + that code generated from a theory forms the basis for implementing + a proof procedure in \isa{SML}. To facilitate interfacing of generated code + with system code, the code generator provides a \isa{code} antiquotation:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{datatype}\isamarkupfalse% +\ form\ {\isacharequal}\ T\ {\isacharbar}\ F\ {\isacharbar}\ And\ form\ form\ {\isacharbar}\ Or\ form\ form% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +\isanewline +% +\isadelimquotett +% +\endisadelimquotett +% +\isatagquotett +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ fun\ eval{\isacharunderscore}form\ % +\isaantiq +code\ T% +\endisaantiq +\ {\isacharequal}\ true\isanewline +\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ % +\isaantiq +code\ F% +\endisaantiq +\ {\isacharequal}\ false\isanewline +\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ {\isacharparenleft}% +\isaantiq +code\ And% +\endisaantiq +\ {\isacharparenleft}p{\isacharcomma}\ q{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ \ \ \ eval{\isacharunderscore}form\ p\ andalso\ eval{\isacharunderscore}form\ q\isanewline +\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ {\isacharparenleft}% +\isaantiq +code\ Or% +\endisaantiq +\ {\isacharparenleft}p{\isacharcomma}\ q{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ \ \ \ eval{\isacharunderscore}form\ p\ orelse\ eval{\isacharunderscore}form\ q{\isacharsemicolon}\isanewline +{\isacharverbatimclose}% +\endisatagquotett +{\isafoldquotett}% +% +\isadelimquotett +% +\endisadelimquotett +% +\begin{isamarkuptext}% +\noindent \isa{code} takes as argument the name of a constant; after the + whole \isa{SML} is read, the necessary code is generated transparently + and the corresponding constant names are inserted. This technique also + allows to use pattern matching on constructors stemming from compiled + \isa{datatypes}. + + For a less simplistic example, theory \hyperlink{theory.Ferrack}{\mbox{\isa{Ferrack}}} is + a good reference.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Imperative data structures% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +If you consider imperative data structures as inevitable for a specific + application, you should consider + \emph{Imperative Functional Programming with Isabelle/HOL} + (\cite{bulwahn-et-al:2008:imperative}); + the framework described there is available in theory \hyperlink{theory.Imperative-HOL}{\mbox{\isa{Imperative{\isacharunderscore}HOL}}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/document/Introduction.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/document/Introduction.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,386 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Introduction}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Introduction\isanewline +\isakeyword{imports}\ Setup\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsection{Introduction and Overview% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +This tutorial introduces a generic code generator for the + \isa{Isabelle} system. + Generic in the sense that the + \qn{target language} for which code shall ultimately be + generated is not fixed but may be an arbitrary state-of-the-art + functional programming language (currently, the implementation + supports \isa{SML} \cite{SML}, \isa{OCaml} \cite{OCaml} and \isa{Haskell} + \cite{haskell-revised-report}). + + Conceptually the code generator framework is part + of Isabelle's \hyperlink{theory.Pure}{\mbox{\isa{Pure}}} meta logic framework; the logic + \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} which is an extension of \hyperlink{theory.Pure}{\mbox{\isa{Pure}}} + already comes with a reasonable framework setup and thus provides + a good working horse for raising code-generation-driven + applications. So, we assume some familiarity and experience + with the ingredients of the \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} distribution theories. + (see also \cite{isa-tutorial}). + + The code generator aims to be usable with no further ado + in most cases while allowing for detailed customisation. + This manifests in the structure of this tutorial: after a short + conceptual introduction with an example (\secref{sec:intro}), + we discuss the generic customisation facilities (\secref{sec:program}). + A further section (\secref{sec:adaption}) is dedicated to the matter of + \qn{adaption} to specific target language environments. After some + further issues (\secref{sec:further}) we conclude with an overview + of some ML programming interfaces (\secref{sec:ml}). + + \begin{warn} + Ultimately, the code generator which this tutorial deals with + is supposed to replace the existing code generator + by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. + So, for the moment, there are two distinct code generators + in Isabelle. In case of ambiguity, we will refer to the framework + described here as \isa{generic\ code\ generator}, to the + other as \isa{SML\ code\ generator}. + Also note that while the framework itself is + object-logic independent, only \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} provides a reasonable + framework setup. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Code generation via shallow embedding \label{sec:intro}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The key concept for understanding \isa{Isabelle}'s code generation is + \emph{shallow embedding}, i.e.~logical entities like constants, types and + classes are identified with corresponding concepts in the target language. + + Inside \hyperlink{theory.HOL}{\mbox{\isa{HOL}}}, the \hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}} and + \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}/\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}/\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}} declarations form + the core of a functional programming language. The default code generator setup + allows to turn those into functional programs immediately. + This means that \qt{naive} code generation can proceed without further ado. + For example, here a simple \qt{implementation} of amortised queues:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{datatype}\isamarkupfalse% +\ {\isacharprime}a\ queue\ {\isacharequal}\ AQueue\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ empty\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{primrec}\isamarkupfalse% +\ enqueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ AQueue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ ys{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{fun}\isamarkupfalse% +\ dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ option\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ \ {\isacharparenleft}case\ rev\ xs\ of\ y\ {\isacharhash}\ ys\ {\isasymRightarrow}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ ys{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Then we can generate code e.g.~for \isa{SML} as follows:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ empty\ dequeue\ enqueue\ \isakeyword{in}\ SML\isanewline +\ \ \isakeyword{module{\isacharunderscore}name}\ Example\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}example{\isachardot}ML{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent resulting in the following code:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun foldl f a [] = a\\ +\hspace*{0pt} ~| foldl f a (x ::~xs) = foldl f (f a x) xs;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun rev xs = foldl (fn xsa => fn x => x ::~xsa) [] xs;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun list{\char95}case f1 f2 (a ::~lista) = f2 a lista\\ +\hspace*{0pt} ~| list{\char95}case f1 f2 [] = f1;\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype 'a queue = AQueue of 'a list * 'a list;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val empty :~'a queue = AQueue ([],~[])\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun dequeue (AQueue ([],~[])) = (NONE,~AQueue ([],~[]))\\ +\hspace*{0pt} ~| dequeue (AQueue (xs,~y ::~ys)) = (SOME y,~AQueue (xs,~ys))\\ +\hspace*{0pt} ~| dequeue (AQueue (v ::~va,~[])) =\\ +\hspace*{0pt} ~~~let\\ +\hspace*{0pt} ~~~~~val y ::~ys = rev (v ::~va);\\ +\hspace*{0pt} ~~~in\\ +\hspace*{0pt} ~~~~~(SOME y,~AQueue ([],~ys))\\ +\hspace*{0pt} ~~~end;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun enqueue x (AQueue (xs,~ys)) = AQueue (x ::~xs,~ys);\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command takes a space-separated list of + constants for which code shall be generated; anything else needed for those + is added implicitly. Then follows a target language identifier + (\isa{SML}, \isa{OCaml} or \isa{Haskell}) and a freely chosen module name. + A file name denotes the destination to store the generated code. Note that + the semantics of the destination depends on the target language: for + \isa{SML} and \isa{OCaml} it denotes a \emph{file}, for \isa{Haskell} + it denotes a \emph{directory} where a file named as the module name + (with extension \isa{{\isachardot}hs}) is written:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{export{\isacharunderscore}code}\isamarkupfalse% +\ empty\ dequeue\ enqueue\ \isakeyword{in}\ Haskell\isanewline +\ \ \isakeyword{module{\isacharunderscore}name}\ Example\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This is how the corresponding code in \isa{Haskell} looks like:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}module Example where {\char123}\\ +\hspace*{0pt}\\ +\hspace*{0pt}\\ +\hspace*{0pt}foldla ::~forall a b.~(a -> b -> a) -> a -> [b] -> a;\\ +\hspace*{0pt}foldla f a [] = a;\\ +\hspace*{0pt}foldla f a (x :~xs) = foldla f (f a x) xs;\\ +\hspace*{0pt}\\ +\hspace*{0pt}rev ::~forall a.~[a] -> [a];\\ +\hspace*{0pt}rev xs = foldla ({\char92}~xsa x -> x :~xsa) [] xs;\\ +\hspace*{0pt}\\ +\hspace*{0pt}list{\char95}case ::~forall t a.~t -> (a -> [a] -> t) -> [a] -> t;\\ +\hspace*{0pt}list{\char95}case f1 f2 (a :~list) = f2 a list;\\ +\hspace*{0pt}list{\char95}case f1 f2 [] = f1;\\ +\hspace*{0pt}\\ +\hspace*{0pt}data Queue a = AQueue [a] [a];\\ +\hspace*{0pt}\\ +\hspace*{0pt}empty ::~forall a.~Queue a;\\ +\hspace*{0pt}empty = AQueue [] [];\\ +\hspace*{0pt}\\ +\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\ +\hspace*{0pt}dequeue (AQueue [] []) = (Nothing,~AQueue [] []);\\ +\hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\ +\hspace*{0pt}dequeue (AQueue (v :~va) []) =\\ +\hspace*{0pt} ~let {\char123}\\ +\hspace*{0pt} ~~~(y :~ys) = rev (v :~va);\\ +\hspace*{0pt} ~{\char125}~in (Just y,~AQueue [] ys);\\ +\hspace*{0pt}\\ +\hspace*{0pt}enqueue ::~forall a.~a -> Queue a -> Queue a;\\ +\hspace*{0pt}enqueue x (AQueue xs ys) = AQueue (x :~xs) ys;\\ +\hspace*{0pt}\\ +\hspace*{0pt}{\char125}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This demonstrates the basic usage of the \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command; + for more details see \secref{sec:further}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Code generator architecture \label{sec:concept}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +What you have seen so far should be already enough in a lot of cases. If you + are content with this, you can quit reading here. Anyway, in order to customise + and adapt the code generator, it is inevitable to gain some understanding + how it works. + + \begin{figure}[h] + \begin{tikzpicture}[x = 4.2cm, y = 1cm] + \tikzstyle entity=[rounded corners, draw, thick, color = black, fill = white]; + \tikzstyle process=[ellipse, draw, thick, color = green, fill = white]; + \tikzstyle process_arrow=[->, semithick, color = green]; + \node (HOL) at (0, 4) [style=entity] {\isa{Isabelle{\isacharslash}HOL} theory}; + \node (eqn) at (2, 2) [style=entity] {code equations}; + \node (iml) at (2, 0) [style=entity] {intermediate language}; + \node (seri) at (1, 0) [style=process] {serialisation}; + \node (SML) at (0, 3) [style=entity] {\isa{SML}}; + \node (OCaml) at (0, 2) [style=entity] {\isa{OCaml}}; + \node (further) at (0, 1) [style=entity] {\isa{{\isasymdots}}}; + \node (Haskell) at (0, 0) [style=entity] {\isa{Haskell}}; + \draw [style=process_arrow] (HOL) .. controls (2, 4) .. + node [style=process, near start] {selection} + node [style=process, near end] {preprocessing} + (eqn); + \draw [style=process_arrow] (eqn) -- node (transl) [style=process] {translation} (iml); + \draw [style=process_arrow] (iml) -- (seri); + \draw [style=process_arrow] (seri) -- (SML); + \draw [style=process_arrow] (seri) -- (OCaml); + \draw [style=process_arrow, dashed] (seri) -- (further); + \draw [style=process_arrow] (seri) -- (Haskell); + \end{tikzpicture} + \caption{Code generator architecture} + \label{fig:arch} + \end{figure} + + The code generator employs a notion of executability + for three foundational executable ingredients known + from functional programming: + \emph{code equations}, \emph{datatypes}, and + \emph{type classes}. A code equation as a first approximation + is a theorem of the form \isa{f\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n\ {\isasymequiv}\ t} + (an equation headed by a constant \isa{f} with arguments + \isa{t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n} and right hand side \isa{t}). + Code generation aims to turn code equations + into a functional program. This is achieved by three major + components which operate sequentially, i.e. the result of one is + the input + of the next in the chain, see diagram \ref{fig:arch}: + + \begin{itemize} + + \item Out of the vast collection of theorems proven in a + \qn{theory}, a reasonable subset modelling + code equations is \qn{selected}. + + \item On those selected theorems, certain + transformations are carried out + (\qn{preprocessing}). Their purpose is to turn theorems + representing non- or badly executable + specifications into equivalent but executable counterparts. + The result is a structured collection of \qn{code theorems}. + + \item Before the selected code equations are continued with, + they can be \qn{preprocessed}, i.e. subjected to theorem + transformations. This \qn{preprocessor} is an interface which + allows to apply + the full expressiveness of ML-based theorem transformations + to code generation; motivating examples are shown below, see + \secref{sec:preproc}. + The result of the preprocessing step is a structured collection + of code equations. + + \item These code equations are \qn{translated} to a program + in an abstract intermediate language. Think of it as a kind + of \qt{Mini-Haskell} with four \qn{statements}: \isa{data} + (for datatypes), \isa{fun} (stemming from code equations), + also \isa{class} and \isa{inst} (for type classes). + + \item Finally, the abstract program is \qn{serialised} into concrete + source code of a target language. + + \end{itemize} + + \noindent From these steps, only the two last are carried out outside the logic; by + keeping this layer as thin as possible, the amount of code to trust is + kept to a minimum.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/document/ML.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/document/ML.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,255 @@ +% +\begin{isabellebody}% +\def\isabellecontext{ML}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ {\isachardoublequoteopen}ML{\isachardoublequoteclose}\isanewline +\isakeyword{imports}\ Setup\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsection{ML system interfaces \label{sec:ml}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Since the code generator framework not only aims to provide + a nice Isar interface but also to form a base for + code-generation-based applications, here a short + description of the most important ML interfaces.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Executable theory content: \isa{Code}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +This Pure module implements the core notions of + executable content of a theory.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Managing executable content% +} +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{Code.add\_eqn}\verb|Code.add_eqn: thm -> theory -> theory| \\ + \indexdef{}{ML}{Code.del\_eqn}\verb|Code.del_eqn: thm -> theory -> theory| \\ + \indexdef{}{ML}{Code.add\_eqnl}\verb|Code.add_eqnl: string * (thm * bool) list lazy -> theory -> theory| \\ + \indexdef{}{ML}{Code.map\_pre}\verb|Code.map_pre: (simpset -> simpset) -> theory -> theory| \\ + \indexdef{}{ML}{Code.map\_post}\verb|Code.map_post: (simpset -> simpset) -> theory -> theory| \\ + \indexdef{}{ML}{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline% +\verb| -> theory -> theory| \\ + \indexdef{}{ML}{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\ + \indexdef{}{ML}{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\ + \indexdef{}{ML}{Code.get\_datatype}\verb|Code.get_datatype: theory -> string|\isasep\isanewline% +\verb| -> (string * sort) list * (string * typ list) list| \\ + \indexdef{}{ML}{Code.get\_datatype\_of\_constr}\verb|Code.get_datatype_of_constr: theory -> string -> string option| + \end{mldecls} + + \begin{description} + + \item \verb|Code.add_eqn|~\isa{thm}~\isa{thy} adds function + theorem \isa{thm} to executable content. + + \item \verb|Code.del_eqn|~\isa{thm}~\isa{thy} removes function + theorem \isa{thm} from executable content, if present. + + \item \verb|Code.add_eqnl|~\isa{{\isacharparenleft}const{\isacharcomma}\ lthms{\isacharparenright}}~\isa{thy} adds + suspended code equations \isa{lthms} for constant + \isa{const} to executable content. + + \item \verb|Code.map_pre|~\isa{f}~\isa{thy} changes + the preprocessor simpset. + + \item \verb|Code.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds + function transformer \isa{f} (named \isa{name}) to executable content; + \isa{f} is a transformer of the code equations belonging + to a certain function definition, depending on the + current theory context. Returning \isa{NONE} indicates that no + transformation took place; otherwise, the whole process will be iterated + with the new code equations. + + \item \verb|Code.del_functrans|~\isa{name}~\isa{thy} removes + function transformer named \isa{name} from executable content. + + \item \verb|Code.add_datatype|~\isa{cs}~\isa{thy} adds + a datatype to executable content, with generation + set \isa{cs}. + + \item \verb|Code.get_datatype_of_constr|~\isa{thy}~\isa{const} + returns type constructor corresponding to + constructor \isa{const}; returns \isa{NONE} + if \isa{const} is no constructor. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Auxiliary% +} +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\ + \indexdef{}{ML}{Code\_Unit.head\_eqn}\verb|Code_Unit.head_eqn: theory -> thm -> string * ((string * sort) list * typ)| \\ + \indexdef{}{ML}{Code\_Unit.rewrite\_eqn}\verb|Code_Unit.rewrite_eqn: simpset -> thm -> thm| \\ + \end{mldecls} + + \begin{description} + + \item \verb|Code_Unit.read_const|~\isa{thy}~\isa{s} + reads a constant as a concrete term expression \isa{s}. + + \item \verb|Code_Unit.head_eqn|~\isa{thy}~\isa{thm} + extracts the constant and its type from a code equation \isa{thm}. + + \item \verb|Code_Unit.rewrite_eqn|~\isa{ss}~\isa{thm} + rewrites a code equation \isa{thm} with a simpset \isa{ss}; + only arguments and right hand side are rewritten, + not the head of the code equation. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Implementing code generator applications% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Implementing code generator applications on top + of the framework set out so far usually not only + involves using those primitive interfaces + but also storing code-dependent data and various + other things.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Data depending on the theory's executable content% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Due to incrementality of code generation, changes in the + theory's executable content have to be propagated in a + certain fashion. Additionally, such changes may occur + not only during theory extension but also during theory + merge, which is a little bit nasty from an implementation + point of view. The framework provides a solution + to this technical challenge by providing a functorial + data slot \verb|CodeDataFun|; on instantiation + of this functor, the following types and operations + are required: + + \medskip + \begin{tabular}{l} + \isa{type\ T} \\ + \isa{val\ empty{\isacharcolon}\ T} \\ + \isa{val\ purge{\isacharcolon}\ theory\ {\isasymrightarrow}\ string\ list\ option\ {\isasymrightarrow}\ T\ {\isasymrightarrow}\ T} + \end{tabular} + + \begin{description} + + \item \isa{T} the type of data to store. + + \item \isa{empty} initial (empty) data. + + \item \isa{purge}~\isa{thy}~\isa{consts} propagates changes in executable content; + \isa{consts} indicates the kind + of change: \verb|NONE| stands for a fundamental change + which invalidates any existing code, \isa{SOME\ consts} + hints that executable content for constants \isa{consts} + has changed. + + \end{description} + + \noindent An instance of \verb|CodeDataFun| provides the following + interface: + + \medskip + \begin{tabular}{l} + \isa{get{\isacharcolon}\ theory\ {\isasymrightarrow}\ T} \\ + \isa{change{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ T{\isacharparenright}\ {\isasymrightarrow}\ T} \\ + \isa{change{\isacharunderscore}yield{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T{\isacharparenright}\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T} + \end{tabular} + + \begin{description} + + \item \isa{get} retrieval of the current data. + + \item \isa{change} update of current data (cached!) + by giving a continuation. + + \item \isa{change{\isacharunderscore}yield} update with side result. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\bigskip + + \emph{Happy proving, happy hacking!}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/document/Program.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/document/Program.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,1238 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Program}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Program\isanewline +\isakeyword{imports}\ Introduction\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsection{Turning Theories into Programs \label{sec:program}% +} +\isamarkuptrue% +% +\isamarkupsubsection{The \isa{Isabelle{\isacharslash}HOL} default setup% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We have already seen how by default equations stemming from + \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}/\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}/\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}} + statements are used for code generation. This default behaviour + can be changed, e.g. by providing different code equations. + All kinds of customisation shown in this section is \emph{safe} + in the sense that the user does not have to worry about + correctness -- all programs generatable that way are partially + correct.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Selecting code equations% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Coming back to our introductory example, we + could provide an alternative code equations for \isa{dequeue} + explicitly:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \ \ \ \ \ else\ dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}cases\ xs{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\ {\isacharparenleft}cases\ {\isachardoublequoteopen}rev\ xs{\isachardoublequoteclose}{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The annotation \isa{{\isacharbrackleft}code{\isacharbrackright}} is an \isa{Isar} + \isa{attribute} which states that the given theorems should be + considered as code equations for a \isa{fun} statement -- + the corresponding constant is determined syntactically. The resulting code:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\ +\hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\ +\hspace*{0pt}dequeue (AQueue xs []) =\\ +\hspace*{0pt} ~(if nulla xs then (Nothing,~AQueue [] [])\\ +\hspace*{0pt} ~~~else dequeue (AQueue [] (rev xs)));% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent You may note that the equality test \isa{xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}} has been + replaced by the predicate \isa{null\ xs}. This is due to the default + setup in the \qn{preprocessor} to be discussed further below (\secref{sec:preproc}). + + Changing the default constructor set of datatypes is also + possible. See \secref{sec:datatypes} for an example. + + As told in \secref{sec:concept}, code generation is based + on a structured collection of code theorems. + For explorative purpose, this collection + may be inspected using the \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} command:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{code{\isacharunderscore}thms}\isamarkupfalse% +\ dequeue% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent prints a table with \emph{all} code equations + for \isa{dequeue}, including + \emph{all} code equations those equations depend + on recursively. + + Similarly, the \hyperlink{command.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}} command shows a graph + visualising dependencies between code equations.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{\isa{class} and \isa{instantiation}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Concerning type classes and code generation, let us examine an example + from abstract algebra:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{class}\isamarkupfalse% +\ semigroup\ {\isacharequal}\isanewline +\ \ \isakeyword{fixes}\ mult\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymotimes}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline +\ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{class}\isamarkupfalse% +\ monoid\ {\isacharequal}\ semigroup\ {\isacharplus}\isanewline +\ \ \isakeyword{fixes}\ neutral\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\ \ \isakeyword{assumes}\ neutl{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ neutr{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instantiation}\isamarkupfalse% +\ nat\ {\isacharcolon}{\isacharcolon}\ monoid\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{primrec}\isamarkupfalse% +\ mult{\isacharunderscore}nat\ \isakeyword{where}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isadigit{0}}\ {\isasymotimes}\ n\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}Suc\ m\ {\isasymotimes}\ n\ {\isacharequal}\ n\ {\isacharplus}\ m\ {\isasymotimes}\ n{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ neutral{\isacharunderscore}nat\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ add{\isacharunderscore}mult{\isacharunderscore}distrib{\isacharcolon}\isanewline +\ \ \isakeyword{fixes}\ n\ m\ q\ {\isacharcolon}{\isacharcolon}\ nat\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isacharparenleft}n\ {\isacharplus}\ m{\isacharparenright}\ {\isasymotimes}\ q\ {\isacharequal}\ n\ {\isasymotimes}\ q\ {\isacharplus}\ m\ {\isasymotimes}\ q{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ n{\isacharparenright}\ simp{\isacharunderscore}all\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ m\ n\ q\ {\isacharcolon}{\isacharcolon}\ nat\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymotimes}\ q\ {\isacharequal}\ m\ {\isasymotimes}\ {\isacharparenleft}n\ {\isasymotimes}\ q{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ m{\isacharparenright}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ add{\isacharunderscore}mult{\isacharunderscore}distrib{\isacharparenright}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}m\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ m{\isacharparenright}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent We define the natural operation of the natural numbers + on monoids:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{primrec}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ monoid{\isacharparenright}\ pow\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ \ \ {\isachardoublequoteopen}pow\ {\isadigit{0}}\ a\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}pow\ {\isacharparenleft}Suc\ n{\isacharparenright}\ a\ {\isacharequal}\ a\ {\isasymotimes}\ pow\ n\ a{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This we use to define the discrete exponentiation function:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{definition}\isamarkupfalse% +\ bexp\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}bexp\ n\ {\isacharequal}\ pow\ n\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The corresponding code:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}module Example where {\char123}\\ +\hspace*{0pt}\\ +\hspace*{0pt}\\ +\hspace*{0pt}data Nat = Zero{\char95}nat | Suc Nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}class Semigroup a where {\char123}\\ +\hspace*{0pt} ~mult ::~a -> a -> a;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}class (Semigroup a) => Monoid a where {\char123}\\ +\hspace*{0pt} ~neutral ::~a;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}pow ::~forall a.~(Monoid a) => Nat -> a -> a;\\ +\hspace*{0pt}pow Zero{\char95}nat a = neutral;\\ +\hspace*{0pt}pow (Suc n) a = mult a (pow n a);\\ +\hspace*{0pt}\\ +\hspace*{0pt}plus{\char95}nat ::~Nat -> Nat -> Nat;\\ +\hspace*{0pt}plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n);\\ +\hspace*{0pt}plus{\char95}nat Zero{\char95}nat n = n;\\ +\hspace*{0pt}\\ +\hspace*{0pt}neutral{\char95}nat ::~Nat;\\ +\hspace*{0pt}neutral{\char95}nat = Suc Zero{\char95}nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}mult{\char95}nat ::~Nat -> Nat -> Nat;\\ +\hspace*{0pt}mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat;\\ +\hspace*{0pt}mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\ +\hspace*{0pt}\\ +\hspace*{0pt}instance Semigroup Nat where {\char123}\\ +\hspace*{0pt} ~mult = mult{\char95}nat;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}instance Monoid Nat where {\char123}\\ +\hspace*{0pt} ~neutral = neutral{\char95}nat;\\ +\hspace*{0pt}{\char125};\\ +\hspace*{0pt}\\ +\hspace*{0pt}bexp ::~Nat -> Nat;\\ +\hspace*{0pt}bexp n = pow n (Suc (Suc Zero{\char95}nat));\\ +\hspace*{0pt}\\ +\hspace*{0pt}{\char125}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This is a convenient place to show how explicit dictionary construction + manifests in generated code (here, the same example in \isa{SML}):% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a semigroup = {\char123}mult :~'a -> 'a -> 'a{\char125};\\ +\hspace*{0pt}fun mult (A{\char95}:'a semigroup) = {\char35}mult A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a monoid = {\char123}Program{\char95}{\char95}semigroup{\char95}monoid :~'a semigroup,~neutral :~'a{\char125};\\ +\hspace*{0pt}fun semigroup{\char95}monoid (A{\char95}:'a monoid) = {\char35}Program{\char95}{\char95}semigroup{\char95}monoid A{\char95};\\ +\hspace*{0pt}fun neutral (A{\char95}:'a monoid) = {\char35}neutral A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun pow A{\char95}~Zero{\char95}nat a = neutral A{\char95}\\ +\hspace*{0pt} ~| pow A{\char95}~(Suc n) a = mult (semigroup{\char95}monoid A{\char95}) a (pow A{\char95}~n a);\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n)\\ +\hspace*{0pt} ~| plus{\char95}nat Zero{\char95}nat n = n;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val neutral{\char95}nat :~nat = Suc Zero{\char95}nat\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat\\ +\hspace*{0pt} ~| mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\ +\hspace*{0pt}\\ +\hspace*{0pt}val semigroup{\char95}nat = {\char123}mult = mult{\char95}nat{\char125}~:~nat semigroup;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val monoid{\char95}nat =\\ +\hspace*{0pt} ~{\char123}Program{\char95}{\char95}semigroup{\char95}monoid = semigroup{\char95}nat,~neutral = neutral{\char95}nat{\char125}~:\\ +\hspace*{0pt} ~nat monoid;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun bexp n = pow monoid{\char95}nat n (Suc (Suc Zero{\char95}nat));\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Note the parameters with trailing underscore (\verb|A_|) + which are the dictionary parameters.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{The preprocessor \label{sec:preproc}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Before selected function theorems are turned into abstract + code, a chain of definitional transformation steps is carried + out: \emph{preprocessing}. In essence, the preprocessor + consists of two components: a \emph{simpset} and \emph{function transformers}. + + The \emph{simpset} allows to employ the full generality of the Isabelle + simplifier. Due to the interpretation of theorems + as code equations, rewrites are applied to the right + hand side and the arguments of the left hand side of an + equation, but never to the constant heading the left hand side. + An important special case are \emph{inline theorems} which may be + declared and undeclared using the + \emph{code inline} or \emph{code inline del} attribute respectively. + + Some common applications:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{itemize} +% +\begin{isamarkuptext}% +\item replacing non-executable constructs by executable ones:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}x\ {\isasymin}\ set\ xs\ {\isasymlongleftrightarrow}\ x\ mem\ xs{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\item eliminating superfluous constants:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\item replacing executable but inconvenient constructs:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongleftrightarrow}\ List{\isachardot}null\ xs{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\end{itemize} +% +\begin{isamarkuptext}% +\noindent \emph{Function transformers} provide a very general interface, + transforming a list of function theorems to another + list of function theorems, provided that neither the heading + constant nor its type change. The \isa{{\isadigit{0}}} / \isa{Suc} + pattern elimination implemented in + theory \isa{Efficient{\isacharunderscore}Nat} (see \secref{eff_nat}) uses this + interface. + + \noindent The current setup of the preprocessor may be inspected using + the \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} command. + \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} provides a convenient + mechanism to inspect the impact of a preprocessor setup + on code equations. + + \begin{warn} + The attribute \emph{code unfold} + associated with the \isa{SML\ code\ generator} also applies to + the \isa{generic\ code\ generator}: + \emph{code unfold} implies \emph{code inline}. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Datatypes \label{sec:datatypes}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Conceptually, any datatype is spanned by a set of + \emph{constructors} of type \isa{{\isasymtau}\ {\isacharequal}\ {\isasymdots}\ {\isasymRightarrow}\ {\isasymkappa}\ {\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymdots}\ {\isasymalpha}\isactrlisub n} where \isa{{\isacharbraceleft}{\isasymalpha}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlisub n{\isacharbraceright}} is exactly the set of \emph{all} type variables in + \isa{{\isasymtau}}. The HOL datatype package by default registers any new + datatype in the table of datatypes, which may be inspected using the + \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} command. + + In some cases, it is appropriate to alter or extend this table. As + an example, we will develop an alternative representation of the + queue example given in \secref{sec:intro}. The amortised + representation is convenient for generating code but exposes its + \qt{implementation} details, which may be cumbersome when proving + theorems about it. Therefore, here a simple, straightforward + representation of queues:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{datatype}\isamarkupfalse% +\ {\isacharprime}a\ queue\ {\isacharequal}\ Queue\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ empty\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ Queue\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{primrec}\isamarkupfalse% +\ enqueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}Queue\ xs{\isacharparenright}\ {\isacharequal}\ Queue\ {\isacharparenleft}xs\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{fun}\isamarkupfalse% +\ dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ option\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}Queue\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}None{\isacharcomma}\ Queue\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}Queue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ x{\isacharcomma}\ Queue\ xs{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This we can use directly for proving; for executing, + we provide an alternative characterisation:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{definition}\isamarkupfalse% +\ AQueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}AQueue\ xs\ ys\ {\isacharequal}\ Queue\ {\isacharparenleft}ys\ {\isacharat}\ rev\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% +\ AQueue% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Here we define a \qt{constructor} \isa{AQueue} which + is defined in terms of \isa{Queue} and interprets its arguments + according to what the \emph{content} of an amortised queue is supposed + to be. Equipped with this, we are able to prove the following equations + for our primitive queue operations which \qt{implement} the simple + queues in an amortised fashion:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ empty{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{unfolding}\isamarkupfalse% +\ AQueue{\isacharunderscore}def\ empty{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ enqueue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ AQueue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ ys{\isachardoublequoteclose}\isanewline +\ \ \isacommand{unfolding}\isamarkupfalse% +\ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ dequeue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \ \ else\ dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{unfolding}\isamarkupfalse% +\ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp{\isacharunderscore}all% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent For completeness, we provide a substitute for the + \isa{case} combinator on queues:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ queue{\isacharunderscore}case{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}queue{\isacharunderscore}case\ f\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ f\ {\isacharparenleft}ys\ {\isacharat}\ rev\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{unfolding}\isamarkupfalse% +\ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The resulting code looks as expected:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun foldl f a [] = a\\ +\hspace*{0pt} ~| foldl f a (x ::~xs) = foldl f (f a x) xs;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun rev xs = foldl (fn xsa => fn x => x ::~xsa) [] xs;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun null [] = true\\ +\hspace*{0pt} ~| null (x ::~xs) = false;\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype 'a queue = AQueue of 'a list * 'a list;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val empty :~'a queue = AQueue ([],~[])\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun dequeue (AQueue (xs,~y ::~ys)) = (SOME y,~AQueue (xs,~ys))\\ +\hspace*{0pt} ~| dequeue (AQueue (xs,~[])) =\\ +\hspace*{0pt} ~~~(if null xs then (NONE,~AQueue ([],~[]))\\ +\hspace*{0pt} ~~~~~else dequeue (AQueue ([],~rev xs)));\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun enqueue x (AQueue (xs,~ys)) = AQueue (x ::~xs,~ys);\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent From this example, it can be glimpsed that using own + constructor sets is a little delicate since it changes the set of + valid patterns for values of that type. Without going into much + detail, here some practical hints: + + \begin{itemize} + + \item When changing the constructor set for datatypes, take care + to provide alternative equations for the \isa{case} combinator. + + \item Values in the target language need not to be normalised -- + different values in the target language may represent the same + value in the logic. + + \item Usually, a good methodology to deal with the subtleties of + pattern matching is to see the type as an abstract type: provide + a set of operations which operate on the concrete representation + of the type, and derive further operations by combinations of + these primitive ones, without relying on a particular + representation. + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Equality and wellsortedness% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Surely you have already noticed how equality is treated + by the code generator:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{primrec}\isamarkupfalse% +\ collect{\isacharunderscore}duplicates\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharparenleft}z{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ z\ {\isasymin}\ set\ xs\isanewline +\ \ \ \ \ \ then\ if\ z\ {\isasymin}\ set\ ys\isanewline +\ \ \ \ \ \ \ \ then\ collect{\isacharunderscore}duplicates\ xs\ ys\ zs\isanewline +\ \ \ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ xs\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs\isanewline +\ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ {\isacharparenleft}z{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs{\isacharparenright}{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent The membership test during preprocessing is rewritten, + resulting in \isa{op\ mem}, which itself + performs an explicit equality check.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\ +\hspace*{0pt}fun eq (A{\char95}:'a eq) = {\char35}eq A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun eqop A{\char95}~a b = eq A{\char95}~a b;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun member A{\char95}~x [] = false\\ +\hspace*{0pt} ~| member A{\char95}~x (y ::~ys) = eqop A{\char95}~x y orelse member A{\char95}~x ys;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun collect{\char95}duplicates A{\char95}~xs ys [] = xs\\ +\hspace*{0pt} ~| collect{\char95}duplicates A{\char95}~xs ys (z ::~zs) =\\ +\hspace*{0pt} ~~~(if member A{\char95}~z xs\\ +\hspace*{0pt} ~~~~~then (if member A{\char95}~z ys then collect{\char95}duplicates A{\char95}~xs ys zs\\ +\hspace*{0pt} ~~~~~~~~~~~~else collect{\char95}duplicates A{\char95}~xs (z ::~ys) zs)\\ +\hspace*{0pt} ~~~~~else collect{\char95}duplicates A{\char95}~(z ::~xs) (z ::~ys) zs);\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Obviously, polymorphic equality is implemented the Haskell + way using a type class. How is this achieved? HOL introduces + an explicit class \isa{eq} with a corresponding operation + \isa{eq{\isacharunderscore}class{\isachardot}eq} such that \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharequal}\ op\ {\isacharequal}}. + The preprocessing framework does the rest by propagating the + \isa{eq} constraints through all dependent code equations. + For datatypes, instances of \isa{eq} are implicitly derived + when possible. For other types, you may instantiate \isa{eq} + manually like any other type class. + + Though this \isa{eq} class is designed to get rarely in + the way, a subtlety + enters the stage when definitions of overloaded constants + are dependent on operational equality. For example, let + us define a lexicographic ordering on tuples + (also see theory \hyperlink{theory.Product-ord}{\mbox{\isa{Product{\isacharunderscore}ord}}}):% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{instantiation}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}order{\isacharcomma}\ order{\isacharparenright}\ order\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}x\ {\isasymle}\ y\ {\isasymlongleftrightarrow}\ fst\ x\ {\isacharless}\ fst\ y\ {\isasymor}\ fst\ x\ {\isacharequal}\ fst\ y\ {\isasymand}\ snd\ x\ {\isasymle}\ snd\ y{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}x\ {\isacharless}\ y\ {\isasymlongleftrightarrow}\ fst\ x\ {\isacharless}\ fst\ y\ {\isasymor}\ fst\ x\ {\isacharequal}\ fst\ y\ {\isasymand}\ snd\ x\ {\isacharless}\ snd\ y{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\ \isacommand{proof}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp{\isacharcolon}\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}prod{\isacharunderscore}def\ intro{\isacharcolon}\ order{\isacharunderscore}less{\isacharunderscore}trans{\isacharparenright}\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ order{\isacharunderscore}prod\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}order{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}order{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Then code generation will fail. Why? The definition + of \isa{op\ {\isasymle}} depends on equality on both arguments, + which are polymorphic and impose an additional \isa{eq} + class constraint, which the preprocessor does not propagate + (for technical reasons). + + The solution is to add \isa{eq} explicitly to the first sort arguments in the + code theorems:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ order{\isacharunderscore}prod{\isacharunderscore}code\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}order{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}order{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Then code generation succeeds:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\ +\hspace*{0pt}fun eq (A{\char95}:'a eq) = {\char35}eq A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a ord = {\char123}less{\char95}eq :~'a -> 'a -> bool,~less :~'a -> 'a -> bool{\char125};\\ +\hspace*{0pt}fun less{\char95}eq (A{\char95}:'a ord) = {\char35}less{\char95}eq A{\char95};\\ +\hspace*{0pt}fun less (A{\char95}:'a ord) = {\char35}less A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun eqop A{\char95}~a b = eq A{\char95}~a b;\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a preorder = {\char123}Orderings{\char95}{\char95}ord{\char95}preorder :~'a ord{\char125};\\ +\hspace*{0pt}fun ord{\char95}preorder (A{\char95}:'a preorder) = {\char35}Orderings{\char95}{\char95}ord{\char95}preorder A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}type 'a order = {\char123}Orderings{\char95}{\char95}preorder{\char95}order :~'a preorder{\char125};\\ +\hspace*{0pt}fun preorder{\char95}order (A{\char95}:'a order) = {\char35}Orderings{\char95}{\char95}preorder{\char95}order A{\char95};\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\ +\hspace*{0pt} ~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\ +\hspace*{0pt} ~~~eqop A1{\char95}~x1 x2 andalso\\ +\hspace*{0pt} ~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2\\ +\hspace*{0pt} ~| less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\ +\hspace*{0pt} ~~~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\ +\hspace*{0pt} ~~~~~eqop A1{\char95}~x1 x2 andalso\\ +\hspace*{0pt} ~~~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2;\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +In some cases, the automatically derived code equations + for equality on a particular type may not be appropriate. + As example, watch the following datatype representing + monomorphic parametric types (where type constructors + are referred to by natural numbers):% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{datatype}\isamarkupfalse% +\ monotype\ {\isacharequal}\ Mono\ nat\ {\isachardoublequoteopen}monotype\ list{\isachardoublequoteclose}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Then code generation for SML would fail with a message + that the generated code contains illegal mutual dependencies: + the theorem \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}{\isacharparenright}\ {\isasymequiv}\ eq{\isacharunderscore}class{\isachardot}eq\ tyco{\isadigit{1}}\ tyco{\isadigit{2}}\ {\isasymand}\ eq{\isacharunderscore}class{\isachardot}eq\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}} already requires the + instance \isa{monotype\ {\isasymColon}\ eq}, which itself requires + \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}{\isacharparenright}\ {\isasymequiv}\ eq{\isacharunderscore}class{\isachardot}eq\ tyco{\isadigit{1}}\ tyco{\isadigit{2}}\ {\isasymand}\ eq{\isacharunderscore}class{\isachardot}eq\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}}; Haskell has no problem with mutually + recursive \isa{instance} and \isa{function} definitions, + but the SML serialiser does not support this. + + In such cases, you have to provide your own equality equations + involving auxiliary constants. In our case, + \isa{list{\isacharunderscore}all{\isadigit{2}}} can do the job:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{lemma}\isamarkupfalse% +\ monotype{\isacharunderscore}eq{\isacharunderscore}list{\isacharunderscore}all{\isadigit{2}}\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline +\ \ \ \ \ eq{\isacharunderscore}class{\isachardot}eq\ tyco{\isadigit{1}}\ tyco{\isadigit{2}}\ {\isasymand}\ list{\isacharunderscore}all{\isadigit{2}}\ eq{\isacharunderscore}class{\isachardot}eq\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq\ list{\isacharunderscore}all{\isadigit{2}}{\isacharunderscore}eq\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent does not depend on instance \isa{monotype\ {\isasymColon}\ eq}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}structure Example = \\ +\hspace*{0pt}struct\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun null [] = true\\ +\hspace*{0pt} ~| null (x ::~xs) = false;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun eq{\char95}nat (Suc a) Zero{\char95}nat = false\\ +\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat (Suc a) = false\\ +\hspace*{0pt} ~| eq{\char95}nat (Suc nat) (Suc nat') = eq{\char95}nat nat nat'\\ +\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat Zero{\char95}nat = true;\\ +\hspace*{0pt}\\ +\hspace*{0pt}datatype monotype = Mono of nat * monotype list;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun list{\char95}all2 p (x ::~xs) (y ::~ys) = p x y andalso list{\char95}all2 p xs ys\\ +\hspace*{0pt} ~| list{\char95}all2 p xs [] = null xs\\ +\hspace*{0pt} ~| list{\char95}all2 p [] ys = null ys;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun eq{\char95}monotype (Mono (tyco1,~typargs1)) (Mono (tyco2,~typargs2)) =\\ +\hspace*{0pt} ~eq{\char95}nat tyco1 tyco2 andalso list{\char95}all2 eq{\char95}monotype typargs1 typargs2;\\ +\hspace*{0pt}\\ +\hspace*{0pt}end;~(*struct Example*)% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\isamarkupsubsection{Explicit partiality% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Partiality usually enters the game by partial patterns, as + in the following example, again for amortised queues:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{definition}\isamarkupfalse% +\ strict{\isacharunderscore}dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ q\ {\isacharequal}\ {\isacharparenleft}case\ dequeue\ q\isanewline +\ \ \ \ of\ {\isacharparenleft}Some\ x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ strict{\isacharunderscore}dequeue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ {\isacharparenleft}case\ rev\ xs\ of\ y\ {\isacharhash}\ ys\ {\isasymRightarrow}\ {\isacharparenleft}y{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ ys{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ strict{\isacharunderscore}dequeue{\isacharunderscore}def\ dequeue{\isacharunderscore}AQueue\ split{\isacharcolon}\ list{\isachardot}splits{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent In the corresponding code, there is no equation + for the pattern \isa{AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}strict{\char95}dequeue ::~forall a.~Queue a -> (a,~Queue a);\\ +\hspace*{0pt}strict{\char95}dequeue (AQueue xs []) =\\ +\hspace*{0pt} ~let {\char123}\\ +\hspace*{0pt} ~~~(y :~ys) = rev xs;\\ +\hspace*{0pt} ~{\char125}~in (y,~AQueue [] ys);\\ +\hspace*{0pt}strict{\char95}dequeue (AQueue xs (y :~ys)) = (y,~AQueue xs ys);% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent In some cases it is desirable to have this + pseudo-\qt{partiality} more explicitly, e.g.~as follows:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{axiomatization}\isamarkupfalse% +\ empty{\isacharunderscore}queue\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\ strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ q\ {\isacharequal}\ {\isacharparenleft}case\ dequeue\ q\ of\ {\isacharparenleft}Some\ x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isacharbar}\ {\isacharunderscore}\ {\isasymRightarrow}\ empty{\isacharunderscore}queue{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ strict{\isacharunderscore}dequeue{\isacharprime}{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ empty{\isacharunderscore}queue\isanewline +\ \ \ \ \ else\ strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ {\isacharparenleft}y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ strict{\isacharunderscore}dequeue{\isacharprime}{\isacharunderscore}def\ dequeue{\isacharunderscore}AQueue\ split{\isacharcolon}\ list{\isachardot}splits{\isacharparenright}% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +Observe that on the right hand side of the definition of \isa{strict{\isacharunderscore}dequeue{\isacharprime}} the constant \isa{empty{\isacharunderscore}queue} occurs + which is unspecified. + + Normally, if constants without any code equations occur in a + program, the code generator complains (since in most cases this is + not what the user expects). But such constants can also be thought + of as function definitions with no equations which always fail, + since there is never a successful pattern match on the left hand + side. In order to categorise a constant into that category + explicitly, use \hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +\isacommand{code{\isacharunderscore}abort}\isamarkupfalse% +\ empty{\isacharunderscore}queue% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent Then the code generator will just insert an error or + exception at the appropriate position:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimquote +% +\endisadelimquote +% +\isatagquote +% +\begin{isamarkuptext}% +\isatypewriter% +\noindent% +\hspace*{0pt}empty{\char95}queue ::~forall a.~a;\\ +\hspace*{0pt}empty{\char95}queue = error {\char34}empty{\char95}queue{\char34};\\ +\hspace*{0pt}\\ +\hspace*{0pt}strict{\char95}dequeue' ::~forall a.~Queue a -> (a,~Queue a);\\ +\hspace*{0pt}strict{\char95}dequeue' (AQueue xs (y :~ys)) = (y,~AQueue xs ys);\\ +\hspace*{0pt}strict{\char95}dequeue' (AQueue xs []) =\\ +\hspace*{0pt} ~(if nulla xs then empty{\char95}queue\\ +\hspace*{0pt} ~~~else strict{\char95}dequeue' (AQueue [] (rev xs)));% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagquote +{\isafoldquote}% +% +\isadelimquote +% +\endisadelimquote +% +\begin{isamarkuptext}% +\noindent This feature however is rarely needed in practice. + Note also that the \isa{HOL} default setup already declares + \isa{undefined} as \hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}, which is most + likely to be used in such situations.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\ \end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/Codegen.hs --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/Codegen.hs Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,23 @@ +module Codegen where { + +import qualified Nat; + +class Null a where { + nulla :: a; +}; + +heada :: forall a. (Codegen.Null a) => [a] -> a; +heada (x : xs) = x; +heada [] = Codegen.nulla; + +null_option :: forall a. Maybe a; +null_option = Nothing; + +instance Codegen.Null (Maybe a) where { + nulla = Codegen.null_option; +}; + +dummy :: Maybe Nat.Nat; +dummy = Codegen.heada [Just (Nat.Suc Nat.Zero_nat), Nothing]; + +} diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/Example.hs --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/Example.hs Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,33 @@ +{-# OPTIONS_GHC -fglasgow-exts #-} + +module Example where { + + +foldla :: forall a b. (a -> b -> a) -> a -> [b] -> a; +foldla f a [] = a; +foldla f a (x : xs) = foldla f (f a x) xs; + +rev :: forall a. [a] -> [a]; +rev xs = foldla (\ xsa x -> x : xsa) [] xs; + +list_case :: forall t a. t -> (a -> [a] -> t) -> [a] -> t; +list_case f1 f2 (a : list) = f2 a list; +list_case f1 f2 [] = f1; + +data Queue a = AQueue [a] [a]; + +empty :: forall a. Queue a; +empty = AQueue [] []; + +dequeue :: forall a. Queue a -> (Maybe a, Queue a); +dequeue (AQueue [] []) = (Nothing, AQueue [] []); +dequeue (AQueue xs (y : ys)) = (Just y, AQueue xs ys); +dequeue (AQueue (v : va) []) = + let { + (y : ys) = rev (v : va); + } in (Just y, AQueue [] ys); + +enqueue :: forall a. a -> Queue a -> Queue a; +enqueue x (AQueue xs ys) = AQueue (x : xs) ys; + +} diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/arbitrary.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/arbitrary.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,9 @@ +structure Codegen = +struct + +val arbitrary_option : 'a option = NONE; + +fun dummy_option [] = arbitrary_option + | dummy_option (x :: xs) = SOME x; + +end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/bool_infix.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/bool_infix.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,19 @@ +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +fun less_nat m (Suc n) = less_eq_nat m n + | less_nat n Zero_nat = false +and less_eq_nat (Suc m) n = less_nat m n + | less_eq_nat Zero_nat n = true; + +end; (*struct Nat*) + +structure Codegen = +struct + +fun in_interval (k, l) n = + Nat.less_eq_nat k n andalso Nat.less_eq_nat n l; + +end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/bool_literal.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/bool_literal.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,31 @@ +structure HOL = +struct + +datatype boola = False | True; + +fun anda x True = x + | anda x False = False + | anda True x = x + | anda False x = False; + +end; (*struct HOL*) + +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +fun less_nat m (Suc n) = less_eq_nat m n + | less_nat n Zero_nat = HOL.False +and less_eq_nat (Suc m) n = less_nat m n + | less_eq_nat Zero_nat n = HOL.True; + +end; (*struct Nat*) + +structure Codegen = +struct + +fun in_interval (k, l) n = + HOL.anda (Nat.less_eq_nat k n) (Nat.less_eq_nat n l); + +end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/bool_mlbool.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/bool_mlbool.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,19 @@ +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +fun less_nat m (Suc n) = less_eq_nat m n + | less_nat n Zero_nat = false +and less_eq_nat (Suc m) n = less_nat m n + | less_eq_nat Zero_nat n = true; + +end; (*struct Nat*) + +structure Codegen = +struct + +fun in_interval (k, l) n = + (Nat.less_eq_nat k n) andalso (Nat.less_eq_nat n l); + +end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/class.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/class.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,24 @@ +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +end; (*struct Nat*) + +structure Codegen = +struct + +type 'a null = {null : 'a}; +fun null (A_:'a null) = #null A_; + +fun head A_ (x :: xs) = x + | head A_ [] = null A_; + +val null_option : 'a option = NONE; + +fun null_optiona () = {null = null_option} : ('a option) null; + +val dummy : Nat.nat option = + head (null_optiona ()) [SOME (Nat.Suc Nat.Zero_nat), NONE]; + +end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/class.ocaml --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/class.ocaml Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,24 @@ +module Nat = +struct + +type nat = Suc of nat | Zero_nat;; + +end;; (*struct Nat*) + +module Codegen = +struct + +type 'a null = {null : 'a};; +let null _A = _A.null;; + +let rec head _A = function x :: xs -> x + | [] -> null _A;; + +let rec null_option = None;; + +let null_optiona () = ({null = null_option} : ('a option) null);; + +let rec dummy + = head (null_optiona ()) [Some (Nat.Suc Nat.Zero_nat); None];; + +end;; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/collect_duplicates.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/collect_duplicates.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,30 @@ +structure HOL = +struct + +type 'a eq = {eq : 'a -> 'a -> bool}; +fun eq (A_:'a eq) = #eq A_; + +fun eqop A_ a = eq A_ a; + +end; (*struct HOL*) + +structure List = +struct + +fun member A_ x (y :: ys) = + (if HOL.eqop A_ y x then true else member A_ x ys) + | member A_ x [] = false; + +end; (*struct List*) + +structure Codegen = +struct + +fun collect_duplicates A_ xs ys (z :: zs) = + (if List.member A_ z xs + then (if List.member A_ z ys then collect_duplicates A_ xs ys zs + else collect_duplicates A_ xs (z :: ys) zs) + else collect_duplicates A_ (z :: xs) (z :: ys) zs) + | collect_duplicates A_ xs ys [] = xs; + +end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/dirty_set.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/dirty_set.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,102 @@ +structure ROOT = +struct + +structure Nat = +struct + +datatype nat = Zero_nat | Suc of nat; + +end; (*struct Nat*) + +structure Integer = +struct + +datatype bit = B0 | B1; + +datatype int = Pls | Min | Bit of int * bit | Number_of_int of int; + +fun pred (Bit (k, B0)) = Bit (pred k, B1) + | pred (Bit (k, B1)) = Bit (k, B0) + | pred Min = Bit (Min, B0) + | pred Pls = Min; + +fun uminus_int (Number_of_int w) = Number_of_int (uminus_int w) + | uminus_int (Bit (k, B0)) = Bit (uminus_int k, B0) + | uminus_int (Bit (k, B1)) = Bit (pred (uminus_int k), B1) + | uminus_int Min = Bit (Pls, B1) + | uminus_int Pls = Pls; + +fun succ (Bit (k, B0)) = Bit (k, B1) + | succ (Bit (k, B1)) = Bit (succ k, B0) + | succ Min = Pls + | succ Pls = Bit (Pls, B1); + +fun plus_int (Number_of_int v) (Number_of_int w) = + Number_of_int (plus_int v w) + | plus_int k Min = pred k + | plus_int k Pls = k + | plus_int (Bit (k, B1)) (Bit (l, B1)) = Bit (plus_int k (succ l), B0) + | plus_int (Bit (k, B1)) (Bit (l, B0)) = Bit (plus_int k l, B1) + | plus_int (Bit (k, B0)) (Bit (l, b)) = Bit (plus_int k l, b) + | plus_int Min k = pred k + | plus_int Pls k = k; + +fun minus_int (Number_of_int v) (Number_of_int w) = + Number_of_int (plus_int v (uminus_int w)) + | minus_int z w = plus_int z (uminus_int w); + +fun less_eq_int (Number_of_int k) (Number_of_int l) = less_eq_int k l + | less_eq_int (Bit (k1, B1)) (Bit (k2, B0)) = less_int k1 k2 + | less_eq_int (Bit (k1, v)) (Bit (k2, B1)) = less_eq_int k1 k2 + | less_eq_int (Bit (k1, B0)) (Bit (k2, v)) = less_eq_int k1 k2 + | less_eq_int (Bit (k, v)) Min = less_eq_int k Min + | less_eq_int (Bit (k, B1)) Pls = less_int k Pls + | less_eq_int (Bit (k, B0)) Pls = less_eq_int k Pls + | less_eq_int Min (Bit (k, B1)) = less_eq_int Min k + | less_eq_int Min (Bit (k, B0)) = less_int Min k + | less_eq_int Min Min = true + | less_eq_int Min Pls = true + | less_eq_int Pls (Bit (k, v)) = less_eq_int Pls k + | less_eq_int Pls Min = false + | less_eq_int Pls Pls = true +and less_int (Number_of_int k) (Number_of_int l) = less_int k l + | less_int (Bit (k1, B0)) (Bit (k2, B1)) = less_eq_int k1 k2 + | less_int (Bit (k1, B1)) (Bit (k2, v)) = less_int k1 k2 + | less_int (Bit (k1, v)) (Bit (k2, B0)) = less_int k1 k2 + | less_int (Bit (k, B1)) Min = less_int k Min + | less_int (Bit (k, B0)) Min = less_eq_int k Min + | less_int (Bit (k, v)) Pls = less_int k Pls + | less_int Min (Bit (k, v)) = less_int Min k + | less_int Min Min = false + | less_int Min Pls = true + | less_int Pls (Bit (k, B1)) = less_eq_int Pls k + | less_int Pls (Bit (k, B0)) = less_int Pls k + | less_int Pls Min = false + | less_int Pls Pls = false; + +fun nat_aux n i = + (if less_eq_int i (Number_of_int Pls) then n + else nat_aux (Nat.Suc n) + (minus_int i (Number_of_int (Bit (Pls, B1))))); + +fun nat i = nat_aux Nat.Zero_nat i; + +end; (*struct Integer*) + +structure Codegen = +struct + +val dummy_set : (Nat.nat -> Nat.nat) list = Nat.Suc :: []; + +val foobar_set : Nat.nat list = + Nat.Zero_nat :: + (Nat.Suc Nat.Zero_nat :: + (Integer.nat + (Integer.Number_of_int + (Integer.Bit + (Integer.Bit (Integer.Pls, Integer.B1), Integer.B0))) + :: [])); + +end; (*struct Codegen*) + +end; (*struct ROOT*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/example.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/example.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,27 @@ +structure Example = +struct + +fun foldl f a [] = a + | foldl f a (x :: xs) = foldl f (f a x) xs; + +fun rev xs = foldl (fn xsa => fn x => x :: xsa) [] xs; + +fun list_case f1 f2 (a :: lista) = f2 a lista + | list_case f1 f2 [] = f1; + +datatype 'a queue = AQueue of 'a list * 'a list; + +val empty : 'a queue = AQueue ([], []) + +fun dequeue (AQueue ([], [])) = (NONE, AQueue ([], [])) + | dequeue (AQueue (xs, y :: ys)) = (SOME y, AQueue (xs, ys)) + | dequeue (AQueue (v :: va, [])) = + let + val y :: ys = rev (v :: va); + in + (SOME y, AQueue ([], ys)) + end; + +fun enqueue x (AQueue (xs, ys)) = AQueue (x :: xs, ys); + +end; (*struct Example*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/fac.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/fac.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,22 @@ +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +val one_nat : nat = Suc Zero_nat; + +fun plus_nat (Suc m) n = plus_nat m (Suc n) + | plus_nat Zero_nat n = n; + +fun times_nat (Suc m) n = plus_nat n (times_nat m n) + | times_nat Zero_nat n = Zero_nat; + +end; (*struct Nat*) + +structure Codegen = +struct + +fun fac (Nat.Suc n) = Nat.times_nat (Nat.Suc n) (fac n) + | fac Nat.Zero_nat = Nat.one_nat; + +end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/integers.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/integers.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,59 @@ +structure ROOT = +struct + +structure Integer = +struct + +datatype bit = B0 | B1; + +datatype int = Pls | Min | Bit of int * bit | Number_of_int of int; + +fun pred (Bit (k, B0)) = Bit (pred k, B1) + | pred (Bit (k, B1)) = Bit (k, B0) + | pred Min = Bit (Min, B0) + | pred Pls = Min; + +fun succ (Bit (k, B0)) = Bit (k, B1) + | succ (Bit (k, B1)) = Bit (succ k, B0) + | succ Min = Pls + | succ Pls = Bit (Pls, B1); + +fun plus_int (Number_of_int v) (Number_of_int w) = + Number_of_int (plus_int v w) + | plus_int k Min = pred k + | plus_int k Pls = k + | plus_int (Bit (k, B1)) (Bit (l, B1)) = Bit (plus_int k (succ l), B0) + | plus_int (Bit (k, B1)) (Bit (l, B0)) = Bit (plus_int k l, B1) + | plus_int (Bit (k, B0)) (Bit (l, b)) = Bit (plus_int k l, b) + | plus_int Min k = pred k + | plus_int Pls k = k; + +fun uminus_int (Number_of_int w) = Number_of_int (uminus_int w) + | uminus_int (Bit (k, B0)) = Bit (uminus_int k, B0) + | uminus_int (Bit (k, B1)) = Bit (pred (uminus_int k), B1) + | uminus_int Min = Bit (Pls, B1) + | uminus_int Pls = Pls; + +fun times_int (Number_of_int v) (Number_of_int w) = + Number_of_int (times_int v w) + | times_int (Bit (k, B0)) l = Bit (times_int k l, B0) + | times_int (Bit (k, B1)) l = plus_int (Bit (times_int k l, B0)) l + | times_int Min k = uminus_int k + | times_int Pls w = Pls; + +end; (*struct Integer*) + +structure Codegen = +struct + +fun double_inc k = + Integer.plus_int + (Integer.times_int + (Integer.Number_of_int + (Integer.Bit (Integer.Bit (Integer.Pls, Integer.B1), Integer.B0))) + k) + (Integer.Number_of_int (Integer.Bit (Integer.Pls, Integer.B1))); + +end; (*struct Codegen*) + +end; (*struct ROOT*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/lexicographic.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/lexicographic.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,19 @@ +structure HOL = +struct + +type 'a eq = {eq : 'a -> 'a -> bool}; +fun eq (A_:'a eq) = #eq A_; + +type 'a ord = {less_eq : 'a -> 'a -> bool, less : 'a -> 'a -> bool}; +fun less_eq (A_:'a ord) = #less_eq A_; +fun less (A_:'a ord) = #less A_; + +end; (*struct HOL*) + +structure Codegen = +struct + +fun less_eq (A1_, A2_) B_ (x1, y1) (x2, y2) = + HOL.less A2_ x1 x2 orelse HOL.eq A1_ x1 x2 andalso HOL.less_eq B_ y1 y2; + +end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/lookup.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/lookup.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,13 @@ +structure ROOT = +struct + +structure Codegen = +struct + +fun lookup ((k, v) :: xs) l = + (if ((k : string) = l) then SOME v else lookup xs l) + | lookup [] l = NONE; + +end; (*struct Codegen*) + +end; (*struct ROOT*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/monotype.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/monotype.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,34 @@ +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +fun eq_nat (Suc a) Zero_nat = false + | eq_nat Zero_nat (Suc a) = false + | eq_nat (Suc nat) (Suc nat') = eq_nat nat nat' + | eq_nat Zero_nat Zero_nat = true; + +end; (*struct Nat*) + +structure List = +struct + +fun null (x :: xs) = false + | null [] = true; + +fun list_all2 p (x :: xs) (y :: ys) = p x y andalso list_all2 p xs ys + | list_all2 p xs [] = null xs + | list_all2 p [] ys = null ys; + +end; (*struct List*) + +structure Codegen = +struct + +datatype monotype = Mono of Nat.nat * monotype list; + +fun eq_monotype (Mono (tyco1, typargs1)) (Mono (tyco2, typargs2)) = + Nat.eq_nat tyco1 tyco2 andalso + List.list_all2 eq_monotype typargs1 typargs2; + +end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/nat_binary.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/nat_binary.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,17 @@ +structure Nat = +struct + +datatype nat = Dig1 of nat | Dig0 of nat | One_nat | Zero_nat; + +fun plus_nat (Dig1 m) (Dig1 n) = Dig0 (plus_nat (plus_nat m n) One_nat) + | plus_nat (Dig1 m) (Dig0 n) = Dig1 (plus_nat m n) + | plus_nat (Dig0 m) (Dig1 n) = Dig1 (plus_nat m n) + | plus_nat (Dig0 m) (Dig0 n) = Dig0 (plus_nat m n) + | plus_nat (Dig1 m) One_nat = Dig0 (plus_nat m One_nat) + | plus_nat One_nat (Dig1 n) = Dig0 (plus_nat n One_nat) + | plus_nat (Dig0 m) One_nat = Dig1 m + | plus_nat One_nat (Dig0 n) = Dig1 n + | plus_nat m Zero_nat = m + | plus_nat Zero_nat n = n; + +end; (*struct Nat*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/pick1.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/pick1.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,44 @@ +structure HOL = +struct + +fun leta s f = f s; + +end; (*struct HOL*) + +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +fun less_nat m (Suc n) = less_eq_nat m n + | less_nat n Zero_nat = false +and less_eq_nat (Suc m) n = less_nat m n + | less_eq_nat Zero_nat n = true; + +fun minus_nat (Suc m) (Suc n) = minus_nat m n + | minus_nat Zero_nat n = Zero_nat + | minus_nat m Zero_nat = m; + +end; (*struct Nat*) + +structure Product_Type = +struct + +fun split f (a, b) = f a b; + +end; (*struct Product_Type*) + +structure Codegen = +struct + +fun pick ((k, v) :: xs) n = + (if Nat.less_nat n k then v else pick xs (Nat.minus_nat n k)) + | pick (x :: xs) n = + let + val a = x; + val (k, v) = a; + in + (if Nat.less_nat n k then v else pick xs (Nat.minus_nat n k)) + end; + +end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/Thy/examples/tree.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/Thy/examples/tree.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,95 @@ +structure HOL = +struct + +type 'a eq = {eq : 'a -> 'a -> bool}; +fun eq (A_:'a eq) = #eq A_; + +type 'a ord = {less_eq : 'a -> 'a -> bool, less : 'a -> 'a -> bool}; +fun less_eq (A_:'a ord) = #less_eq A_; +fun less (A_:'a ord) = #less A_; + +fun eqop A_ a = eq A_ a; + +end; (*struct HOL*) + +structure Orderings = +struct + +type 'a preorder = {Orderings__ord_preorder : 'a HOL.ord}; +fun ord_preorder (A_:'a preorder) = #Orderings__ord_preorder A_; + +type 'a order = {Orderings__preorder_order : 'a preorder}; +fun preorder_order (A_:'a order) = #Orderings__preorder_order A_; + +type 'a linorder = {Orderings__order_linorder : 'a order}; +fun order_linorder (A_:'a linorder) = #Orderings__order_linorder A_; + +end; (*struct Orderings*) + +structure Nat = +struct + +datatype nat = Suc of nat | Zero_nat; + +fun eq_nat (Suc a) Zero_nat = false + | eq_nat Zero_nat (Suc a) = false + | eq_nat (Suc nat) (Suc nat') = eq_nat nat nat' + | eq_nat Zero_nat Zero_nat = true; + +val eq_nata = {eq = eq_nat} : nat HOL.eq; + +fun less_nat m (Suc n) = less_eq_nat m n + | less_nat n Zero_nat = false +and less_eq_nat (Suc m) n = less_nat m n + | less_eq_nat Zero_nat n = true; + +val ord_nat = {less_eq = less_eq_nat, less = less_nat} : nat HOL.ord; + +val preorder_nat = {Orderings__ord_preorder = ord_nat} : + nat Orderings.preorder; + +val order_nat = {Orderings__preorder_order = preorder_nat} : + nat Orderings.order; + +val linorder_nat = {Orderings__order_linorder = order_nat} : + nat Orderings.linorder; + +end; (*struct Nat*) + +structure Codegen = +struct + +datatype ('a, 'b) searchtree = + Branch of ('a, 'b) searchtree * 'a * ('a, 'b) searchtree | + Leaf of 'a * 'b; + +fun update (A1_, A2_) (it, entry) (Leaf (key, vala)) = + (if HOL.eqop A1_ it key then Leaf (key, entry) + else (if HOL.less_eq + ((Orderings.ord_preorder o Orderings.preorder_order o + Orderings.order_linorder) + A2_) + it key + then Branch (Leaf (it, entry), it, Leaf (key, vala)) + else Branch (Leaf (key, vala), it, Leaf (it, entry)))) + | update (A1_, A2_) (it, entry) (Branch (t1, key, t2)) = + (if HOL.less_eq + ((Orderings.ord_preorder o Orderings.preorder_order o + Orderings.order_linorder) + A2_) + it key + then Branch (update (A1_, A2_) (it, entry) t1, key, t2) + else Branch (t1, key, update (A1_, A2_) (it, entry) t2)); + +val example : (Nat.nat, (Nat.nat list)) searchtree = + update (Nat.eq_nata, Nat.linorder_nat) + (Nat.Suc (Nat.Suc (Nat.Suc (Nat.Suc Nat.Zero_nat))), + [Nat.Suc (Nat.Suc Nat.Zero_nat), Nat.Suc (Nat.Suc Nat.Zero_nat)]) + (update (Nat.eq_nata, Nat.linorder_nat) + (Nat.Suc (Nat.Suc (Nat.Suc Nat.Zero_nat)), + [Nat.Suc (Nat.Suc (Nat.Suc Nat.Zero_nat))]) + (update (Nat.eq_nata, Nat.linorder_nat) + (Nat.Suc (Nat.Suc Nat.Zero_nat), [Nat.Suc (Nat.Suc Nat.Zero_nat)]) + (Leaf (Nat.Suc Nat.Zero_nat, [])))); + +end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/codegen.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/codegen.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,53 @@ + +\documentclass[12pt,a4paper,fleqn]{article} +\usepackage{latexsym,graphicx} +\usepackage[refpage]{nomencl} +\usepackage{../iman,../extra,../isar,../proof} +\usepackage{../isabelle,../isabellesym} +\usepackage{style} +\usepackage{pgf} +\usepackage{pgflibraryshapes} +\usepackage{tikz} +\usepackage{../pdfsetup} + +\hyphenation{Isabelle} +\hyphenation{Isar} +\isadroptag{theory} + +\title{\includegraphics[scale=0.5]{isabelle_isar} + \\[4ex] Code generation from Isabelle/HOL theories} +\author{\emph{Florian Haftmann}} + +\begin{document} + +\maketitle + +\begin{abstract} + \noindent This tutorial gives an introduction to a generic code generator framework in Isabelle + for generating executable code in functional programming languages from logical + specifications in Isabelle/HOL. +\end{abstract} + +\thispagestyle{empty}\clearpage + +\pagenumbering{roman} +\clearfirst + +\input{Thy/document/Introduction.tex} +\input{Thy/document/Program.tex} +\input{Thy/document/Adaption.tex} +\input{Thy/document/Further.tex} +\input{Thy/document/ML.tex} + +\begingroup +\bibliographystyle{plain} \small\raggedright\frenchspacing +\bibliography{../manual} +\endgroup + +\end{document} + + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/codegen_process.pdf Binary file doc-src/Codegen/codegen_process.pdf has changed diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/codegen_process.ps --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/codegen_process.ps Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,586 @@ +%!PS-Adobe-2.0 +%%Creator: dot version 2.2 (Mon Sep 12 23:33:36 UTC 2005) +%%For: (haftmann) Florian Haftmann +%%Title: _anonymous_0 +%%Pages: (atend) +%%BoundingBox: 35 35 451 291 +%%EndComments +save +%%BeginProlog +/DotDict 200 dict def +DotDict begin + +/setupLatin1 { +mark +/EncodingVector 256 array def + EncodingVector 0 + +ISOLatin1Encoding 0 255 getinterval putinterval + +EncodingVector + dup 306 /AE + dup 301 /Aacute + dup 302 /Acircumflex + dup 304 /Adieresis + dup 300 /Agrave + dup 305 /Aring + dup 303 /Atilde + dup 307 /Ccedilla + dup 311 /Eacute + dup 312 /Ecircumflex + dup 313 /Edieresis + dup 310 /Egrave + dup 315 /Iacute + dup 316 /Icircumflex + dup 317 /Idieresis + dup 314 /Igrave + dup 334 /Udieresis + dup 335 /Yacute + dup 376 /thorn + dup 337 /germandbls + dup 341 /aacute + dup 342 /acircumflex + dup 344 /adieresis + dup 346 /ae + dup 340 /agrave + dup 345 /aring + dup 347 /ccedilla + dup 351 /eacute + dup 352 /ecircumflex + dup 353 /edieresis + dup 350 /egrave + dup 355 /iacute + dup 356 /icircumflex + dup 357 /idieresis + dup 354 /igrave + dup 360 /dcroat + dup 361 /ntilde + dup 363 /oacute + dup 364 /ocircumflex + dup 366 /odieresis + dup 362 /ograve + dup 365 /otilde + dup 370 /oslash + dup 372 /uacute + dup 373 /ucircumflex + dup 374 /udieresis + dup 371 /ugrave + dup 375 /yacute + dup 377 /ydieresis + +% Set up ISO Latin 1 character encoding +/starnetISO { + dup dup findfont dup length dict begin + { 1 index /FID ne { def }{ pop pop } ifelse + } forall + /Encoding EncodingVector def + currentdict end definefont +} def +/Times-Roman starnetISO def +/Times-Italic starnetISO def +/Times-Bold starnetISO def +/Times-BoldItalic starnetISO def +/Helvetica starnetISO def +/Helvetica-Oblique starnetISO def +/Helvetica-Bold starnetISO def +/Helvetica-BoldOblique starnetISO def +/Courier starnetISO def +/Courier-Oblique starnetISO def +/Courier-Bold starnetISO def +/Courier-BoldOblique starnetISO def +cleartomark +} bind def + +%%BeginResource: procset graphviz 0 0 +/coord-font-family /Times-Roman def +/default-font-family /Times-Roman def +/coordfont coord-font-family findfont 8 scalefont def + +/InvScaleFactor 1.0 def +/set_scale { + dup 1 exch div /InvScaleFactor exch def + dup scale +} bind def + +% styles +/solid { [] 0 setdash } bind def +/dashed { [9 InvScaleFactor mul dup ] 0 setdash } bind def +/dotted { [1 InvScaleFactor mul 6 InvScaleFactor mul] 0 setdash } bind def +/invis {/fill {newpath} def /stroke {newpath} def /show {pop newpath} def} bind def +/bold { 2 setlinewidth } bind def +/filled { } bind def +/unfilled { } bind def +/rounded { } bind def +/diagonals { } bind def + +% hooks for setting color +/nodecolor { sethsbcolor } bind def +/edgecolor { sethsbcolor } bind def +/graphcolor { sethsbcolor } bind def +/nopcolor {pop pop pop} bind def + +/beginpage { % i j npages + /npages exch def + /j exch def + /i exch def + /str 10 string def + npages 1 gt { + gsave + coordfont setfont + 0 0 moveto + (\() show i str cvs show (,) show j str cvs show (\)) show + grestore + } if +} bind def + +/set_font { + findfont exch + scalefont setfont +} def + +% draw aligned label in bounding box aligned to current point +/alignedtext { % width adj text + /text exch def + /adj exch def + /width exch def + gsave + width 0 gt { + text stringwidth pop adj mul 0 rmoveto + } if + [] 0 setdash + text show + grestore +} def + +/boxprim { % xcorner ycorner xsize ysize + 4 2 roll + moveto + 2 copy + exch 0 rlineto + 0 exch rlineto + pop neg 0 rlineto + closepath +} bind def + +/ellipse_path { + /ry exch def + /rx exch def + /y exch def + /x exch def + matrix currentmatrix + newpath + x y translate + rx ry scale + 0 0 1 0 360 arc + setmatrix +} bind def + +/endpage { showpage } bind def +/showpage { } def + +/layercolorseq + [ % layer color sequence - darkest to lightest + [0 0 0] + [.2 .8 .8] + [.4 .8 .8] + [.6 .8 .8] + [.8 .8 .8] + ] +def + +/layerlen layercolorseq length def + +/setlayer {/maxlayer exch def /curlayer exch def + layercolorseq curlayer 1 sub layerlen mod get + aload pop sethsbcolor + /nodecolor {nopcolor} def + /edgecolor {nopcolor} def + /graphcolor {nopcolor} def +} bind def + +/onlayer { curlayer ne {invis} if } def + +/onlayers { + /myupper exch def + /mylower exch def + curlayer mylower lt + curlayer myupper gt + or + {invis} if +} def + +/curlayer 0 def + +%%EndResource +%%EndProlog +%%BeginSetup +14 default-font-family set_font +1 setmiterlimit +% /arrowlength 10 def +% /arrowwidth 5 def + +% make sure pdfmark is harmless for PS-interpreters other than Distiller +/pdfmark where {pop} {userdict /pdfmark /cleartomark load put} ifelse +% make '<<' and '>>' safe on PS Level 1 devices +/languagelevel where {pop languagelevel}{1} ifelse +2 lt { + userdict (<<) cvn ([) cvn load put + userdict (>>) cvn ([) cvn load put +} if + +%%EndSetup +%%Page: 1 1 +%%PageBoundingBox: 36 36 451 291 +%%PageOrientation: Portrait +gsave +35 35 416 256 boxprim clip newpath +36 36 translate +0 0 1 beginpage +0 0 translate 0 rotate +[ /CropBox [36 36 451 291] /PAGES pdfmark +0.000 0.000 0.000 graphcolor +14.00 /Times-Roman set_font + +% theory +gsave 10 dict begin +newpath 93 254 moveto +1 254 lineto +1 214 lineto +93 214 lineto +closepath +stroke +gsave 10 dict begin +8 237 moveto +(Isabelle/HOL) +[4.56 5.52 6.24 6.96 6.24 3.84 3.84 6.24 3.84 10.08 10.08 8.64] +xshow +16 221 moveto +(Isar theory) +[4.56 5.52 6.24 4.56 3.6 4.08 6.96 6.24 6.96 4.8 6.96] +xshow +end grestore +end grestore + +% selection +gsave 10 dict begin +183 234 38 18 ellipse_path +stroke +gsave 10 dict begin +158 229 moveto +(selection) +[5.52 6.24 3.84 6.24 6.24 3.84 3.84 6.96 6.96] +xshow +end grestore +end grestore + +% theory -> selection +newpath 94 234 moveto +107 234 121 234 135 234 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 135 238 moveto +145 234 lineto +135 231 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 135 238 moveto +145 234 lineto +135 231 lineto +closepath +stroke +end grestore + +% sml +gsave 10 dict begin +newpath 74 144 moveto +20 144 lineto +20 108 lineto +74 108 lineto +closepath +stroke +gsave 10 dict begin +32 121 moveto +(SML) +[7.68 12.48 8.64] +xshow +end grestore +end grestore + +% other +gsave 10 dict begin +gsave 10 dict begin +41 67 moveto +(...) +[3.6 3.6 3.6] +xshow +end grestore +end grestore + +% haskell +gsave 10 dict begin +newpath 77 36 moveto +17 36 lineto +17 0 lineto +77 0 lineto +closepath +stroke +gsave 10 dict begin +25 13 moveto +(Haskell) +[10.08 6.24 5.52 6.72 6.24 3.84 3.84] +xshow +end grestore +end grestore + +% preprocessing +gsave 10 dict begin +183 180 52 18 ellipse_path +stroke +gsave 10 dict begin +143 175 moveto +(preprocessing) +[6.96 4.56 6.24 6.96 4.56 6.96 6.24 6.24 5.52 5.52 3.84 6.96 6.96] +xshow +end grestore +end grestore + +% selection -> preprocessing +newpath 183 216 moveto +183 213 183 211 183 208 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 187 208 moveto +183 198 lineto +180 208 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 187 208 moveto +183 198 lineto +180 208 lineto +closepath +stroke +end grestore + +% def_eqn +gsave 10 dict begin +newpath 403 198 moveto +283 198 lineto +283 162 lineto +403 162 lineto +closepath +stroke +gsave 10 dict begin +291 175 moveto +(defining equations) +[6.96 6.24 4.8 3.84 6.96 3.84 6.96 6.96 3.6 6.24 6.72 6.96 6.24 3.84 3.84 6.96 6.96 5.52] +xshow +end grestore +end grestore + +% preprocessing -> def_eqn +newpath 236 180 moveto +248 180 260 180 273 180 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 273 184 moveto +283 180 lineto +273 177 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 273 184 moveto +283 180 lineto +273 177 lineto +closepath +stroke +end grestore + +% serialization +gsave 10 dict begin +183 72 47 18 ellipse_path +stroke +gsave 10 dict begin +148 67 moveto +(serialization) +[5.52 6.24 4.8 3.84 6.24 3.84 3.84 6.24 6.24 3.84 3.84 6.96 6.96] +xshow +end grestore +end grestore + +% serialization -> sml +newpath 150 85 moveto +129 93 104 103 83 111 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 82 108 moveto +74 115 lineto +85 114 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 82 108 moveto +74 115 lineto +85 114 lineto +closepath +stroke +end grestore + +% serialization -> other +gsave 10 dict begin +dotted +newpath 135 72 moveto +119 72 100 72 84 72 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 84 69 moveto +74 72 lineto +84 76 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 84 69 moveto +74 72 lineto +84 76 lineto +closepath +stroke +end grestore +end grestore + +% serialization -> haskell +newpath 150 59 moveto +131 51 107 42 86 34 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 88 31 moveto +77 30 lineto +85 37 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 88 31 moveto +77 30 lineto +85 37 lineto +closepath +stroke +end grestore + +% translation +gsave 10 dict begin +343 126 43 18 ellipse_path +stroke +gsave 10 dict begin +313 121 moveto +(translation) +[3.84 4.56 6.24 6.96 5.52 3.84 6.24 3.84 3.84 6.96 6.96] +xshow +end grestore +end grestore + +% def_eqn -> translation +newpath 343 162 moveto +343 159 343 157 343 154 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 347 154 moveto +343 144 lineto +340 154 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 347 154 moveto +343 144 lineto +340 154 lineto +closepath +stroke +end grestore + +% iml +gsave 10 dict begin +newpath 413 90 moveto +273 90 lineto +273 54 lineto +413 54 lineto +closepath +stroke +gsave 10 dict begin +280 67 moveto +(intermediate language) +[3.84 6.96 3.84 6.24 4.8 10.8 6.24 6.96 3.84 6.24 3.84 6.24 3.6 3.84 6.24 6.96 6.96 6.96 6.24 6.72 6.24] +xshow +end grestore +end grestore + +% translation -> iml +newpath 343 108 moveto +343 105 343 103 343 100 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 347 100 moveto +343 90 lineto +340 100 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 347 100 moveto +343 90 lineto +340 100 lineto +closepath +stroke +end grestore + +% iml -> serialization +newpath 272 72 moveto +262 72 251 72 241 72 curveto +stroke +gsave 10 dict begin +solid +1 setlinewidth +0.000 0.000 0.000 edgecolor +newpath 241 69 moveto +231 72 lineto +241 76 lineto +closepath +fill +0.000 0.000 0.000 edgecolor +newpath 241 69 moveto +231 72 lineto +241 76 lineto +closepath +stroke +end grestore +endpage +showpage +grestore +%%PageTrailer +%%EndPage: 1 +%%Trailer +%%Pages: 1 +end +restore +%%EOF diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Codegen/style.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Codegen/style.sty Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,63 @@ + +%% toc +\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} +\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} + +%% paragraphs +\setlength{\parindent}{1em} + +%% references +\newcommand{\secref}[1]{\S\ref{#1}} + +%% logical markup +\newcommand{\strong}[1]{{\bfseries {#1}}} +\newcommand{\qn}[1]{\emph{#1}} + +%% typographic conventions +\newcommand{\qt}[1]{``{#1}''} + +%% verbatim text +\newcommand{\isatypewriter}{\fontsize{9pt}{0pt}\tt\renewcommand{\baselinestretch}{1}\setlength{\baselineskip}{9pt}} + +%% quote environment +\isakeeptag{quote} +\renewenvironment{quote} + {\list{}{\leftmargin2em\rightmargin0pt}\parindent0pt\parskip0pt\item\relax} + {\endlist} +\renewcommand{\isatagquote}{\begin{quote}} +\renewcommand{\endisatagquote}{\end{quote}} +\newcommand{\quotebreak}{\\[1.2ex]} + +\isakeeptag{quotett} +\renewcommand{\isatagquotett}{\begin{quote}\isabellestyle{tt}\isastyle} +\renewcommand{\endisatagquotett}{\end{quote}\isabellestyle{it}\isastyle} + +%% a trick +\newcommand{\isasymSML}{SML} + +%% presentation +\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} + +%% character detail +\renewcommand{\isadigit}[1]{\isamath{#1}} +\binperiod +\underscoreoff + +%% format +\pagestyle{headings} +\isabellestyle{it} + +%% ml reference +\newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup} + +\isakeeptag{mlref} +\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small} +\renewcommand{\endisatagmlref}{\endgroup} + +\isabellestyle{it} + + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "implementation" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Dirs --- a/doc-src/Dirs Wed Mar 04 11:05:02 2009 +0100 +++ b/doc-src/Dirs Wed Mar 04 11:05:29 2009 +0100 @@ -1,1 +1,1 @@ -Ref System Logics HOL ZF Inductive TutorialI IsarOverview IsarRef IsarImplementation Locales LaTeXsugar IsarAdvanced/Classes IsarAdvanced/Codegen IsarAdvanced/Functions +Intro Ref System Logics HOL ZF Inductive TutorialI IsarOverview IsarRef IsarImplementation Locales LaTeXsugar Classes Codegen Functions diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/IsaMakefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/IsaMakefile Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,33 @@ + +## targets + +default: Thy +images: +test: Thy + +all: images test + + +## global settings + +SRC = $(ISABELLE_HOME)/src +OUT = $(ISABELLE_OUTPUT) +LOG = $(OUT)/log + +USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -d false -C false -D document + + +## Thy + +THY = $(LOG)/HOL-Thy.gz + +Thy: $(THY) + +$(THY): Thy/ROOT.ML Thy/Functions.thy + @$(USEDIR) HOL Thy + + +## clean + +clean: + @rm -f $(THY) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/Makefile Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,38 @@ +# +# $Id$ +# + +## targets + +default: dvi + + +## dependencies + +include ../Makefile.in + +NAME = functions + +FILES = $(NAME).tex Thy/document/Functions.tex intro.tex conclusion.tex \ + style.sty ../iman.sty ../extra.sty ../isar.sty \ + ../isabelle.sty ../isabellesym.sty ../pdfsetup.sty \ + ../manual.bib ../proof.sty + +dvi: $(NAME).dvi + +$(NAME).dvi: $(FILES) isabelle_isar.eps + $(LATEX) $(NAME) + $(BIBTEX) $(NAME) + $(LATEX) $(NAME) + $(LATEX) $(NAME) + +pdf: $(NAME).pdf + +$(NAME).pdf: $(FILES) isabelle_isar.pdf + $(PDFLATEX) $(NAME) + $(BIBTEX) $(NAME) + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) + $(FIXBOOKMARKS) $(NAME).out + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/Thy/Functions.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/Thy/Functions.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,1264 @@ +(* Title: doc-src/IsarAdvanced/Functions/Thy/Fundefs.thy + Author: Alexander Krauss, TU Muenchen + +Tutorial for function definitions with the new "function" package. +*) + +theory Functions +imports Main +begin + +section {* Function Definitions for Dummies *} + +text {* + In most cases, defining a recursive function is just as simple as other definitions: +*} + +fun fib :: "nat \ nat" +where + "fib 0 = 1" +| "fib (Suc 0) = 1" +| "fib (Suc (Suc n)) = fib n + fib (Suc n)" + +text {* + The syntax is rather self-explanatory: We introduce a function by + giving its name, its type, + and a set of defining recursive equations. + If we leave out the type, the most general type will be + inferred, which can sometimes lead to surprises: Since both @{term + "1::nat"} and @{text "+"} are overloaded, we would end up + with @{text "fib :: nat \ 'a::{one,plus}"}. +*} + +text {* + The function always terminates, since its argument gets smaller in + every recursive call. + Since HOL is a logic of total functions, termination is a + fundamental requirement to prevent inconsistencies\footnote{From the + \qt{definition} @{text "f(n) = f(n) + 1"} we could prove + @{text "0 = 1"} by subtracting @{text "f(n)"} on both sides.}. + Isabelle tries to prove termination automatically when a definition + is made. In \S\ref{termination}, we will look at cases where this + fails and see what to do then. +*} + +subsection {* Pattern matching *} + +text {* \label{patmatch} + Like in functional programming, we can use pattern matching to + define functions. At the moment we will only consider \emph{constructor + patterns}, which only consist of datatype constructors and + variables. Furthermore, patterns must be linear, i.e.\ all variables + on the left hand side of an equation must be distinct. In + \S\ref{genpats} we discuss more general pattern matching. + + If patterns overlap, the order of the equations is taken into + account. The following function inserts a fixed element between any + two elements of a list: +*} + +fun sep :: "'a \ 'a list \ 'a list" +where + "sep a (x#y#xs) = x # a # sep a (y # xs)" +| "sep a xs = xs" + +text {* + Overlapping patterns are interpreted as \qt{increments} to what is + already there: The second equation is only meant for the cases where + the first one does not match. Consequently, Isabelle replaces it + internally by the remaining cases, making the patterns disjoint: +*} + +thm sep.simps + +text {* @{thm [display] sep.simps[no_vars]} *} + +text {* + \noindent The equations from function definitions are automatically used in + simplification: +*} + +lemma "sep 0 [1, 2, 3] = [1, 0, 2, 0, 3]" +by simp + +subsection {* Induction *} + +text {* + + Isabelle provides customized induction rules for recursive + functions. These rules follow the recursive structure of the + definition. Here is the rule @{text sep.induct} arising from the + above definition of @{const sep}: + + @{thm [display] sep.induct} + + We have a step case for list with at least two elements, and two + base cases for the zero- and the one-element list. Here is a simple + proof about @{const sep} and @{const map} +*} + +lemma "map f (sep x ys) = sep (f x) (map f ys)" +apply (induct x ys rule: sep.induct) + +txt {* + We get three cases, like in the definition. + + @{subgoals [display]} +*} + +apply auto +done +text {* + + With the \cmd{fun} command, you can define about 80\% of the + functions that occur in practice. The rest of this tutorial explains + the remaining 20\%. +*} + + +section {* fun vs.\ function *} + +text {* + The \cmd{fun} command provides a + convenient shorthand notation for simple function definitions. In + this mode, Isabelle tries to solve all the necessary proof obligations + automatically. If any proof fails, the definition is + rejected. This can either mean that the definition is indeed faulty, + or that the default proof procedures are just not smart enough (or + rather: not designed) to handle the definition. + + By expanding the abbreviation to the more verbose \cmd{function} command, these proof obligations become visible and can be analyzed or + solved manually. The expansion from \cmd{fun} to \cmd{function} is as follows: + +\end{isamarkuptext} + + +\[\left[\;\begin{minipage}{0.25\textwidth}\vspace{6pt} +\cmd{fun} @{text "f :: \"}\\% +\cmd{where}\\% +\hspace*{2ex}{\it equations}\\% +\hspace*{2ex}\vdots\vspace*{6pt} +\end{minipage}\right] +\quad\equiv\quad +\left[\;\begin{minipage}{0.48\textwidth}\vspace{6pt} +\cmd{function} @{text "("}\cmd{sequential}@{text ") f :: \"}\\% +\cmd{where}\\% +\hspace*{2ex}{\it equations}\\% +\hspace*{2ex}\vdots\\% +\cmd{by} @{text "pat_completeness auto"}\\% +\cmd{termination by} @{text "lexicographic_order"}\vspace{6pt} +\end{minipage} +\right]\] + +\begin{isamarkuptext} + \vspace*{1em} + \noindent Some details have now become explicit: + + \begin{enumerate} + \item The \cmd{sequential} option enables the preprocessing of + pattern overlaps which we already saw. Without this option, the equations + must already be disjoint and complete. The automatic completion only + works with constructor patterns. + + \item A function definition produces a proof obligation which + expresses completeness and compatibility of patterns (we talk about + this later). The combination of the methods @{text "pat_completeness"} and + @{text "auto"} is used to solve this proof obligation. + + \item A termination proof follows the definition, started by the + \cmd{termination} command. This will be explained in \S\ref{termination}. + \end{enumerate} + Whenever a \cmd{fun} command fails, it is usually a good idea to + expand the syntax to the more verbose \cmd{function} form, to see + what is actually going on. + *} + + +section {* Termination *} + +text {*\label{termination} + The method @{text "lexicographic_order"} is the default method for + termination proofs. It can prove termination of a + certain class of functions by searching for a suitable lexicographic + combination of size measures. Of course, not all functions have such + a simple termination argument. For them, we can specify the termination + relation manually. +*} + +subsection {* The {\tt relation} method *} +text{* + Consider the following function, which sums up natural numbers up to + @{text "N"}, using a counter @{text "i"}: +*} + +function sum :: "nat \ nat \ nat" +where + "sum i N = (if i > N then 0 else i + sum (Suc i) N)" +by pat_completeness auto + +text {* + \noindent The @{text "lexicographic_order"} method fails on this example, because none of the + arguments decreases in the recursive call, with respect to the standard size ordering. + To prove termination manually, we must provide a custom wellfounded relation. + + The termination argument for @{text "sum"} is based on the fact that + the \emph{difference} between @{text "i"} and @{text "N"} gets + smaller in every step, and that the recursion stops when @{text "i"} + is greater than @{text "N"}. Phrased differently, the expression + @{text "N + 1 - i"} always decreases. + + We can use this expression as a measure function suitable to prove termination. +*} + +termination sum +apply (relation "measure (\(i,N). N + 1 - i)") + +txt {* + The \cmd{termination} command sets up the termination goal for the + specified function @{text "sum"}. If the function name is omitted, it + implicitly refers to the last function definition. + + The @{text relation} method takes a relation of + type @{typ "('a \ 'a) set"}, where @{typ "'a"} is the argument type of + the function. If the function has multiple curried arguments, then + these are packed together into a tuple, as it happened in the above + example. + + The predefined function @{term[source] "measure :: ('a \ nat) \ ('a \ 'a) set"} constructs a + wellfounded relation from a mapping into the natural numbers (a + \emph{measure function}). + + After the invocation of @{text "relation"}, we must prove that (a) + the relation we supplied is wellfounded, and (b) that the arguments + of recursive calls indeed decrease with respect to the + relation: + + @{subgoals[display,indent=0]} + + These goals are all solved by @{text "auto"}: +*} + +apply auto +done + +text {* + Let us complicate the function a little, by adding some more + recursive calls: +*} + +function foo :: "nat \ nat \ nat" +where + "foo i N = (if i > N + then (if N = 0 then 0 else foo 0 (N - 1)) + else i + foo (Suc i) N)" +by pat_completeness auto + +text {* + When @{text "i"} has reached @{text "N"}, it starts at zero again + and @{text "N"} is decremented. + This corresponds to a nested + loop where one index counts up and the other down. Termination can + be proved using a lexicographic combination of two measures, namely + the value of @{text "N"} and the above difference. The @{const + "measures"} combinator generalizes @{text "measure"} by taking a + list of measure functions. +*} + +termination +by (relation "measures [\(i, N). N, \(i,N). N + 1 - i]") auto + +subsection {* How @{text "lexicographic_order"} works *} + +(*fun fails :: "nat \ nat list \ nat" +where + "fails a [] = a" +| "fails a (x#xs) = fails (x + a) (x # xs)" +*) + +text {* + To see how the automatic termination proofs work, let's look at an + example where it fails\footnote{For a detailed discussion of the + termination prover, see \cite{bulwahnKN07}}: + +\end{isamarkuptext} +\cmd{fun} @{text "fails :: \"nat \ nat list \ nat\""}\\% +\cmd{where}\\% +\hspace*{2ex}@{text "\"fails a [] = a\""}\\% +|\hspace*{1.5ex}@{text "\"fails a (x#xs) = fails (x + a) (x#xs)\""}\\ +\begin{isamarkuptext} + +\noindent Isabelle responds with the following error: + +\begin{isabelle} +*** Unfinished subgoals:\newline +*** (a, 1, <):\newline +*** \ 1.~@{text "\x. x = 0"}\newline +*** (a, 1, <=):\newline +*** \ 1.~False\newline +*** (a, 2, <):\newline +*** \ 1.~False\newline +*** Calls:\newline +*** a) @{text "(a, x # xs) -->> (x + a, x # xs)"}\newline +*** Measures:\newline +*** 1) @{text "\x. size (fst x)"}\newline +*** 2) @{text "\x. size (snd x)"}\newline +*** Result matrix:\newline +*** \ \ \ \ 1\ \ 2 \newline +*** a: ? <= \newline +*** Could not find lexicographic termination order.\newline +*** At command "fun".\newline +\end{isabelle} +*} + + +text {* + The key to this error message is the matrix at the bottom. The rows + of that matrix correspond to the different recursive calls (In our + case, there is just one). The columns are the function's arguments + (expressed through different measure functions, which map the + argument tuple to a natural number). + + The contents of the matrix summarize what is known about argument + descents: The second argument has a weak descent (@{text "<="}) at the + recursive call, and for the first argument nothing could be proved, + which is expressed by @{text "?"}. In general, there are the values + @{text "<"}, @{text "<="} and @{text "?"}. + + For the failed proof attempts, the unfinished subgoals are also + printed. Looking at these will often point to a missing lemma. + +% As a more real example, here is quicksort: +*} +(* +function qs :: "nat list \ nat list" +where + "qs [] = []" +| "qs (x#xs) = qs [y\xs. y < x] @ x # qs [y\xs. y \ x]" +by pat_completeness auto + +termination apply lexicographic_order + +text {* If we try @{text "lexicographic_order"} method, we get the + following error *} +termination by (lexicographic_order simp:l2) + +lemma l: "x \ y \ x < Suc y" by arith + +function + +*) + +section {* Mutual Recursion *} + +text {* + If two or more functions call one another mutually, they have to be defined + in one step. Here are @{text "even"} and @{text "odd"}: +*} + +function even :: "nat \ bool" + and odd :: "nat \ bool" +where + "even 0 = True" +| "odd 0 = False" +| "even (Suc n) = odd n" +| "odd (Suc n) = even n" +by pat_completeness auto + +text {* + To eliminate the mutual dependencies, Isabelle internally + creates a single function operating on the sum + type @{typ "nat + nat"}. Then, @{const even} and @{const odd} are + defined as projections. Consequently, termination has to be proved + simultaneously for both functions, by specifying a measure on the + sum type: +*} + +termination +by (relation "measure (\x. case x of Inl n \ n | Inr n \ n)") auto + +text {* + We could also have used @{text lexicographic_order}, which + supports mutual recursive termination proofs to a certain extent. +*} + +subsection {* Induction for mutual recursion *} + +text {* + + When functions are mutually recursive, proving properties about them + generally requires simultaneous induction. The induction rule @{text "even_odd.induct"} + generated from the above definition reflects this. + + Let us prove something about @{const even} and @{const odd}: +*} + +lemma even_odd_mod2: + "even n = (n mod 2 = 0)" + "odd n = (n mod 2 = 1)" + +txt {* + We apply simultaneous induction, specifying the induction variable + for both goals, separated by \cmd{and}: *} + +apply (induct n and n rule: even_odd.induct) + +txt {* + We get four subgoals, which correspond to the clauses in the + definition of @{const even} and @{const odd}: + @{subgoals[display,indent=0]} + Simplification solves the first two goals, leaving us with two + statements about the @{text "mod"} operation to prove: +*} + +apply simp_all + +txt {* + @{subgoals[display,indent=0]} + + \noindent These can be handled by Isabelle's arithmetic decision procedures. + +*} + +apply arith +apply arith +done + +text {* + In proofs like this, the simultaneous induction is really essential: + Even if we are just interested in one of the results, the other + one is necessary to strengthen the induction hypothesis. If we leave + out the statement about @{const odd} and just write @{term True} instead, + the same proof fails: +*} + +lemma failed_attempt: + "even n = (n mod 2 = 0)" + "True" +apply (induct n rule: even_odd.induct) + +txt {* + \noindent Now the third subgoal is a dead end, since we have no + useful induction hypothesis available: + + @{subgoals[display,indent=0]} +*} + +oops + +section {* General pattern matching *} +text{*\label{genpats} *} + +subsection {* Avoiding automatic pattern splitting *} + +text {* + + Up to now, we used pattern matching only on datatypes, and the + patterns were always disjoint and complete, and if they weren't, + they were made disjoint automatically like in the definition of + @{const "sep"} in \S\ref{patmatch}. + + This automatic splitting can significantly increase the number of + equations involved, and this is not always desirable. The following + example shows the problem: + + Suppose we are modeling incomplete knowledge about the world by a + three-valued datatype, which has values @{term "T"}, @{term "F"} + and @{term "X"} for true, false and uncertain propositions, respectively. +*} + +datatype P3 = T | F | X + +text {* \noindent Then the conjunction of such values can be defined as follows: *} + +fun And :: "P3 \ P3 \ P3" +where + "And T p = p" +| "And p T = p" +| "And p F = F" +| "And F p = F" +| "And X X = X" + + +text {* + This definition is useful, because the equations can directly be used + as simplification rules. But the patterns overlap: For example, + the expression @{term "And T T"} is matched by both the first and + the second equation. By default, Isabelle makes the patterns disjoint by + splitting them up, producing instances: +*} + +thm And.simps + +text {* + @{thm[indent=4] And.simps} + + \vspace*{1em} + \noindent There are several problems with this: + + \begin{enumerate} + \item If the datatype has many constructors, there can be an + explosion of equations. For @{const "And"}, we get seven instead of + five equations, which can be tolerated, but this is just a small + example. + + \item Since splitting makes the equations \qt{less general}, they + do not always match in rewriting. While the term @{term "And x F"} + can be simplified to @{term "F"} with the original equations, a + (manual) case split on @{term "x"} is now necessary. + + \item The splitting also concerns the induction rule @{text + "And.induct"}. Instead of five premises it now has seven, which + means that our induction proofs will have more cases. + + \item In general, it increases clarity if we get the same definition + back which we put in. + \end{enumerate} + + If we do not want the automatic splitting, we can switch it off by + leaving out the \cmd{sequential} option. However, we will have to + prove that our pattern matching is consistent\footnote{This prevents + us from defining something like @{term "f x = True"} and @{term "f x + = False"} simultaneously.}: +*} + +function And2 :: "P3 \ P3 \ P3" +where + "And2 T p = p" +| "And2 p T = p" +| "And2 p F = F" +| "And2 F p = F" +| "And2 X X = X" + +txt {* + \noindent Now let's look at the proof obligations generated by a + function definition. In this case, they are: + + @{subgoals[display,indent=0]}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em} + + The first subgoal expresses the completeness of the patterns. It has + the form of an elimination rule and states that every @{term x} of + the function's input type must match at least one of the patterns\footnote{Completeness could + be equivalently stated as a disjunction of existential statements: +@{term "(\p. x = (T, p)) \ (\p. x = (p, T)) \ (\p. x = (p, F)) \ + (\p. x = (F, p)) \ (x = (X, X))"}, and you can use the method @{text atomize_elim} to get that form instead.}. If the patterns just involve + datatypes, we can solve it with the @{text "pat_completeness"} + method: +*} + +apply pat_completeness + +txt {* + The remaining subgoals express \emph{pattern compatibility}. We do + allow that an input value matches multiple patterns, but in this + case, the result (i.e.~the right hand sides of the equations) must + also be equal. For each pair of two patterns, there is one such + subgoal. Usually this needs injectivity of the constructors, which + is used automatically by @{text "auto"}. +*} + +by auto + + +subsection {* Non-constructor patterns *} + +text {* + Most of Isabelle's basic types take the form of inductive datatypes, + and usually pattern matching works on the constructors of such types. + However, this need not be always the case, and the \cmd{function} + command handles other kind of patterns, too. + + One well-known instance of non-constructor patterns are + so-called \emph{$n+k$-patterns}, which are a little controversial in + the functional programming world. Here is the initial fibonacci + example with $n+k$-patterns: +*} + +function fib2 :: "nat \ nat" +where + "fib2 0 = 1" +| "fib2 1 = 1" +| "fib2 (n + 2) = fib2 n + fib2 (Suc n)" + +(*<*)ML_val "goals_limit := 1"(*>*) +txt {* + This kind of matching is again justified by the proof of pattern + completeness and compatibility. + The proof obligation for pattern completeness states that every natural number is + either @{term "0::nat"}, @{term "1::nat"} or @{term "n + + (2::nat)"}: + + @{subgoals[display,indent=0]} + + This is an arithmetic triviality, but unfortunately the + @{text arith} method cannot handle this specific form of an + elimination rule. However, we can use the method @{text + "atomize_elim"} to do an ad-hoc conversion to a disjunction of + existentials, which can then be solved by the arithmetic decision procedure. + Pattern compatibility and termination are automatic as usual. +*} +(*<*)ML_val "goals_limit := 10"(*>*) +apply atomize_elim +apply arith +apply auto +done +termination by lexicographic_order +text {* + We can stretch the notion of pattern matching even more. The + following function is not a sensible functional program, but a + perfectly valid mathematical definition: +*} + +function ev :: "nat \ bool" +where + "ev (2 * n) = True" +| "ev (2 * n + 1) = False" +apply atomize_elim +by arith+ +termination by (relation "{}") simp + +text {* + This general notion of pattern matching gives you a certain freedom + in writing down specifications. However, as always, such freedom should + be used with care: + + If we leave the area of constructor + patterns, we have effectively departed from the world of functional + programming. This means that it is no longer possible to use the + code generator, and expect it to generate ML code for our + definitions. Also, such a specification might not work very well together with + simplification. Your mileage may vary. +*} + + +subsection {* Conditional equations *} + +text {* + The function package also supports conditional equations, which are + similar to guards in a language like Haskell. Here is Euclid's + algorithm written with conditional patterns\footnote{Note that the + patterns are also overlapping in the base case}: +*} + +function gcd :: "nat \ nat \ nat" +where + "gcd x 0 = x" +| "gcd 0 y = y" +| "x < y \ gcd (Suc x) (Suc y) = gcd (Suc x) (y - x)" +| "\ x < y \ gcd (Suc x) (Suc y) = gcd (x - y) (Suc y)" +by (atomize_elim, auto, arith) +termination by lexicographic_order + +text {* + By now, you can probably guess what the proof obligations for the + pattern completeness and compatibility look like. + + Again, functions with conditional patterns are not supported by the + code generator. +*} + + +subsection {* Pattern matching on strings *} + +text {* + As strings (as lists of characters) are normal datatypes, pattern + matching on them is possible, but somewhat problematic. Consider the + following definition: + +\end{isamarkuptext} +\noindent\cmd{fun} @{text "check :: \"string \ bool\""}\\% +\cmd{where}\\% +\hspace*{2ex}@{text "\"check (''good'') = True\""}\\% +@{text "| \"check s = False\""} +\begin{isamarkuptext} + + \noindent An invocation of the above \cmd{fun} command does not + terminate. What is the problem? Strings are lists of characters, and + characters are a datatype with a lot of constructors. Splitting the + catch-all pattern thus leads to an explosion of cases, which cannot + be handled by Isabelle. + + There are two things we can do here. Either we write an explicit + @{text "if"} on the right hand side, or we can use conditional patterns: +*} + +function check :: "string \ bool" +where + "check (''good'') = True" +| "s \ ''good'' \ check s = False" +by auto + + +section {* Partiality *} + +text {* + In HOL, all functions are total. A function @{term "f"} applied to + @{term "x"} always has the value @{term "f x"}, and there is no notion + of undefinedness. + This is why we have to do termination + proofs when defining functions: The proof justifies that the + function can be defined by wellfounded recursion. + + However, the \cmd{function} package does support partiality to a + certain extent. Let's look at the following function which looks + for a zero of a given function f. +*} + +function (*<*)(domintros, tailrec)(*>*)findzero :: "(nat \ nat) \ nat \ nat" +where + "findzero f n = (if f n = 0 then n else findzero f (Suc n))" +by pat_completeness auto +(*<*)declare findzero.simps[simp del](*>*) + +text {* + \noindent Clearly, any attempt of a termination proof must fail. And without + that, we do not get the usual rules @{text "findzero.simps"} and + @{text "findzero.induct"}. So what was the definition good for at all? +*} + +subsection {* Domain predicates *} + +text {* + The trick is that Isabelle has not only defined the function @{const findzero}, but also + a predicate @{term "findzero_dom"} that characterizes the values where the function + terminates: the \emph{domain} of the function. If we treat a + partial function just as a total function with an additional domain + predicate, we can derive simplification and + induction rules as we do for total functions. They are guarded + by domain conditions and are called @{text psimps} and @{text + pinduct}: +*} + +text {* + \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.psimps}\end{minipage} + \hfill(@{text "findzero.psimps"}) + \vspace{1em} + + \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.pinduct}\end{minipage} + \hfill(@{text "findzero.pinduct"}) +*} + +text {* + Remember that all we + are doing here is use some tricks to make a total function appear + as if it was partial. We can still write the term @{term "findzero + (\x. 1) 0"} and like any other term of type @{typ nat} it is equal + to some natural number, although we might not be able to find out + which one. The function is \emph{underdefined}. + + But it is defined enough to prove something interesting about it. We + can prove that if @{term "findzero f n"} + terminates, it indeed returns a zero of @{term f}: +*} + +lemma findzero_zero: "findzero_dom (f, n) \ f (findzero f n) = 0" + +txt {* \noindent We apply induction as usual, but using the partial induction + rule: *} + +apply (induct f n rule: findzero.pinduct) + +txt {* \noindent This gives the following subgoals: + + @{subgoals[display,indent=0]} + + \noindent The hypothesis in our lemma was used to satisfy the first premise in + the induction rule. However, we also get @{term + "findzero_dom (f, n)"} as a local assumption in the induction step. This + allows to unfold @{term "findzero f n"} using the @{text psimps} + rule, and the rest is trivial. Since the @{text psimps} rules carry the + @{text "[simp]"} attribute by default, we just need a single step: + *} +apply simp +done + +text {* + Proofs about partial functions are often not harder than for total + functions. Fig.~\ref{findzero_isar} shows a slightly more + complicated proof written in Isar. It is verbose enough to show how + partiality comes into play: From the partial induction, we get an + additional domain condition hypothesis. Observe how this condition + is applied when calls to @{term findzero} are unfolded. +*} + +text_raw {* +\begin{figure} +\hrule\vspace{6pt} +\begin{minipage}{0.8\textwidth} +\isabellestyle{it} +\isastyle\isamarkuptrue +*} +lemma "\findzero_dom (f, n); x \ {n ..< findzero f n}\ \ f x \ 0" +proof (induct rule: findzero.pinduct) + fix f n assume dom: "findzero_dom (f, n)" + and IH: "\f n \ 0; x \ {Suc n ..< findzero f (Suc n)}\ \ f x \ 0" + and x_range: "x \ {n ..< findzero f n}" + have "f n \ 0" + proof + assume "f n = 0" + with dom have "findzero f n = n" by simp + with x_range show False by auto + qed + + from x_range have "x = n \ x \ {Suc n ..< findzero f n}" by auto + thus "f x \ 0" + proof + assume "x = n" + with `f n \ 0` show ?thesis by simp + next + assume "x \ {Suc n ..< findzero f n}" + with dom and `f n \ 0` have "x \ {Suc n ..< findzero f (Suc n)}" by simp + with IH and `f n \ 0` + show ?thesis by simp + qed +qed +text_raw {* +\isamarkupfalse\isabellestyle{tt} +\end{minipage}\vspace{6pt}\hrule +\caption{A proof about a partial function}\label{findzero_isar} +\end{figure} +*} + +subsection {* Partial termination proofs *} + +text {* + Now that we have proved some interesting properties about our + function, we should turn to the domain predicate and see if it is + actually true for some values. Otherwise we would have just proved + lemmas with @{term False} as a premise. + + Essentially, we need some introduction rules for @{text + findzero_dom}. The function package can prove such domain + introduction rules automatically. But since they are not used very + often (they are almost never needed if the function is total), this + functionality is disabled by default for efficiency reasons. So we have to go + back and ask for them explicitly by passing the @{text + "(domintros)"} option to the function package: + +\vspace{1ex} +\noindent\cmd{function} @{text "(domintros) findzero :: \"(nat \ nat) \ nat \ nat\""}\\% +\cmd{where}\isanewline% +\ \ \ldots\\ + + \noindent Now the package has proved an introduction rule for @{text findzero_dom}: +*} + +thm findzero.domintros + +text {* + @{thm[display] findzero.domintros} + + Domain introduction rules allow to show that a given value lies in the + domain of a function, if the arguments of all recursive calls + are in the domain as well. They allow to do a \qt{single step} in a + termination proof. Usually, you want to combine them with a suitable + induction principle. + + Since our function increases its argument at recursive calls, we + need an induction principle which works \qt{backwards}. We will use + @{text inc_induct}, which allows to do induction from a fixed number + \qt{downwards}: + + \begin{center}@{thm inc_induct}\hfill(@{text "inc_induct"})\end{center} + + Figure \ref{findzero_term} gives a detailed Isar proof of the fact + that @{text findzero} terminates if there is a zero which is greater + or equal to @{term n}. First we derive two useful rules which will + solve the base case and the step case of the induction. The + induction is then straightforward, except for the unusual induction + principle. + +*} + +text_raw {* +\begin{figure} +\hrule\vspace{6pt} +\begin{minipage}{0.8\textwidth} +\isabellestyle{it} +\isastyle\isamarkuptrue +*} +lemma findzero_termination: + assumes "x \ n" and "f x = 0" + shows "findzero_dom (f, n)" +proof - + have base: "findzero_dom (f, x)" + by (rule findzero.domintros) (simp add:`f x = 0`) + + have step: "\i. findzero_dom (f, Suc i) + \ findzero_dom (f, i)" + by (rule findzero.domintros) simp + + from `x \ n` show ?thesis + proof (induct rule:inc_induct) + show "findzero_dom (f, x)" by (rule base) + next + fix i assume "findzero_dom (f, Suc i)" + thus "findzero_dom (f, i)" by (rule step) + qed +qed +text_raw {* +\isamarkupfalse\isabellestyle{tt} +\end{minipage}\vspace{6pt}\hrule +\caption{Termination proof for @{text findzero}}\label{findzero_term} +\end{figure} +*} + +text {* + Again, the proof given in Fig.~\ref{findzero_term} has a lot of + detail in order to explain the principles. Using more automation, we + can also have a short proof: +*} + +lemma findzero_termination_short: + assumes zero: "x >= n" + assumes [simp]: "f x = 0" + shows "findzero_dom (f, n)" +using zero +by (induct rule:inc_induct) (auto intro: findzero.domintros) + +text {* + \noindent It is simple to combine the partial correctness result with the + termination lemma: +*} + +lemma findzero_total_correctness: + "f x = 0 \ f (findzero f 0) = 0" +by (blast intro: findzero_zero findzero_termination) + +subsection {* Definition of the domain predicate *} + +text {* + Sometimes it is useful to know what the definition of the domain + predicate looks like. Actually, @{text findzero_dom} is just an + abbreviation: + + @{abbrev[display] findzero_dom} + + The domain predicate is the \emph{accessible part} of a relation @{const + findzero_rel}, which was also created internally by the function + package. @{const findzero_rel} is just a normal + inductive predicate, so we can inspect its definition by + looking at the introduction rules @{text findzero_rel.intros}. + In our case there is just a single rule: + + @{thm[display] findzero_rel.intros} + + The predicate @{const findzero_rel} + describes the \emph{recursion relation} of the function + definition. The recursion relation is a binary relation on + the arguments of the function that relates each argument to its + recursive calls. In general, there is one introduction rule for each + recursive call. + + The predicate @{term "accp findzero_rel"} is the accessible part of + that relation. An argument belongs to the accessible part, if it can + be reached in a finite number of steps (cf.~its definition in @{text + "Wellfounded.thy"}). + + Since the domain predicate is just an abbreviation, you can use + lemmas for @{const accp} and @{const findzero_rel} directly. Some + lemmas which are occasionally useful are @{text accpI}, @{text + accp_downward}, and of course the introduction and elimination rules + for the recursion relation @{text "findzero.intros"} and @{text "findzero.cases"}. +*} + +(*lemma findzero_nicer_domintros: + "f x = 0 \ findzero_dom (f, x)" + "findzero_dom (f, Suc x) \ findzero_dom (f, x)" +by (rule accpI, erule findzero_rel.cases, auto)+ +*) + +subsection {* A Useful Special Case: Tail recursion *} + +text {* + The domain predicate is our trick that allows us to model partiality + in a world of total functions. The downside of this is that we have + to carry it around all the time. The termination proof above allowed + us to replace the abstract @{term "findzero_dom (f, n)"} by the more + concrete @{term "(x \ n \ f x = (0::nat))"}, but the condition is still + there and can only be discharged for special cases. + In particular, the domain predicate guards the unfolding of our + function, since it is there as a condition in the @{text psimp} + rules. + + Now there is an important special case: We can actually get rid + of the condition in the simplification rules, \emph{if the function + is tail-recursive}. The reason is that for all tail-recursive + equations there is a total function satisfying them, even if they + are non-terminating. + +% A function is tail recursive, if each call to the function is either +% equal +% +% So the outer form of the +% +%if it can be written in the following +% form: +% {term[display] "f x = (if COND x then BASE x else f (LOOP x))"} + + + The function package internally does the right construction and can + derive the unconditional simp rules, if we ask it to do so. Luckily, + our @{const "findzero"} function is tail-recursive, so we can just go + back and add another option to the \cmd{function} command: + +\vspace{1ex} +\noindent\cmd{function} @{text "(domintros, tailrec) findzero :: \"(nat \ nat) \ nat \ nat\""}\\% +\cmd{where}\isanewline% +\ \ \ldots\\% + + + \noindent Now, we actually get unconditional simplification rules, even + though the function is partial: +*} + +thm findzero.simps + +text {* + @{thm[display] findzero.simps} + + \noindent Of course these would make the simplifier loop, so we better remove + them from the simpset: +*} + +declare findzero.simps[simp del] + +text {* + Getting rid of the domain conditions in the simplification rules is + not only useful because it simplifies proofs. It is also required in + order to use Isabelle's code generator to generate ML code + from a function definition. + Since the code generator only works with equations, it cannot be + used with @{text "psimp"} rules. Thus, in order to generate code for + partial functions, they must be defined as a tail recursion. + Luckily, many functions have a relatively natural tail recursive + definition. +*} + +section {* Nested recursion *} + +text {* + Recursive calls which are nested in one another frequently cause + complications, since their termination proof can depend on a partial + correctness property of the function itself. + + As a small example, we define the \qt{nested zero} function: +*} + +function nz :: "nat \ nat" +where + "nz 0 = 0" +| "nz (Suc n) = nz (nz n)" +by pat_completeness auto + +text {* + If we attempt to prove termination using the identity measure on + naturals, this fails: +*} + +termination + apply (relation "measure (\n. n)") + apply auto + +txt {* + We get stuck with the subgoal + + @{subgoals[display]} + + Of course this statement is true, since we know that @{const nz} is + the zero function. And in fact we have no problem proving this + property by induction. +*} +(*<*)oops(*>*) +lemma nz_is_zero: "nz_dom n \ nz n = 0" + by (induct rule:nz.pinduct) auto + +text {* + We formulate this as a partial correctness lemma with the condition + @{term "nz_dom n"}. This allows us to prove it with the @{text + pinduct} rule before we have proved termination. With this lemma, + the termination proof works as expected: +*} + +termination + by (relation "measure (\n. n)") (auto simp: nz_is_zero) + +text {* + As a general strategy, one should prove the statements needed for + termination as a partial property first. Then they can be used to do + the termination proof. This also works for less trivial + examples. Figure \ref{f91} defines the 91-function, a well-known + challenge problem due to John McCarthy, and proves its termination. +*} + +text_raw {* +\begin{figure} +\hrule\vspace{6pt} +\begin{minipage}{0.8\textwidth} +\isabellestyle{it} +\isastyle\isamarkuptrue +*} + +function f91 :: "nat \ nat" +where + "f91 n = (if 100 < n then n - 10 else f91 (f91 (n + 11)))" +by pat_completeness auto + +lemma f91_estimate: + assumes trm: "f91_dom n" + shows "n < f91 n + 11" +using trm by induct auto + +termination +proof + let ?R = "measure (\x. 101 - x)" + show "wf ?R" .. + + fix n :: nat assume "\ 100 < n" -- "Assumptions for both calls" + + thus "(n + 11, n) \ ?R" by simp -- "Inner call" + + assume inner_trm: "f91_dom (n + 11)" -- "Outer call" + with f91_estimate have "n + 11 < f91 (n + 11) + 11" . + with `\ 100 < n` show "(f91 (n + 11), n) \ ?R" by simp +qed + +text_raw {* +\isamarkupfalse\isabellestyle{tt} +\end{minipage} +\vspace{6pt}\hrule +\caption{McCarthy's 91-function}\label{f91} +\end{figure} +*} + + +section {* Higher-Order Recursion *} + +text {* + Higher-order recursion occurs when recursive calls + are passed as arguments to higher-order combinators such as @{const + map}, @{term filter} etc. + As an example, imagine a datatype of n-ary trees: +*} + +datatype 'a tree = + Leaf 'a +| Branch "'a tree list" + + +text {* \noindent We can define a function which swaps the left and right subtrees recursively, using the + list functions @{const rev} and @{const map}: *} + +fun mirror :: "'a tree \ 'a tree" +where + "mirror (Leaf n) = Leaf n" +| "mirror (Branch l) = Branch (rev (map mirror l))" + +text {* + Although the definition is accepted without problems, let us look at the termination proof: +*} + +termination proof + txt {* + + As usual, we have to give a wellfounded relation, such that the + arguments of the recursive calls get smaller. But what exactly are + the arguments of the recursive calls when mirror is given as an + argument to @{const map}? Isabelle gives us the + subgoals + + @{subgoals[display,indent=0]} + + So the system seems to know that @{const map} only + applies the recursive call @{term "mirror"} to elements + of @{term "l"}, which is essential for the termination proof. + + This knowledge about @{const map} is encoded in so-called congruence rules, + which are special theorems known to the \cmd{function} command. The + rule for @{const map} is + + @{thm[display] map_cong} + + You can read this in the following way: Two applications of @{const + map} are equal, if the list arguments are equal and the functions + coincide on the elements of the list. This means that for the value + @{term "map f l"} we only have to know how @{term f} behaves on + the elements of @{term l}. + + Usually, one such congruence rule is + needed for each higher-order construct that is used when defining + new functions. In fact, even basic functions like @{const + If} and @{const Let} are handled by this mechanism. The congruence + rule for @{const If} states that the @{text then} branch is only + relevant if the condition is true, and the @{text else} branch only if it + is false: + + @{thm[display] if_cong} + + Congruence rules can be added to the + function package by giving them the @{term fundef_cong} attribute. + + The constructs that are predefined in Isabelle, usually + come with the respective congruence rules. + But if you define your own higher-order functions, you may have to + state and prove the required congruence rules yourself, if you want to use your + functions in recursive definitions. +*} +(*<*)oops(*>*) + +subsection {* Congruence Rules and Evaluation Order *} + +text {* + Higher order logic differs from functional programming languages in + that it has no built-in notion of evaluation order. A program is + just a set of equations, and it is not specified how they must be + evaluated. + + However for the purpose of function definition, we must talk about + evaluation order implicitly, when we reason about termination. + Congruence rules express that a certain evaluation order is + consistent with the logical definition. + + Consider the following function. +*} + +function f :: "nat \ bool" +where + "f n = (n = 0 \ f (n - 1))" +(*<*)by pat_completeness auto(*>*) + +text {* + For this definition, the termination proof fails. The default configuration + specifies no congruence rule for disjunction. We have to add a + congruence rule that specifies left-to-right evaluation order: + + \vspace{1ex} + \noindent @{thm disj_cong}\hfill(@{text "disj_cong"}) + \vspace{1ex} + + Now the definition works without problems. Note how the termination + proof depends on the extra condition that we get from the congruence + rule. + + However, as evaluation is not a hard-wired concept, we + could just turn everything around by declaring a different + congruence rule. Then we can make the reverse definition: +*} + +lemma disj_cong2[fundef_cong]: + "(\ Q' \ P = P') \ (Q = Q') \ (P \ Q) = (P' \ Q')" + by blast + +fun f' :: "nat \ bool" +where + "f' n = (f' (n - 1) \ n = 0)" + +text {* + \noindent These examples show that, in general, there is no \qt{best} set of + congruence rules. + + However, such tweaking should rarely be necessary in + practice, as most of the time, the default set of congruence rules + works well. +*} + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/Thy/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/Thy/ROOT.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,4 @@ + +(* $Id$ *) + +use_thy "Functions"; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/Thy/document/Functions.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/Thy/document/Functions.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,1985 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Functions}% +% +\isadelimtheory +\isanewline +\isanewline +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Functions\isanewline +\isakeyword{imports}\ Main\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsection{Function Definitions for Dummies% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In most cases, defining a recursive function is just as simple as other definitions:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{fun}\isamarkupfalse% +\ fib\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}fib\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}fib\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}fib\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ fib\ n\ {\isacharplus}\ fib\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +The syntax is rather self-explanatory: We introduce a function by + giving its name, its type, + and a set of defining recursive equations. + If we leave out the type, the most general type will be + inferred, which can sometimes lead to surprises: Since both \isa{{\isadigit{1}}} and \isa{{\isacharplus}} are overloaded, we would end up + with \isa{fib\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ {\isacharprime}a{\isacharcolon}{\isacharcolon}{\isacharbraceleft}one{\isacharcomma}plus{\isacharbraceright}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +The function always terminates, since its argument gets smaller in + every recursive call. + Since HOL is a logic of total functions, termination is a + fundamental requirement to prevent inconsistencies\footnote{From the + \qt{definition} \isa{f{\isacharparenleft}n{\isacharparenright}\ {\isacharequal}\ f{\isacharparenleft}n{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}} we could prove + \isa{{\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}} by subtracting \isa{f{\isacharparenleft}n{\isacharparenright}} on both sides.}. + Isabelle tries to prove termination automatically when a definition + is made. In \S\ref{termination}, we will look at cases where this + fails and see what to do then.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Pattern matching% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\label{patmatch} + Like in functional programming, we can use pattern matching to + define functions. At the moment we will only consider \emph{constructor + patterns}, which only consist of datatype constructors and + variables. Furthermore, patterns must be linear, i.e.\ all variables + on the left hand side of an equation must be distinct. In + \S\ref{genpats} we discuss more general pattern matching. + + If patterns overlap, the order of the equations is taken into + account. The following function inserts a fixed element between any + two elements of a list:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{fun}\isamarkupfalse% +\ sep\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}sep\ a\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}sep\ a\ xs\ \ \ \ \ \ \ {\isacharequal}\ xs{\isachardoublequoteclose}% +\begin{isamarkuptext}% +Overlapping patterns are interpreted as \qt{increments} to what is + already there: The second equation is only meant for the cases where + the first one does not match. Consequently, Isabelle replaces it + internally by the remaining cases, making the patterns disjoint:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{thm}\isamarkupfalse% +\ sep{\isachardot}simps% +\begin{isamarkuptext}% +\begin{isabelle}% +sep\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}\isasep\isanewline% +sep\ a\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\isasep\isanewline% +sep\ a\ {\isacharbrackleft}v{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}v{\isacharbrackright}% +\end{isabelle}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\noindent The equations from function definitions are automatically used in + simplification:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isachardoublequoteopen}sep\ {\isadigit{0}}\ {\isacharbrackleft}{\isadigit{1}}{\isacharcomma}\ {\isadigit{2}}{\isacharcomma}\ {\isadigit{3}}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isadigit{1}}{\isacharcomma}\ {\isadigit{0}}{\isacharcomma}\ {\isadigit{2}}{\isacharcomma}\ {\isadigit{0}}{\isacharcomma}\ {\isadigit{3}}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Induction% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle provides customized induction rules for recursive + functions. These rules follow the recursive structure of the + definition. Here is the rule \isa{sep{\isachardot}induct} arising from the + above definition of \isa{sep}: + + \begin{isabelle}% +{\isasymlbrakk}{\isasymAnd}a\ x\ y\ xs{\isachardot}\ {\isacharquery}P\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharquery}P\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}{\isacharsemicolon}\ {\isasymAnd}a{\isachardot}\ {\isacharquery}P\ a\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\ {\isasymAnd}a\ v{\isachardot}\ {\isacharquery}P\ a\ {\isacharbrackleft}v{\isacharbrackright}{\isasymrbrakk}\isanewline +{\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}a{\isadigit{0}}{\isachardot}{\isadigit{0}}\ {\isacharquery}a{\isadigit{1}}{\isachardot}{\isadigit{0}}% +\end{isabelle} + + We have a step case for list with at least two elements, and two + base cases for the zero- and the one-element list. Here is a simple + proof about \isa{sep} and \isa{map}% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isachardoublequoteopen}map\ f\ {\isacharparenleft}sep\ x\ ys{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}map\ f\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}induct\ x\ ys\ rule{\isacharcolon}\ sep{\isachardot}induct{\isacharparenright}% +\begin{isamarkuptxt}% +We get three cases, like in the definition. + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a\ x\ y\ xs{\isachardot}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a{\isachardot}\ map\ f\ {\isacharparenleft}sep\ a\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline +\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}a\ v{\isachardot}\ map\ f\ {\isacharparenleft}sep\ a\ {\isacharbrackleft}v{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharbrackleft}v{\isacharbrackright}{\isacharparenright}% +\end{isabelle}% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ auto\ \isanewline +\isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +With the \cmd{fun} command, you can define about 80\% of the + functions that occur in practice. The rest of this tutorial explains + the remaining 20\%.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{fun vs.\ function% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The \cmd{fun} command provides a + convenient shorthand notation for simple function definitions. In + this mode, Isabelle tries to solve all the necessary proof obligations + automatically. If any proof fails, the definition is + rejected. This can either mean that the definition is indeed faulty, + or that the default proof procedures are just not smart enough (or + rather: not designed) to handle the definition. + + By expanding the abbreviation to the more verbose \cmd{function} command, these proof obligations become visible and can be analyzed or + solved manually. The expansion from \cmd{fun} to \cmd{function} is as follows: + +\end{isamarkuptext} + + +\[\left[\;\begin{minipage}{0.25\textwidth}\vspace{6pt} +\cmd{fun} \isa{f\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}}\\% +\cmd{where}\\% +\hspace*{2ex}{\it equations}\\% +\hspace*{2ex}\vdots\vspace*{6pt} +\end{minipage}\right] +\quad\equiv\quad +\left[\;\begin{minipage}{0.48\textwidth}\vspace{6pt} +\cmd{function} \isa{{\isacharparenleft}}\cmd{sequential}\isa{{\isacharparenright}\ f\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}}\\% +\cmd{where}\\% +\hspace*{2ex}{\it equations}\\% +\hspace*{2ex}\vdots\\% +\cmd{by} \isa{pat{\isacharunderscore}completeness\ auto}\\% +\cmd{termination by} \isa{lexicographic{\isacharunderscore}order}\vspace{6pt} +\end{minipage} +\right]\] + +\begin{isamarkuptext} + \vspace*{1em} + \noindent Some details have now become explicit: + + \begin{enumerate} + \item The \cmd{sequential} option enables the preprocessing of + pattern overlaps which we already saw. Without this option, the equations + must already be disjoint and complete. The automatic completion only + works with constructor patterns. + + \item A function definition produces a proof obligation which + expresses completeness and compatibility of patterns (we talk about + this later). The combination of the methods \isa{pat{\isacharunderscore}completeness} and + \isa{auto} is used to solve this proof obligation. + + \item A termination proof follows the definition, started by the + \cmd{termination} command. This will be explained in \S\ref{termination}. + \end{enumerate} + Whenever a \cmd{fun} command fails, it is usually a good idea to + expand the syntax to the more verbose \cmd{function} form, to see + what is actually going on.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Termination% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\label{termination} + The method \isa{lexicographic{\isacharunderscore}order} is the default method for + termination proofs. It can prove termination of a + certain class of functions by searching for a suitable lexicographic + combination of size measures. Of course, not all functions have such + a simple termination argument. For them, we can specify the termination + relation manually.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{The {\tt relation} method% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Consider the following function, which sums up natural numbers up to + \isa{N}, using a counter \isa{i}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ sum\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}sum\ i\ N\ {\isacharequal}\ {\isacharparenleft}if\ i\ {\isachargreater}\ N\ then\ {\isadigit{0}}\ else\ i\ {\isacharplus}\ sum\ {\isacharparenleft}Suc\ i{\isacharparenright}\ N{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent The \isa{lexicographic{\isacharunderscore}order} method fails on this example, because none of the + arguments decreases in the recursive call, with respect to the standard size ordering. + To prove termination manually, we must provide a custom wellfounded relation. + + The termination argument for \isa{sum} is based on the fact that + the \emph{difference} between \isa{i} and \isa{N} gets + smaller in every step, and that the recursion stops when \isa{i} + is greater than \isa{N}. Phrased differently, the expression + \isa{N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i} always decreases. + + We can use this expression as a measure function suitable to prove termination.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{termination}\isamarkupfalse% +\ sum\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% +\begin{isamarkuptxt}% +The \cmd{termination} command sets up the termination goal for the + specified function \isa{sum}. If the function name is omitted, it + implicitly refers to the last function definition. + + The \isa{relation} method takes a relation of + type \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set}, where \isa{{\isacharprime}a} is the argument type of + the function. If the function has multiple curried arguments, then + these are packed together into a tuple, as it happened in the above + example. + + The predefined function \isa{{\isachardoublequote}measure\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set{\isachardoublequote}} constructs a + wellfounded relation from a mapping into the natural numbers (a + \emph{measure function}). + + After the invocation of \isa{relation}, we must prove that (a) + the relation we supplied is wellfounded, and (b) that the arguments + of recursive calls indeed decrease with respect to the + relation: + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ wf\ {\isacharparenleft}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}\ N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharparenright}{\isacharparenright}\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}i\ N{\isachardot}\ {\isasymnot}\ N\ {\isacharless}\ i\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isacharparenleft}Suc\ i{\isacharcomma}\ N{\isacharparenright}{\isacharcomma}\ i{\isacharcomma}\ N{\isacharparenright}\ {\isasymin}\ measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}\ N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharparenright}% +\end{isabelle} + + These goals are all solved by \isa{auto}:% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ auto\isanewline +\isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Let us complicate the function a little, by adding some more + recursive calls:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ foo\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}foo\ i\ N\ {\isacharequal}\ {\isacharparenleft}if\ i\ {\isachargreater}\ N\ \isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ then\ {\isacharparenleft}if\ N\ {\isacharequal}\ {\isadigit{0}}\ then\ {\isadigit{0}}\ else\ foo\ {\isadigit{0}}\ {\isacharparenleft}N\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ else\ i\ {\isacharplus}\ foo\ {\isacharparenleft}Suc\ i{\isacharparenright}\ N{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +When \isa{i} has reached \isa{N}, it starts at zero again + and \isa{N} is decremented. + This corresponds to a nested + loop where one index counts up and the other down. Termination can + be proved using a lexicographic combination of two measures, namely + the value of \isa{N} and the above difference. The \isa{measures} combinator generalizes \isa{measure} by taking a + list of measure functions.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{termination}\isamarkupfalse% +\ \isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measures\ {\isacharbrackleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}\ N{\isacharparenright}{\isachardot}\ N{\isacharcomma}\ {\isasymlambda}{\isacharparenleft}i{\isacharcomma}N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharbrackright}{\isachardoublequoteclose}{\isacharparenright}\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{How \isa{lexicographic{\isacharunderscore}order} works% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +To see how the automatic termination proofs work, let's look at an + example where it fails\footnote{For a detailed discussion of the + termination prover, see \cite{bulwahnKN07}}: + +\end{isamarkuptext} +\cmd{fun} \isa{fails\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymRightarrow}\ nat\ list\ {\isasymRightarrow}\ nat{\isachardoublequote}}\\% +\cmd{where}\\% +\hspace*{2ex}\isa{{\isachardoublequote}fails\ a\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ a{\isachardoublequote}}\\% +|\hspace*{1.5ex}\isa{{\isachardoublequote}fails\ a\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ fails\ {\isacharparenleft}x\ {\isacharplus}\ a{\isacharparenright}\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}{\isachardoublequote}}\\ +\begin{isamarkuptext} + +\noindent Isabelle responds with the following error: + +\begin{isabelle} +*** Unfinished subgoals:\newline +*** (a, 1, <):\newline +*** \ 1.~\isa{{\isasymAnd}x{\isachardot}\ x\ {\isacharequal}\ {\isadigit{0}}}\newline +*** (a, 1, <=):\newline +*** \ 1.~False\newline +*** (a, 2, <):\newline +*** \ 1.~False\newline +*** Calls:\newline +*** a) \isa{{\isacharparenleft}a{\isacharcomma}\ x\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharminus}{\isacharminus}{\isachargreater}{\isachargreater}\ {\isacharparenleft}x\ {\isacharplus}\ a{\isacharcomma}\ x\ {\isacharhash}\ xs{\isacharparenright}}\newline +*** Measures:\newline +*** 1) \isa{{\isasymlambda}x{\isachardot}\ size\ {\isacharparenleft}fst\ x{\isacharparenright}}\newline +*** 2) \isa{{\isasymlambda}x{\isachardot}\ size\ {\isacharparenleft}snd\ x{\isacharparenright}}\newline +*** Result matrix:\newline +*** \ \ \ \ 1\ \ 2 \newline +*** a: ? <= \newline +*** Could not find lexicographic termination order.\newline +*** At command "fun".\newline +\end{isabelle}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +The key to this error message is the matrix at the bottom. The rows + of that matrix correspond to the different recursive calls (In our + case, there is just one). The columns are the function's arguments + (expressed through different measure functions, which map the + argument tuple to a natural number). + + The contents of the matrix summarize what is known about argument + descents: The second argument has a weak descent (\isa{{\isacharless}{\isacharequal}}) at the + recursive call, and for the first argument nothing could be proved, + which is expressed by \isa{{\isacharquery}}. In general, there are the values + \isa{{\isacharless}}, \isa{{\isacharless}{\isacharequal}} and \isa{{\isacharquery}}. + + For the failed proof attempts, the unfinished subgoals are also + printed. Looking at these will often point to a missing lemma. + +% As a more real example, here is quicksort:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Mutual Recursion% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +If two or more functions call one another mutually, they have to be defined + in one step. Here are \isa{even} and \isa{odd}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ even\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ odd\ \ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}even\ {\isadigit{0}}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}odd\ {\isadigit{0}}\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}even\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ odd\ n{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}odd\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ even\ n{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +To eliminate the mutual dependencies, Isabelle internally + creates a single function operating on the sum + type \isa{nat\ {\isacharplus}\ nat}. Then, \isa{even} and \isa{odd} are + defined as projections. Consequently, termination has to be proved + simultaneously for both functions, by specifying a measure on the + sum type:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{termination}\isamarkupfalse% +\ \isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ case\ x\ of\ Inl\ n\ {\isasymRightarrow}\ n\ {\isacharbar}\ Inr\ n\ {\isasymRightarrow}\ n{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +We could also have used \isa{lexicographic{\isacharunderscore}order}, which + supports mutual recursive termination proofs to a certain extent.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Induction for mutual recursion% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +When functions are mutually recursive, proving properties about them + generally requires simultaneous induction. The induction rule \isa{even{\isacharunderscore}odd{\isachardot}induct} + generated from the above definition reflects this. + + Let us prove something about \isa{even} and \isa{odd}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ even{\isacharunderscore}odd{\isacharunderscore}mod{\isadigit{2}}{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}odd\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +We apply simultaneous induction, specifying the induction variable + for both goals, separated by \cmd{and}:% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}induct\ n\ \isakeyword{and}\ n\ rule{\isacharcolon}\ even{\isacharunderscore}odd{\isachardot}induct{\isacharparenright}% +\begin{isamarkuptxt}% +We get four subgoals, which correspond to the clauses in the + definition of \isa{even} and \isa{odd}: + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ even\ {\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline +\ {\isadigit{2}}{\isachardot}\ odd\ {\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}\isanewline +\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ odd\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}\ {\isasymLongrightarrow}\ even\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline +\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ odd\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}% +\end{isabelle} + Simplification solves the first two goals, leaving us with two + statements about the \isa{mod} operation to prove:% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ simp{\isacharunderscore}all% +\begin{isamarkuptxt}% +\begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ odd\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isacharparenright}% +\end{isabelle} + + \noindent These can be handled by Isabelle's arithmetic decision procedures.% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ arith\isanewline +\isacommand{apply}\isamarkupfalse% +\ arith\isanewline +\isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +In proofs like this, the simultaneous induction is really essential: + Even if we are just interested in one of the results, the other + one is necessary to strengthen the induction hypothesis. If we leave + out the statement about \isa{odd} and just write \isa{True} instead, + the same proof fails:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ failed{\isacharunderscore}attempt{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ {\isachardoublequoteopen}True{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}induct\ n\ rule{\isacharcolon}\ even{\isacharunderscore}odd{\isachardot}induct{\isacharparenright}% +\begin{isamarkuptxt}% +\noindent Now the third subgoal is a dead end, since we have no + useful induction hypothesis available: + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ even\ {\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline +\ {\isadigit{2}}{\isachardot}\ True\isanewline +\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ True\ {\isasymLongrightarrow}\ even\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline +\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ True% +\end{isabelle}% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{oops}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsection{General pattern matching% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\label{genpats}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Avoiding automatic pattern splitting% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Up to now, we used pattern matching only on datatypes, and the + patterns were always disjoint and complete, and if they weren't, + they were made disjoint automatically like in the definition of + \isa{sep} in \S\ref{patmatch}. + + This automatic splitting can significantly increase the number of + equations involved, and this is not always desirable. The following + example shows the problem: + + Suppose we are modeling incomplete knowledge about the world by a + three-valued datatype, which has values \isa{T}, \isa{F} + and \isa{X} for true, false and uncertain propositions, respectively.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ P{\isadigit{3}}\ {\isacharequal}\ T\ {\isacharbar}\ F\ {\isacharbar}\ X% +\begin{isamarkuptext}% +\noindent Then the conjunction of such values can be defined as follows:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{fun}\isamarkupfalse% +\ And\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}And\ T\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And\ p\ T\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And\ p\ F\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And\ F\ p\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And\ X\ X\ {\isacharequal}\ X{\isachardoublequoteclose}% +\begin{isamarkuptext}% +This definition is useful, because the equations can directly be used + as simplification rules. But the patterns overlap: For example, + the expression \isa{And\ T\ T} is matched by both the first and + the second equation. By default, Isabelle makes the patterns disjoint by + splitting them up, producing instances:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{thm}\isamarkupfalse% +\ And{\isachardot}simps% +\begin{isamarkuptext}% +\isa{And\ T\ {\isacharquery}p\ {\isacharequal}\ {\isacharquery}p\isasep\isanewline% +And\ F\ T\ {\isacharequal}\ F\isasep\isanewline% +And\ X\ T\ {\isacharequal}\ X\isasep\isanewline% +And\ F\ F\ {\isacharequal}\ F\isasep\isanewline% +And\ X\ F\ {\isacharequal}\ F\isasep\isanewline% +And\ F\ X\ {\isacharequal}\ F\isasep\isanewline% +And\ X\ X\ {\isacharequal}\ X} + + \vspace*{1em} + \noindent There are several problems with this: + + \begin{enumerate} + \item If the datatype has many constructors, there can be an + explosion of equations. For \isa{And}, we get seven instead of + five equations, which can be tolerated, but this is just a small + example. + + \item Since splitting makes the equations \qt{less general}, they + do not always match in rewriting. While the term \isa{And\ x\ F} + can be simplified to \isa{F} with the original equations, a + (manual) case split on \isa{x} is now necessary. + + \item The splitting also concerns the induction rule \isa{And{\isachardot}induct}. Instead of five premises it now has seven, which + means that our induction proofs will have more cases. + + \item In general, it increases clarity if we get the same definition + back which we put in. + \end{enumerate} + + If we do not want the automatic splitting, we can switch it off by + leaving out the \cmd{sequential} option. However, we will have to + prove that our pattern matching is consistent\footnote{This prevents + us from defining something like \isa{f\ x\ {\isacharequal}\ True} and \isa{f\ x\ {\isacharequal}\ False} simultaneously.}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ And{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}And{\isadigit{2}}\ T\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ p\ T\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ p\ F\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ F\ p\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ X\ X\ {\isacharequal}\ X{\isachardoublequoteclose}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +\noindent Now let's look at the proof obligations generated by a + function definition. In this case, they are: + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ {\isasymlbrakk}{\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ F{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ \ }{\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ p{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ x\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ }{\isasymLongrightarrow}\ P\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}T{\isacharcomma}\ pa{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ pa\isanewline +\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ pa\isanewline +\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ F{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline +\ {\isadigit{5}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ pa{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline +\ {\isadigit{6}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ X\isanewline +\ {\isadigit{7}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ pa\isanewline +\ {\isadigit{8}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ F{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline +\ {\isadigit{9}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ pa{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline +\ {\isadigit{1}}{\isadigit{0}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ X% +\end{isabelle}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em} + + The first subgoal expresses the completeness of the patterns. It has + the form of an elimination rule and states that every \isa{x} of + the function's input type must match at least one of the patterns\footnote{Completeness could + be equivalently stated as a disjunction of existential statements: +\isa{{\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ F{\isacharparenright}{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ p{\isacharparenright}{\isacharparenright}\ {\isasymor}\ x\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}}, and you can use the method \isa{atomize{\isacharunderscore}elim} to get that form instead.}. If the patterns just involve + datatypes, we can solve it with the \isa{pat{\isacharunderscore}completeness} + method:% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ pat{\isacharunderscore}completeness% +\begin{isamarkuptxt}% +The remaining subgoals express \emph{pattern compatibility}. We do + allow that an input value matches multiple patterns, but in this + case, the result (i.e.~the right hand sides of the equations) must + also be equal. For each pair of two patterns, there is one such + subgoal. Usually this needs injectivity of the constructors, which + is used automatically by \isa{auto}.% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{by}\isamarkupfalse% +\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Non-constructor patterns% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Most of Isabelle's basic types take the form of inductive datatypes, + and usually pattern matching works on the constructors of such types. + However, this need not be always the case, and the \cmd{function} + command handles other kind of patterns, too. + + One well-known instance of non-constructor patterns are + so-called \emph{$n+k$-patterns}, which are a little controversial in + the functional programming world. Here is the initial fibonacci + example with $n+k$-patterns:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ fib{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}fib{\isadigit{2}}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}fib{\isadigit{2}}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}fib{\isadigit{2}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ fib{\isadigit{2}}\ n\ {\isacharplus}\ fib{\isadigit{2}}\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +This kind of matching is again justified by the proof of pattern + completeness and compatibility. + The proof obligation for pattern completeness states that every natural number is + either \isa{{\isadigit{0}}}, \isa{{\isadigit{1}}} or \isa{n\ {\isacharplus}\ {\isadigit{2}}}: + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ {\isasymlbrakk}x\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ x\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}n{\isachardot}\ x\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline +\ {\isadigit{3}}{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline +\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\isanewline +\ {\isadigit{5}}{\isachardot}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline +\ {\isadigit{6}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{1}}\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\isanewline +\ {\isadigit{7}}{\isachardot}\ {\isasymAnd}n\ na{\isachardot}\isanewline +\isaindent{\ {\isadigit{7}}{\isachardot}\ \ \ \ }n\ {\isacharplus}\ {\isadigit{2}}\ {\isacharequal}\ na\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\isanewline +\isaindent{\ {\isadigit{7}}{\isachardot}\ \ \ \ }fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ na\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ na{\isacharparenright}% +\end{isabelle} + + This is an arithmetic triviality, but unfortunately the + \isa{arith} method cannot handle this specific form of an + elimination rule. However, we can use the method \isa{atomize{\isacharunderscore}elim} to do an ad-hoc conversion to a disjunction of + existentials, which can then be solved by the arithmetic decision procedure. + Pattern compatibility and termination are automatic as usual.% +\end{isamarkuptxt}% +\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ atomize{\isacharunderscore}elim\isanewline +\isacommand{apply}\isamarkupfalse% +\ arith\isanewline +\isacommand{apply}\isamarkupfalse% +\ auto\isanewline +\isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{termination}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ lexicographic{\isacharunderscore}order% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +We can stretch the notion of pattern matching even more. The + following function is not a sensible functional program, but a + perfectly valid mathematical definition:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ ev\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}ev\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}ev\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ atomize{\isacharunderscore}elim\isanewline +\isacommand{by}\isamarkupfalse% +\ arith{\isacharplus}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{termination}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}relation\ {\isachardoublequoteopen}{\isacharbraceleft}{\isacharbraceright}{\isachardoublequoteclose}{\isacharparenright}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +This general notion of pattern matching gives you a certain freedom + in writing down specifications. However, as always, such freedom should + be used with care: + + If we leave the area of constructor + patterns, we have effectively departed from the world of functional + programming. This means that it is no longer possible to use the + code generator, and expect it to generate ML code for our + definitions. Also, such a specification might not work very well together with + simplification. Your mileage may vary.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Conditional equations% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The function package also supports conditional equations, which are + similar to guards in a language like Haskell. Here is Euclid's + algorithm written with conditional patterns\footnote{Note that the + patterns are also overlapping in the base case}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ gcd\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}gcd\ x\ {\isadigit{0}}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}gcd\ {\isadigit{0}}\ y\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}x\ {\isacharless}\ y\ {\isasymLongrightarrow}\ gcd\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}\ {\isacharequal}\ gcd\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}y\ {\isacharminus}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}{\isasymnot}\ x\ {\isacharless}\ y\ {\isasymLongrightarrow}\ gcd\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}\ {\isacharequal}\ gcd\ {\isacharparenleft}x\ {\isacharminus}\ y{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}atomize{\isacharunderscore}elim{\isacharcomma}\ auto{\isacharcomma}\ arith{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{termination}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ lexicographic{\isacharunderscore}order% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +By now, you can probably guess what the proof obligations for the + pattern completeness and compatibility look like. + + Again, functions with conditional patterns are not supported by the + code generator.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Pattern matching on strings% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +As strings (as lists of characters) are normal datatypes, pattern + matching on them is possible, but somewhat problematic. Consider the + following definition: + +\end{isamarkuptext} +\noindent\cmd{fun} \isa{check\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}string\ {\isasymRightarrow}\ bool{\isachardoublequote}}\\% +\cmd{where}\\% +\hspace*{2ex}\isa{{\isachardoublequote}check\ {\isacharparenleft}{\isacharprime}{\isacharprime}good{\isacharprime}{\isacharprime}{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequote}}\\% +\isa{{\isacharbar}\ {\isachardoublequote}check\ s\ {\isacharequal}\ False{\isachardoublequote}} +\begin{isamarkuptext} + + \noindent An invocation of the above \cmd{fun} command does not + terminate. What is the problem? Strings are lists of characters, and + characters are a datatype with a lot of constructors. Splitting the + catch-all pattern thus leads to an explosion of cases, which cannot + be handled by Isabelle. + + There are two things we can do here. Either we write an explicit + \isa{if} on the right hand side, or we can use conditional patterns:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ check\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}string\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}check\ {\isacharparenleft}{\isacharprime}{\isacharprime}good{\isacharprime}{\isacharprime}{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}s\ {\isasymnoteq}\ {\isacharprime}{\isacharprime}good{\isacharprime}{\isacharprime}\ {\isasymLongrightarrow}\ check\ s\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsection{Partiality% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In HOL, all functions are total. A function \isa{f} applied to + \isa{x} always has the value \isa{f\ x}, and there is no notion + of undefinedness. + This is why we have to do termination + proofs when defining functions: The proof justifies that the + function can be defined by wellfounded recursion. + + However, the \cmd{function} package does support partiality to a + certain extent. Let's look at the following function which looks + for a zero of a given function f.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ findzero\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}findzero\ f\ n\ {\isacharequal}\ {\isacharparenleft}if\ f\ n\ {\isacharequal}\ {\isadigit{0}}\ then\ n\ else\ findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Clearly, any attempt of a termination proof must fail. And without + that, we do not get the usual rules \isa{findzero{\isachardot}simps} and + \isa{findzero{\isachardot}induct}. So what was the definition good for at all?% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Domain predicates% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The trick is that Isabelle has not only defined the function \isa{findzero}, but also + a predicate \isa{findzero{\isacharunderscore}dom} that characterizes the values where the function + terminates: the \emph{domain} of the function. If we treat a + partial function just as a total function with an additional domain + predicate, we can derive simplification and + induction rules as we do for total functions. They are guarded + by domain conditions and are called \isa{psimps} and \isa{pinduct}:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\noindent\begin{minipage}{0.79\textwidth}\begin{isabelle}% +findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ {\isacharquery}n{\isacharparenright}\ {\isasymLongrightarrow}\isanewline +findzero\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isacharparenleft}if\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isadigit{0}}\ then\ {\isacharquery}n\ else\ findzero\ {\isacharquery}f\ {\isacharparenleft}Suc\ {\isacharquery}n{\isacharparenright}{\isacharparenright}% +\end{isabelle}\end{minipage} + \hfill(\isa{findzero{\isachardot}psimps}) + \vspace{1em} + + \noindent\begin{minipage}{0.79\textwidth}\begin{isabelle}% +{\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}a{\isadigit{0}}{\isachardot}{\isadigit{0}}{\isacharcomma}\ {\isacharquery}a{\isadigit{1}}{\isachardot}{\isadigit{0}}{\isacharparenright}{\isacharsemicolon}\isanewline +\isaindent{\ }{\isasymAnd}f\ n{\isachardot}\ {\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isacharsemicolon}\ f\ n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isacharquery}P\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ f\ n{\isasymrbrakk}\isanewline +{\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}a{\isadigit{0}}{\isachardot}{\isadigit{0}}\ {\isacharquery}a{\isadigit{1}}{\isachardot}{\isadigit{0}}% +\end{isabelle}\end{minipage} + \hfill(\isa{findzero{\isachardot}pinduct})% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +Remember that all we + are doing here is use some tricks to make a total function appear + as if it was partial. We can still write the term \isa{findzero\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ {\isadigit{1}}{\isacharparenright}\ {\isadigit{0}}} and like any other term of type \isa{nat} it is equal + to some natural number, although we might not be able to find out + which one. The function is \emph{underdefined}. + + But it is defined enough to prove something interesting about it. We + can prove that if \isa{findzero\ f\ n} + terminates, it indeed returns a zero of \isa{f}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ findzero{\isacharunderscore}zero{\isacharcolon}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ n{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +\noindent We apply induction as usual, but using the partial induction + rule:% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}induct\ f\ n\ rule{\isacharcolon}\ findzero{\isachardot}pinduct{\isacharparenright}% +\begin{isamarkuptxt}% +\noindent This gives the following subgoals: + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}f\ n{\isachardot}\ {\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isacharsemicolon}\ f\ n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}{\isasymrbrakk}\isanewline +\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}f\ n{\isachardot}\ }{\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ n{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}% +\end{isabelle} + + \noindent The hypothesis in our lemma was used to satisfy the first premise in + the induction rule. However, we also get \isa{findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}} as a local assumption in the induction step. This + allows to unfold \isa{findzero\ f\ n} using the \isa{psimps} + rule, and the rest is trivial. Since the \isa{psimps} rules carry the + \isa{{\isacharbrackleft}simp{\isacharbrackright}} attribute by default, we just need a single step:% +\end{isamarkuptxt}% +\isamarkuptrue% +\isacommand{apply}\isamarkupfalse% +\ simp\isanewline +\isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Proofs about partial functions are often not harder than for total + functions. Fig.~\ref{findzero_isar} shows a slightly more + complicated proof written in Isar. It is verbose enough to show how + partiality comes into play: From the partial induction, we get an + additional domain condition hypothesis. Observe how this condition + is applied when calls to \isa{findzero} are unfolded.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{figure} +\hrule\vspace{6pt} +\begin{minipage}{0.8\textwidth} +\isabellestyle{it} +\isastyle\isamarkuptrue +\isacommand{lemma}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isacharsemicolon}\ x\ {\isasymin}\ {\isacharbraceleft}n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ f\ x\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}induct\ rule{\isacharcolon}\ findzero{\isachardot}pinduct{\isacharparenright}\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ f\ n\ \isacommand{assume}\isamarkupfalse% +\ dom{\isacharcolon}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ IH{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharsemicolon}\ x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharbraceright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ f\ x\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ x{\isacharunderscore}range{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ \isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}f\ n\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ dom\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}findzero\ f\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ x{\isacharunderscore}range\ \isacommand{show}\isamarkupfalse% +\ False\ \isacommand{by}\isamarkupfalse% +\ auto\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ x{\isacharunderscore}range\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isacharequal}\ n\ {\isasymor}\ x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ auto\isanewline +\ \ \isacommand{thus}\isamarkupfalse% +\ {\isachardoublequoteopen}f\ x\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ {\isacharbackquoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ dom\ \isakeyword{and}\ {\isacharbackquoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharbackquoteclose}\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ IH\ \isakeyword{and}\ {\isacharbackquoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharbackquoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupfalse\isabellestyle{tt} +\end{minipage}\vspace{6pt}\hrule +\caption{A proof about a partial function}\label{findzero_isar} +\end{figure} +% +\isamarkupsubsection{Partial termination proofs% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Now that we have proved some interesting properties about our + function, we should turn to the domain predicate and see if it is + actually true for some values. Otherwise we would have just proved + lemmas with \isa{False} as a premise. + + Essentially, we need some introduction rules for \isa{findzero{\isacharunderscore}dom}. The function package can prove such domain + introduction rules automatically. But since they are not used very + often (they are almost never needed if the function is total), this + functionality is disabled by default for efficiency reasons. So we have to go + back and ask for them explicitly by passing the \isa{{\isacharparenleft}domintros{\isacharparenright}} option to the function package: + +\vspace{1ex} +\noindent\cmd{function} \isa{{\isacharparenleft}domintros{\isacharparenright}\ findzero\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}nat\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequote}}\\% +\cmd{where}\isanewline% +\ \ \ldots\\ + + \noindent Now the package has proved an introduction rule for \isa{findzero{\isacharunderscore}dom}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{thm}\isamarkupfalse% +\ findzero{\isachardot}domintros% +\begin{isamarkuptext}% +\begin{isabelle}% +{\isacharparenleft}{\isadigit{0}}\ {\isacharless}\ {\isacharquery}f\ {\isacharquery}n\ {\isasymLongrightarrow}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ Suc\ {\isacharquery}n{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ {\isacharquery}n{\isacharparenright}% +\end{isabelle} + + Domain introduction rules allow to show that a given value lies in the + domain of a function, if the arguments of all recursive calls + are in the domain as well. They allow to do a \qt{single step} in a + termination proof. Usually, you want to combine them with a suitable + induction principle. + + Since our function increases its argument at recursive calls, we + need an induction principle which works \qt{backwards}. We will use + \isa{inc{\isacharunderscore}induct}, which allows to do induction from a fixed number + \qt{downwards}: + + \begin{center}\isa{{\isasymlbrakk}{\isacharquery}i\ {\isasymle}\ {\isacharquery}j{\isacharsemicolon}\ {\isacharquery}P\ {\isacharquery}j{\isacharsemicolon}\ {\isasymAnd}i{\isachardot}\ {\isasymlbrakk}i\ {\isacharless}\ {\isacharquery}j{\isacharsemicolon}\ {\isacharquery}P\ {\isacharparenleft}Suc\ i{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ i{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}i}\hfill(\isa{inc{\isacharunderscore}induct})\end{center} + + Figure \ref{findzero_term} gives a detailed Isar proof of the fact + that \isa{findzero} terminates if there is a zero which is greater + or equal to \isa{n}. First we derive two useful rules which will + solve the base case and the step case of the induction. The + induction is then straightforward, except for the unusual induction + principle.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{figure} +\hrule\vspace{6pt} +\begin{minipage}{0.8\textwidth} +\isabellestyle{it} +\isastyle\isamarkuptrue +\isacommand{lemma}\isamarkupfalse% +\ findzero{\isacharunderscore}termination{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}x\ {\isasymge}\ n{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\ \isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ base{\isacharcolon}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ findzero{\isachardot}domintros{\isacharparenright}\ {\isacharparenleft}simp\ add{\isacharcolon}{\isacharbackquoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}{\isacharbackquoteclose}{\isacharparenright}\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ step{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}i{\isachardot}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ Suc\ i{\isacharparenright}\ \isanewline +\ \ \ \ {\isasymLongrightarrow}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ i{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ findzero{\isachardot}domintros{\isacharparenright}\ simp\isanewline +\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ {\isacharbackquoteopen}x\ {\isasymge}\ n{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}induct\ rule{\isacharcolon}inc{\isacharunderscore}induct{\isacharparenright}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ x{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ base{\isacharparenright}\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ i\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ Suc\ i{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{thus}\isamarkupfalse% +\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ i{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ step{\isacharparenright}\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupfalse\isabellestyle{tt} +\end{minipage}\vspace{6pt}\hrule +\caption{Termination proof for \isa{findzero}}\label{findzero_term} +\end{figure} +% +\begin{isamarkuptext}% +Again, the proof given in Fig.~\ref{findzero_term} has a lot of + detail in order to explain the principles. Using more automation, we + can also have a short proof:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ findzero{\isacharunderscore}termination{\isacharunderscore}short{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ zero{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isachargreater}{\isacharequal}\ n{\isachardoublequoteclose}\ \isanewline +\ \ \isakeyword{assumes}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{using}\isamarkupfalse% +\ zero\isanewline +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ rule{\isacharcolon}inc{\isacharunderscore}induct{\isacharparenright}\ {\isacharparenleft}auto\ intro{\isacharcolon}\ findzero{\isachardot}domintros{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent It is simple to combine the partial correctness result with the + termination lemma:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ findzero{\isacharunderscore}total{\isacharunderscore}correctness{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}blast\ intro{\isacharcolon}\ findzero{\isacharunderscore}zero\ findzero{\isacharunderscore}termination{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Definition of the domain predicate% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Sometimes it is useful to know what the definition of the domain + predicate looks like. Actually, \isa{findzero{\isacharunderscore}dom} is just an + abbreviation: + + \begin{isabelle}% +findzero{\isacharunderscore}dom\ {\isasymequiv}\ accp\ findzero{\isacharunderscore}rel% +\end{isabelle} + + The domain predicate is the \emph{accessible part} of a relation \isa{findzero{\isacharunderscore}rel}, which was also created internally by the function + package. \isa{findzero{\isacharunderscore}rel} is just a normal + inductive predicate, so we can inspect its definition by + looking at the introduction rules \isa{findzero{\isacharunderscore}rel{\isachardot}intros}. + In our case there is just a single rule: + + \begin{isabelle}% +{\isacharquery}f\ {\isacharquery}n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ findzero{\isacharunderscore}rel\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ Suc\ {\isacharquery}n{\isacharparenright}\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ {\isacharquery}n{\isacharparenright}% +\end{isabelle} + + The predicate \isa{findzero{\isacharunderscore}rel} + describes the \emph{recursion relation} of the function + definition. The recursion relation is a binary relation on + the arguments of the function that relates each argument to its + recursive calls. In general, there is one introduction rule for each + recursive call. + + The predicate \isa{findzero{\isacharunderscore}dom} is the accessible part of + that relation. An argument belongs to the accessible part, if it can + be reached in a finite number of steps (cf.~its definition in \isa{Wellfounded{\isachardot}thy}). + + Since the domain predicate is just an abbreviation, you can use + lemmas for \isa{accp} and \isa{findzero{\isacharunderscore}rel} directly. Some + lemmas which are occasionally useful are \isa{accpI}, \isa{accp{\isacharunderscore}downward}, and of course the introduction and elimination rules + for the recursion relation \isa{findzero{\isachardot}intros} and \isa{findzero{\isachardot}cases}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{A Useful Special Case: Tail recursion% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The domain predicate is our trick that allows us to model partiality + in a world of total functions. The downside of this is that we have + to carry it around all the time. The termination proof above allowed + us to replace the abstract \isa{findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}} by the more + concrete \isa{n\ {\isasymle}\ x\ {\isasymand}\ f\ x\ {\isacharequal}\ {\isadigit{0}}}, but the condition is still + there and can only be discharged for special cases. + In particular, the domain predicate guards the unfolding of our + function, since it is there as a condition in the \isa{psimp} + rules. + + Now there is an important special case: We can actually get rid + of the condition in the simplification rules, \emph{if the function + is tail-recursive}. The reason is that for all tail-recursive + equations there is a total function satisfying them, even if they + are non-terminating. + +% A function is tail recursive, if each call to the function is either +% equal +% +% So the outer form of the +% +%if it can be written in the following +% form: +% {term[display] "f x = (if COND x then BASE x else f (LOOP x))"} + + + The function package internally does the right construction and can + derive the unconditional simp rules, if we ask it to do so. Luckily, + our \isa{findzero} function is tail-recursive, so we can just go + back and add another option to the \cmd{function} command: + +\vspace{1ex} +\noindent\cmd{function} \isa{{\isacharparenleft}domintros{\isacharcomma}\ tailrec{\isacharparenright}\ findzero\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}nat\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequote}}\\% +\cmd{where}\isanewline% +\ \ \ldots\\% + + + \noindent Now, we actually get unconditional simplification rules, even + though the function is partial:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{thm}\isamarkupfalse% +\ findzero{\isachardot}simps% +\begin{isamarkuptext}% +\begin{isabelle}% +findzero\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isacharparenleft}if\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isadigit{0}}\ then\ {\isacharquery}n\ else\ findzero\ {\isacharquery}f\ {\isacharparenleft}Suc\ {\isacharquery}n{\isacharparenright}{\isacharparenright}% +\end{isabelle} + + \noindent Of course these would make the simplifier loop, so we better remove + them from the simpset:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{declare}\isamarkupfalse% +\ findzero{\isachardot}simps{\isacharbrackleft}simp\ del{\isacharbrackright}% +\begin{isamarkuptext}% +Getting rid of the domain conditions in the simplification rules is + not only useful because it simplifies proofs. It is also required in + order to use Isabelle's code generator to generate ML code + from a function definition. + Since the code generator only works with equations, it cannot be + used with \isa{psimp} rules. Thus, in order to generate code for + partial functions, they must be defined as a tail recursion. + Luckily, many functions have a relatively natural tail recursive + definition.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Nested recursion% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Recursive calls which are nested in one another frequently cause + complications, since their termination proof can depend on a partial + correctness property of the function itself. + + As a small example, we define the \qt{nested zero} function:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ nz\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}nz\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}nz\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ nz\ {\isacharparenleft}nz\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +If we attempt to prove termination using the identity measure on + naturals, this fails:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{termination}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}n{\isachardot}\ n{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\ \ \isacommand{apply}\isamarkupfalse% +\ auto% +\begin{isamarkuptxt}% +We get stuck with the subgoal + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ nz{\isacharunderscore}dom\ n\ {\isasymLongrightarrow}\ nz\ n\ {\isacharless}\ Suc\ n% +\end{isabelle} + + Of course this statement is true, since we know that \isa{nz} is + the zero function. And in fact we have no problem proving this + property by induction.% +\end{isamarkuptxt}% +\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isacommand{lemma}\isamarkupfalse% +\ nz{\isacharunderscore}is{\isacharunderscore}zero{\isacharcolon}\ {\isachardoublequoteopen}nz{\isacharunderscore}dom\ n\ {\isasymLongrightarrow}\ nz\ n\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ rule{\isacharcolon}nz{\isachardot}pinduct{\isacharparenright}\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +We formulate this as a partial correctness lemma with the condition + \isa{nz{\isacharunderscore}dom\ n}. This allows us to prove it with the \isa{pinduct} rule before we have proved termination. With this lemma, + the termination proof works as expected:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{termination}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}n{\isachardot}\ n{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\ {\isacharparenleft}auto\ simp{\isacharcolon}\ nz{\isacharunderscore}is{\isacharunderscore}zero{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +As a general strategy, one should prove the statements needed for + termination as a partial property first. Then they can be used to do + the termination proof. This also works for less trivial + examples. Figure \ref{f91} defines the 91-function, a well-known + challenge problem due to John McCarthy, and proves its termination.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{figure} +\hrule\vspace{6pt} +\begin{minipage}{0.8\textwidth} +\isabellestyle{it} +\isastyle\isamarkuptrue +\isacommand{function}\isamarkupfalse% +\ f{\isadigit{9}}{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}f{\isadigit{9}}{\isadigit{1}}\ n\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharless}\ n\ then\ n\ {\isacharminus}\ {\isadigit{1}}{\isadigit{0}}\ else\ f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}estimate{\isacharcolon}\ \isanewline +\ \ \isakeyword{assumes}\ trm{\isacharcolon}\ {\isachardoublequoteopen}f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}dom\ n{\isachardoublequoteclose}\ \isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}n\ {\isacharless}\ f{\isadigit{9}}{\isadigit{1}}\ n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{using}\isamarkupfalse% +\ trm\ \isacommand{by}\isamarkupfalse% +\ induct\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{termination}\isamarkupfalse% +\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{let}\isamarkupfalse% +\ {\isacharquery}R\ {\isacharequal}\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{1}}\ {\isacharminus}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}wf\ {\isacharquery}R{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ n\ {\isacharcolon}{\isacharcolon}\ nat\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymnot}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharless}\ n{\isachardoublequoteclose}\ % +\isamarkupcmt{Assumptions for both calls% +} +\isanewline +\isanewline +\ \ \isacommand{thus}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharcomma}\ n{\isacharparenright}\ {\isasymin}\ {\isacharquery}R{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\ % +\isamarkupcmt{Inner call% +} +\isanewline +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ inner{\isacharunderscore}trm{\isacharcolon}\ {\isachardoublequoteopen}f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}dom\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\ % +\isamarkupcmt{Outer call% +} +\isanewline +\ \ \isacommand{with}\isamarkupfalse% +\ f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}estimate\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}\ {\isacharless}\ f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{with}\isamarkupfalse% +\ {\isacharbackquoteopen}{\isasymnot}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharless}\ n{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\isacharcomma}\ n{\isacharparenright}\ {\isasymin}\ {\isacharquery}R{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupfalse\isabellestyle{tt} +\end{minipage} +\vspace{6pt}\hrule +\caption{McCarthy's 91-function}\label{f91} +\end{figure} +% +\isamarkupsection{Higher-Order Recursion% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Higher-order recursion occurs when recursive calls + are passed as arguments to higher-order combinators such as \isa{map}, \isa{filter} etc. + As an example, imagine a datatype of n-ary trees:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ {\isacharprime}a\ tree\ {\isacharequal}\ \isanewline +\ \ Leaf\ {\isacharprime}a\ \isanewline +{\isacharbar}\ Branch\ {\isachardoublequoteopen}{\isacharprime}a\ tree\ list{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent We can define a function which swaps the left and right subtrees recursively, using the + list functions \isa{rev} and \isa{map}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{fun}\isamarkupfalse% +\ mirror\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ tree{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}mirror\ {\isacharparenleft}Leaf\ n{\isacharparenright}\ {\isacharequal}\ Leaf\ n{\isachardoublequoteclose}\isanewline +{\isacharbar}\ {\isachardoublequoteopen}mirror\ {\isacharparenleft}Branch\ l{\isacharparenright}\ {\isacharequal}\ Branch\ {\isacharparenleft}rev\ {\isacharparenleft}map\ mirror\ l{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +Although the definition is accepted without problems, let us look at the termination proof:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{termination}\isamarkupfalse% +% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +% +\begin{isamarkuptxt}% +As usual, we have to give a wellfounded relation, such that the + arguments of the recursive calls get smaller. But what exactly are + the arguments of the recursive calls when mirror is given as an + argument to \isa{map}? Isabelle gives us the + subgoals + + \begin{isabelle}% +\ {\isadigit{1}}{\isachardot}\ wf\ {\isacharquery}R\isanewline +\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}l\ x{\isachardot}\ x\ {\isasymin}\ set\ l\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ Branch\ l{\isacharparenright}\ {\isasymin}\ {\isacharquery}R% +\end{isabelle} + + So the system seems to know that \isa{map} only + applies the recursive call \isa{mirror} to elements + of \isa{l}, which is essential for the termination proof. + + This knowledge about \isa{map} is encoded in so-called congruence rules, + which are special theorems known to the \cmd{function} command. The + rule for \isa{map} is + + \begin{isabelle}% +{\isasymlbrakk}{\isacharquery}xs\ {\isacharequal}\ {\isacharquery}ys{\isacharsemicolon}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ set\ {\isacharquery}ys\ {\isasymLongrightarrow}\ {\isacharquery}f\ x\ {\isacharequal}\ {\isacharquery}g\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ map\ {\isacharquery}f\ {\isacharquery}xs\ {\isacharequal}\ map\ {\isacharquery}g\ {\isacharquery}ys% +\end{isabelle} + + You can read this in the following way: Two applications of \isa{map} are equal, if the list arguments are equal and the functions + coincide on the elements of the list. This means that for the value + \isa{map\ f\ l} we only have to know how \isa{f} behaves on + the elements of \isa{l}. + + Usually, one such congruence rule is + needed for each higher-order construct that is used when defining + new functions. In fact, even basic functions like \isa{If} and \isa{Let} are handled by this mechanism. The congruence + rule for \isa{If} states that the \isa{then} branch is only + relevant if the condition is true, and the \isa{else} branch only if it + is false: + + \begin{isabelle}% +{\isasymlbrakk}{\isacharquery}b\ {\isacharequal}\ {\isacharquery}c{\isacharsemicolon}\ {\isacharquery}c\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isacharequal}\ {\isacharquery}u{\isacharsemicolon}\ {\isasymnot}\ {\isacharquery}c\ {\isasymLongrightarrow}\ {\isacharquery}y\ {\isacharequal}\ {\isacharquery}v{\isasymrbrakk}\isanewline +{\isasymLongrightarrow}\ {\isacharparenleft}if\ {\isacharquery}b\ then\ {\isacharquery}x\ else\ {\isacharquery}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ {\isacharquery}c\ then\ {\isacharquery}u\ else\ {\isacharquery}v{\isacharparenright}% +\end{isabelle} + + Congruence rules can be added to the + function package by giving them the \isa{fundef{\isacharunderscore}cong} attribute. + + The constructs that are predefined in Isabelle, usually + come with the respective congruence rules. + But if you define your own higher-order functions, you may have to + state and prove the required congruence rules yourself, if you want to use your + functions in recursive definitions.% +\end{isamarkuptxt}% +\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Congruence Rules and Evaluation Order% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Higher order logic differs from functional programming languages in + that it has no built-in notion of evaluation order. A program is + just a set of equations, and it is not specified how they must be + evaluated. + + However for the purpose of function definition, we must talk about + evaluation order implicitly, when we reason about termination. + Congruence rules express that a certain evaluation order is + consistent with the logical definition. + + Consider the following function.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{function}\isamarkupfalse% +\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}f\ n\ {\isacharequal}\ {\isacharparenleft}n\ {\isacharequal}\ {\isadigit{0}}\ {\isasymor}\ f\ {\isacharparenleft}n\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +For this definition, the termination proof fails. The default configuration + specifies no congruence rule for disjunction. We have to add a + congruence rule that specifies left-to-right evaluation order: + + \vspace{1ex} + \noindent \isa{{\isasymlbrakk}{\isacharquery}P\ {\isacharequal}\ {\isacharquery}P{\isacharprime}{\isacharsemicolon}\ {\isasymnot}\ {\isacharquery}P{\isacharprime}\ {\isasymLongrightarrow}\ {\isacharquery}Q\ {\isacharequal}\ {\isacharquery}Q{\isacharprime}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isacharquery}P\ {\isasymor}\ {\isacharquery}Q{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}P{\isacharprime}\ {\isasymor}\ {\isacharquery}Q{\isacharprime}{\isacharparenright}}\hfill(\isa{disj{\isacharunderscore}cong}) + \vspace{1ex} + + Now the definition works without problems. Note how the termination + proof depends on the extra condition that we get from the congruence + rule. + + However, as evaluation is not a hard-wired concept, we + could just turn everything around by declaring a different + congruence rule. Then we can make the reverse definition:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ disj{\isacharunderscore}cong{\isadigit{2}}{\isacharbrackleft}fundef{\isacharunderscore}cong{\isacharbrackright}{\isacharcolon}\ \isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}{\isasymnot}\ Q{\isacharprime}\ {\isasymLongrightarrow}\ P\ {\isacharequal}\ P{\isacharprime}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}Q\ {\isacharequal}\ Q{\isacharprime}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}P{\isacharprime}\ {\isasymor}\ Q{\isacharprime}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{fun}\isamarkupfalse% +\ f{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}f{\isacharprime}\ n\ {\isacharequal}\ {\isacharparenleft}f{\isacharprime}\ {\isacharparenleft}n\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}\ {\isasymor}\ n\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent These examples show that, in general, there is no \qt{best} set of + congruence rules. + + However, such tweaking should rarely be necessary in + practice, as most of the time, the default set of congruence rules + works well.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/Thy/document/session.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/Thy/document/session.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,6 @@ +\input{Functions.tex} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/conclusion.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/conclusion.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,7 @@ +\section{Conclusion} + +\fixme{} + + + + diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/functions.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/functions.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,91 @@ + +\documentclass[a4paper,fleqn]{article} + +\usepackage{latexsym,graphicx} +\usepackage[refpage]{nomencl} +\usepackage{../iman,../extra,../isar,../proof} +\usepackage{../isabelle,../isabellesym} +\usepackage{style} +\usepackage{mathpartir} +\usepackage{amsthm} +\usepackage{../pdfsetup} + +\newcommand{\cmd}[1]{\isacommand{#1}} + +\newcommand{\isasymINFIX}{\cmd{infix}} +\newcommand{\isasymLOCALE}{\cmd{locale}} +\newcommand{\isasymINCLUDES}{\cmd{includes}} +\newcommand{\isasymDATATYPE}{\cmd{datatype}} +\newcommand{\isasymAXCLASS}{\cmd{axclass}} +\newcommand{\isasymDEFINES}{\cmd{defines}} +\newcommand{\isasymNOTES}{\cmd{notes}} +\newcommand{\isasymCLASS}{\cmd{class}} +\newcommand{\isasymINSTANCE}{\cmd{instance}} +\newcommand{\isasymLEMMA}{\cmd{lemma}} +\newcommand{\isasymPROOF}{\cmd{proof}} +\newcommand{\isasymQED}{\cmd{qed}} +\newcommand{\isasymFIX}{\cmd{fix}} +\newcommand{\isasymASSUME}{\cmd{assume}} +\newcommand{\isasymSHOW}{\cmd{show}} +\newcommand{\isasymNOTE}{\cmd{note}} +\newcommand{\isasymCODEGEN}{\cmd{code\_gen}} +\newcommand{\isasymPRINTCODETHMS}{\cmd{print\_codethms}} +\newcommand{\isasymFUN}{\cmd{fun}} +\newcommand{\isasymFUNCTION}{\cmd{function}} +\newcommand{\isasymPRIMREC}{\cmd{primrec}} +\newcommand{\isasymRECDEF}{\cmd{recdef}} + +\newcommand{\qt}[1]{``#1''} +\newcommand{\qtt}[1]{"{}{#1}"{}} +\newcommand{\qn}[1]{\emph{#1}} +\newcommand{\strong}[1]{{\bfseries #1}} +\newcommand{\fixme}[1][!]{\strong{FIXME: #1}} + +\newtheorem{exercise}{Exercise}{\bf}{\itshape} +%\newtheorem*{thmstar}{Theorem}{\bf}{\itshape} + +\hyphenation{Isabelle} +\hyphenation{Isar} + +\isadroptag{theory} +\title{Defining Recursive Functions in Isabelle/HOL} +\author{Alexander Krauss} + +\isabellestyle{tt} +\renewcommand{\isastyletxt}{\isastyletext}% use same formatting for txt and text + +\begin{document} + +\date{\ \\} +\maketitle + +\begin{abstract} + This tutorial describes the use of the new \emph{function} package, + which provides general recursive function definitions for Isabelle/HOL. + We start with very simple examples and then gradually move on to more + advanced topics such as manual termination proofs, nested recursion, + partiality, tail recursion and congruence rules. +\end{abstract} + +%\thispagestyle{empty}\clearpage + +%\pagenumbering{roman} +%\clearfirst + +\input{intro.tex} +\input{Thy/document/Functions.tex} +%\input{conclusion.tex} + +\begingroup +%\tocentry{\bibname} +\bibliographystyle{plain} \small\raggedright\frenchspacing +\bibliography{../manual} +\endgroup + +\end{document} + + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/intro.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/intro.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,55 @@ +\section{Introduction} + +Starting from Isabelle 2007, new facilities for recursive +function definitions~\cite{krauss2006} are available. They provide +better support for general recursive definitions than previous +packages. But despite all tool support, function definitions can +sometimes be a difficult thing. + +This tutorial is an example-guided introduction to the practical use +of the package and related tools. It should help you get started with +defining functions quickly. For the more difficult definitions we will +discuss what problems can arise, and how they can be solved. + +We assume that you have mastered the fundamentals of Isabelle/HOL +and are able to write basic specifications and proofs. To start out +with Isabelle in general, consult the Isabelle/HOL tutorial +\cite{isa-tutorial}. + + + +\paragraph{Structure of this tutorial.} +Section 2 introduces the syntax and basic operation of the \cmd{fun} +command, which provides full automation with reasonable default +behavior. The impatient reader can stop after that +section, and consult the remaining sections only when needed. +Section 3 introduces the more verbose \cmd{function} command which +gives fine-grained control. This form should be used +whenever the short form fails. +After that we discuss more specialized issues: +termination, mutual, nested and higher-order recursion, partiality, pattern matching +and others. + + +\paragraph{Some background.} +Following the LCF tradition, the package is realized as a definitional +extension: Recursive definitions are internally transformed into a +non-recursive form, such that the function can be defined using +standard definition facilities. Then the recursive specification is +derived from the primitive definition. This is a complex task, but it +is fully automated and mostly transparent to the user. Definitional +extensions are valuable because they are conservative by construction: +The \qt{new} concept of general wellfounded recursion is completely reduced +to existing principles. + + + + +The new \cmd{function} command, and its short form \cmd{fun} have mostly +replaced the traditional \cmd{recdef} command \cite{slind-tfl}. They solve +a few of technical issues around \cmd{recdef}, and allow definitions +which were not previously possible. + + + + diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/isabelle_isar.eps --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/isabelle_isar.eps Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,6488 @@ +%!PS-Adobe-2.0 EPSF-1.2 +%%Title: isabelle_any +%%Creator: FreeHand 5.5 +%%CreationDate: 24.09.1998 21:04 Uhr +%%BoundingBox: 0 0 202 178 +%%FHPathName:MacSystem:Home:Markus:TUM:Isabelle Logo:export:isabelle_any +%ALDOriginalFile:MacSystem:Home:Markus:TUM:Isabelle Logo:export:isabelle_any +%ALDBoundingBox: -153 -386 442 456 +%%FHPageNum:1 +%%DocumentSuppliedResources: procset Altsys_header 4 0 +%%ColorUsage: Color +%%DocumentProcessColors: Cyan Magenta Yellow Black +%%DocumentNeededResources: font Symbol +%%+ font ZapfHumanist601BT-Bold +%%DocumentFonts: Symbol +%%+ ZapfHumanist601BT-Bold +%%DocumentNeededFonts: Symbol +%%+ ZapfHumanist601BT-Bold +%%EndComments +%!PS-AdobeFont-1.0: ZapfHumanist601BT-Bold 003.001 +%%CreationDate: Mon Jun 22 16:09:28 1992 +%%VMusage: 35200 38400 +% Bitstream Type 1 Font Program +% Copyright 1990-1992 as an unpublished work by Bitstream Inc., Cambridge, MA. +% All rights reserved. +% Confidential and proprietary to Bitstream Inc. +% U.S. GOVERNMENT RESTRICTED RIGHTS +% This software typeface product is provided with RESTRICTED RIGHTS. Use, +% duplication or disclosure by the Government is subject to restrictions +% as set forth in the license agreement and in FAR 52.227-19 (c) (2) (May, 1987), +% when applicable, or the applicable provisions of the DOD FAR supplement +% 252.227-7013 subdivision (a) (15) (April, 1988) or subdivision (a) (17) +% (April, 1988). Contractor/manufacturer is Bitstream Inc., +% 215 First Street, Cambridge, MA 02142. +% Bitstream is a registered trademark of Bitstream Inc. +11 dict begin +/FontInfo 9 dict dup begin + /version (003.001) readonly def + /Notice (Copyright 1990-1992 as an unpublished work by Bitstream Inc. All rights reserved. Confidential.) readonly def + /FullName (Zapf Humanist 601 Bold) readonly def + /FamilyName (Zapf Humanist 601) readonly def + /Weight (Bold) readonly def + /ItalicAngle 0 def + /isFixedPitch false def + /UnderlinePosition -136 def + /UnderlineThickness 85 def +end readonly def +/FontName /ZapfHumanist601BT-Bold def +/PaintType 0 def +/FontType 1 def +/FontMatrix [0.001 0 0 0.001 0 0] readonly def +/Encoding StandardEncoding def +/FontBBox {-167 -275 1170 962} readonly def +/UniqueID 15530396 def +currentdict end +currentfile eexec +a2951840838a4133839ca9d22e2b99f2b61c767cd675080aacfcb24e19cd +1336739bb64994c56737090b4cec92c9945ff0745ef7ffc61bb0a9a3b849 +e7e98740e56c0b5af787559cc6956ab31e33cf8553d55c0b0e818ef5ec6b +f48162eac42e7380ca921dae1c82b38fd6bcf2001abb5d001a56157094cf +e27d8f4eac9693e88372d20358b47e0c3876558ebf757a1fbc5c1cddf62b +3c57bf727ef1c4879422c142a084d1c7462ac293e097fabe3a3ecfcd8271 +f259833bac7912707218ec9a3063bf7385e02d8c1058ac06df00b33b8c01 +8768b278010eb4dd58c7ba59321899741cb7215d8a55bee8d3398c887f02 +e1f4869387f89141de693fcb429c7884c22dcdeddcaa62b7f5060249dfab +cfc351201f7d188b6ed68a228abda4d33b3d269ac09cde172bc045e67449 +c0f25d224efbe8c9f9d2968a01edbfb039123c365ed0db58ad38aabe015b +8881191dd23092f6d53d5c1cd68ebd038e098d32cb24b433b7d5d89c28ee +05ea0b6070bb785a2974b5a160ee4cf8b6d8c73445d36720af0530441cd9 +61bc0c367f1af1ec1c5ab7255ddda153c1868aba58cd5b44835535d85326 +5d7fed5ff7118adb5d5b76cc3b72e5ff27e21eb857261b3afb7688fca12d +1663b0d8bdc1dd47a84b65b47d3e76d3b8fa8b319f17e1bb22b45a7482fd +f9ad1b6129e09ae47f15cd2447484cd2d64f59ab0f2f876c81e7d87ccdf9 +005aa8fc093d02db51a075d571e925f2d309a1c535a1e59d34215c6cd33e +3c38997b4956461f376399901a8d0943dca6a335baac93fc8482c0659f04 +329c6f040b35828ea6dd1bd1858f2a9be4ef77731b5b75a1c536c6bc9479 +0821e5d88f6e2981835dbfd65ec254ebcf2cf49c917d121cd3bbb476a12b +69c15f17d9c17bb15ad1e7d31d2afcf58c8f0ad526d68615a0f1ac3b1d1c +d3beafeea3cf56c8f8a66367c70df9159f0b1b3157ccfd010045c4718e0e +625c0891e85790c9b97b85517c74c9d55eaca31a01cddc64898bf0eeadf3 +53391a185e507dcb0a6f52661a56357ac818dfc740a540aadf02f4e7a79d +8008a77cd30abde337025b01217d6a68c306abe145b7260c3478fa5f366f +b2d37259ead8a8ec2db2f09ae0eb3a682d27b0d73a60935f80254c24426a +003a87a29a4370cbf1b2ef1e19ad8466ec725fd5b463d06011a5e0da2258 +ff6c1483c4532bc21f2ed9b99b929b2800ddefc1a98d12ba085adc210bac +e0274b69e24d16af015a51ca73edf779a7534a887aa780337ad966839881 +edc22ba72038aa1a96a6deba24ad676795da711b92f8cf5f54cb4322ec04 +40ef9e15b11e3005f3ff69376ecb29bb66e8fc1b685f2b05fb162fcb35aa +d7eb2a8ec39b97ab1ff05ef02f8dbbc12085a5cd252cc4010fab7f63dfd5 +7fa1be86f724d37db5faef17ae8e8e537444e8e9db917b270344183473af +7f47d5703a160d8ef1e772438620d3336b2fbcf6433612e4b5e64fae0329 +8a3c4010c17d93f13ba66d549c69dd58c7d26ddc90285fed831918563a16 +2a7ac2511e2f18c9eb3df905a9dcba65a31cc1c39f07458abb11b4c60346 +aea19070e126982f1dde336e79be0ecd69a8afbe2493d064d4d2ff38788b +b3038125961302db9761403c3b8019ec641e107425002205a77ae2ae0f4f +7550d623dd03f0ec0199f42a9a8b89514e2e21baca9b3c8c96ca48cbf9dc +ee6d6241d713e014b49e83ad85e62a6b2f70b58e4cc72d41ea6fcbdd3b5c +890c8af0d24200658773b1628c6cc9aaaabb08865ee4c7ff2c513ad7aa23 +155a321213fa94731683a3e282a0e60aa3c87aade3da231465bdd9097f2c +89a1af8e5b9649143f2d9482546501ea97e8bea2f5d5eea97d4f19bb6835 +3138d3babb2461c08d537491aaede1f23d734c6f099eb5bef6e2ffaaf138 +e5ab71b8b41599091037e440127a4eaedf208c20c8a2fc62eadab191d1ab +4d5531f826aa6b9fff2797a7f54673e0a3fae09a93a0dfafb8b11d60dc69 +5acf9b7e1a47c31d0b5a0b85b7b50cddff5ac831651d9c7469c2139c7a89 +7d2f868f36c65156921803eccfdbdd1618595ab6d2a9230ef523a1b5ee51 +f2a0d200fc0e94aff7f546593ff2a3eb865d129895af01b8ab6e4616fe20 +9123b6e2b7e0817adc3cdb78ae8b0b1d75f2986ebd8fb24c4de92ac9e8c3 +6afa520636bcad2e6a03d11c268d97fa578561f6e7523e042c4cc73a0eac +7a841907450e83d8e7a8de4db5085f6e8b25dc85b59e018380f4b9523a7f +02cbeec989f0221b7681ec427519062b429dcd8fc2e4f81173519f88e2e4 +3798b90a84060920f6ae789afd6a9182e7fec87cd2d4235d37a65c3f3bcc +c742c89cbe5f3e2ba6c4f40ebba162e12569d20684cc167685285a087e7a +0a995fe1939bf25c09553512ba2cf77ef21d2ef8da1c73ba6e5826f4099e +27d8bc7b3545fc592b75228e70127c15a382774357457cd4586d80dc0bd6 +065aee32acfd5c0523303cece85a3dbf46853b917618c0330146f527c15b +dbb9f6526964368b2b8593eed1551dad75659565d54c6a0a52da7a8e366f +dd009ef853491c0fb995e19933cba1dbdc8902721c3ea6017ffdd5851cb8 +3c8bada46075ac121afe13a70e87529e40693157adcc999ed4657e017adf +f7dbac4bc0d204f204c6f47b769aaf714f9ec1d25226f24d0a1b53e28ac5 +374ab99755852c1431b2486df5fd637e2005a25303345a1c95a15a1189ba +f6f6883de1ad46d48427b137c2003d210ab2b2f5680f2633939f289d7553 +eb943adf8127f1c3ee7d6453b5566393700ad74ab86eb9a89f8b0380af55 +6b62f51b7dbd0c5dcc9a9fb658944d7ad5845d58dedc2d38200d0ef7cb0f +76041dc104ef3ab89c1dc2f6a75635d48051c8a7dd9f5e60253a53957ec8 +9d1266566d7ed20d79dfc2807b397d7cf056bdaccdb72528a88aa4987682 +c909b2fe1e35a71c2f29e89a2bf32173967e79610367ce4574ba6a1cc031 +cfb176fc0313f74f91a866ef9954b95b29caf917a6b919586f54d23cb7ce +23305886ae7760ebd6263df0d3c511ac7afc361df78bc2621f66d3268b99 +078fa59124f0eb9476496c938eb4584e87455dc6f2faa999e938460b31c6 +28021c652acfa12d4556aa4302bbcd043e60389480b796c3fc0b2e51b81e +c2afa4a34335318a1c5a842dcaa120df414acba2e79ab5cc61c99e98108c +5cb907a96b30d731131782f9df79aabfc16a20ace8852d047497982e11c8 +26321addf679de8a96a2d18743f6f2c3b2bc397370b10ad273fcfb76b48b +9dad27cf89ca2c370149cd48ab956e9bbce01cbf8e1f0c661b99cf19b36e +87b6165dd85ae3f3674525e17d85971093e110520d17f2d6253611e35ec9 +e17000e48e2926454c1e8eb4098e0115b49536f2c33505eb5e574f7a414b +e176398c5ddf6d846ea5ddf2a5e94c0422e0115c57a8c9e56bf8ba962c82 +698c96bd6138baaca7347e44352cc62df4eeba364954ad921a5a43280980 +264df4a7fb29d005423179f7bd1d98b4280d62ce23c927551f1ffc2b8f17 +0a9c23656c0c91b640cdcfdbd88089ffb28d3ac68bad25dbbed82c083696 +1f9f86a6183cc1061ffdb32279796569d05b31c946955402d0be1fb9f2bf +304d1ad8e1e357be49e6e2ee67f7f1e7bc699d46a5f7853fe659ba2e1930 +0d3e3ea658b9862701dcab08fdd23bf1d751777f25efbe9e02d12b5612b3 +c3fc2275190346b94ec4024e4ade08e54d75c0b4c48f4956b9182e8ce997 +74b54da4a9318c099d89f1ce3b6803a95f48b9fb8b845372be25e54478e8 +49e4707ea03a36e134efa661e4e6250d89649ae074cfd9d6b9e2071c1099 +3b8a5a5ebc3e1cb228c97565aef7f254e3f90af8a3dd281c83792755719d +c6a5b3bab4aa6be5afe9624050eee8dfb13b018f4088c932cd48ace38dfe +b1b4218dba8f7fada6628076acf1b54db0c95d4fb12232f1fa9d1ba848f9 +fe80c65b75d6946c00fe78839142c5302707d9269b24565dbcc551aca060 +b4d0f99b961dd3cc795a982063ac42e9fc81fc98add42744a9f92e74b00d +637ee4606ea2099b6c763493c3159f8e52a90dafca682010c0e92bc9038a +10abb066c75c8d97f7ad6fb1a37136e52cf2093c4fa485fe12adad10e4d0 +83b78b023628ddc5326cbf8392516027a9f3de4945f93488e4a1997efd2a +22c2c136dbac1bdb829e082beac48cdd06dcb17bacf09451c7b636bd49a8 +fc60cb1d2292250bea78e1dd276725ab4c526b66ddabf4e1b2bf0a2571df +2490df70735f5da321fac74fe4fab54444e53dace9832cff38a70c58104a +4f0c0545dcf7a1a9ecb54e0e32d6d8195d30b2c98a567741fcf325a4ddeb +244a1a35676e246ddc835fac13b569f35f22ec93191eca3efbe17ff9a950 +d08f863710b2bbecec969068c498eb2338b68f3fc3f5447449fe4de2b065 +e068ecd9255d369b2bb6c4b0b7423774bed69294758aca2bdb1a8d5bf618 +d3fa09462f7838e8a79b7a53bebe6dacb0a1561eaa074bc6f069f6a06fb2 +c4a5cb13e2172bce9be659a82665da5cded73da84322bb16aa6e19ac1958 +7515cb5d2b65e65e803f76700ce5efd3df7fe4ed431fae0e8e286b1d5056 +a0d18df926b2be7a93c503ab903abd4788680a6201fdc299f2cb5d6a9b6e +2048109c8d1fb633a54128938594b2cce86a7e0185e7d59e6536584039ec +9e30ff7be6ddba9fdba82de7415fdc47de84d97afb1aa3ba495bd91dee9d +f3b21ee1164987dd8510925087cd0919f1085cba6e4dd3c7384d26667f94 +ad7f736a60d8bd76dfaa4b53c23536fc309ff2317c48ee0107ff2ca4d1b3 +f78c5a27b901c931128bdb636094ef0cd543a5b62b7dbe10ed83aed5780c +a16067a4a7e8b7c5bf8a8e822149bc1b5bcdabe13a7f6aa6eaeff24a42f4 +a58a2b70f545103540169137fda9abb589f734b6776cb50402d6123ce802 +10dce05e3697a98c9411cf60a02c278c91e03d540b936cd00c668960e357 +1aeaf4d94cfb496b259ec0d8fdba9199fb46634ff177bc8d310ea1314eef +d46c927a981c58e88743ed4e07d80fe841edee812e3053412bf2e710146c +b25dec8ea70c38bb1f6e4db3c2e8ba521963c1584eeb60ea1e9555058f13 +e98307c13cbd15c26b611f543149b1ddf88dd6296ae703f58daeb67f1b03 +ab5b24c72d5770cb9d8ed242c4faaad1dd940ada00e98ff3a61799d13355 +aba916910aa9a6e5ee8af438d0ba8235346fcd139b9d2cb7db7bd3f298a3 +94ff0aff3b9042f32a004e042c346a5ea35917f229848a9c5a32909b0090 +4aa715640277a6ada99f8b2927fda22899ff1347f42bac73e2bd4bbf3945 +55fd7dd30d5c6dadf5c259fdb2455d8c607de1c5da588e20159f18e4e8da +b714e532d888a0386c52c2b85964251af003ac9d10c0c8b9b3465e1dde48 +2e74a29e17a7cf6c9a43b5af1646f0b8508f98e9a1279ec3908073d89dcb +aa093e8dd1004c1ecccce0803095b0069d4be7a1eb14b02efc37d137dfe3 +f0190bc9628069abc257f45d0e050e60c7f5281277937dd986fcd5b94a2b +845a1a75addd74a142800f18ef408c08a2c2ad16a93298f148c0ae7d2990 +ded147f92f053809a60d4f992a227167aad5b5eb2bbe8a4a77dc77a08b72 +6acb34422e2532eec7e274e4a42d71ee584f5f3df5a3a5b379974ede73ab +5f1b0a2dbfcc8cfac4747ca26eb6030dc2f85a104b754432977850c093b9 +97ed90af711b544ff682e7b1eac82b2b7b44014b09c17ecf084c994a095d +9eeef5391305caf093b62ac9916f982a436d521fcf4d75c5b8e4d92266fd +e99a58aa39d7693ecd1796b8851761d64bbca39a6d5a0b4533ae47123327 +f98d0ad0e8b36625cc3647b55459552906d8a1d5766845ffac101980efcf +79657e365510be5db557cefef21193ca3cf3dad175ee2e7ae91d174fdc06 +2ff5c51ffe2f021122e96df042019d3a1883e662537ec1b69c11fbb6e750 +0132eabf5803c816153ecbff60ca3b3b39708c26cb1751afb2e65d8e5f4a +c4397a88fb1f112672fcdd24e4ba545c5b2a7968c17b62f8e2530a8acbff +cfca82c64b7abcab84e2c4a0a7ced67b15669301fe9ff2c756e70ff7ce33 +497be6acc4ac5617e1f043bd8a87416299a39bf17fc31c02d72d75fdc2a1 +e60669fa4d5e4a49d9afea2f53f4626680e9c0dfca223529efa415c4343a +b6067aa800c484457ea050eaaa5d3fafeedd0eec72f327e02c6b3912b5a8 +c404de4839c9c4a99da42681cde43316606a34c7d2f02269de1aab776857 +e668f35946af4d618d36d444bdc02b1f63ea25b6260b4fb606ac8575b5c9 +782a5de4037350d5753b1537537ccb008c454eeb264e6cd4727c999e240e +0ac89e95a896b67d54910a3531345f64198ad394b5ceb52881f1dd9e6beb +95862dc188d45b3e46aacb5fe40097947dab9bc3c1ee46bfc9b1b3ed6167 +efd0d65ceb043d7b24c1456676e4baa47b1209a315f199bb3a91f4374cd9 +cc0b40d3f09f19f8dd8a46915eee55eeeeb3c7b8f437106ee491ef0f4ff9 +2c5c6f779e0fbe7bd5293964bb645ca362b106abeb773571d9ae83b786a3 +d5a4ea3ea970daadc46cc5e6037f76fd20e0fffc47cf4e7af9522b91f96b +3297720fd45d9bc2200622ad2ca9445556c8a8202b1991bc63da360d021d +55be2528e043f803e08da99b91ab9cfc5e65b2655d78206b4aecd445a7b0 +1caa0d06b0a55e8f04b70b60b04a860c8e1ab439f4910051e3f7441b47c7 +8aa3ab8519f181a9e833f3242fa58d02ed76bf0031f71f9def8484ecced8 +b6e41aca56176b6b32a2443d12492c8a0f5ba8a3e227219dfd1dd23fcb48 +fcfd255dbbf3e198874e607399db8d8498e719f00e9ed8bdd96c88817606 +357a0063c23854e64ad4e59ddd5060845b2c4cddd00c40081458f8ee02c7 +303c11747bd104440046bf2d09794fca2c4beb23ed1b66d9ccb9a4dd57ad +a24943461ecc00704c916bdc621bfddb17913dfb0f3513b65f3ab015786a +caa51ee9546bc8ddf87e2e104137e35ddf8f8d23724e9a53824169bc7cfa +99562656e6f1c888d4dbff0b269c5d1e733e5f212d91297610201eb43249 +35e336dd0052738db2d64f3e89429903bb5c1810009cf766e9a06223dd2f +219b706394a121dc029af55c6ada9052af59682ef7c51e121cf16f0319ac +0aa9512ef900c548d673fe361da19052808797e958209072e145d46ec8cc +a89fafd76630eff30ae979973bdf0f8c9e469d8edd3b1c93731c72f976b7 +d81142bc15c376403f967affaa5f482efd57c6f91970729d16db851f0ed3 +ea7d82f409307b5b436886c1beda94a1fa3ab1b60686f6574c844fb2c0b3 +a07174dc4f27b4fed2f8bd4d5436be4b343e5efdf0400d235bd910255341 +a20770804a26f8437e9bce6da8e9f8258a343c7aee291f1510be306ae67a +ab1d7696453530c02fd153bbe49dbf62baad6146029cbd1656cdb76c952c +b93edfee76fe33832930be59636bb947e8ad285f20f663cccf484fba97d6 +7446c7b6c6f5857428bb1737d9ae801df75d9cb4d7bd59ef7a4cbadde928 +38f15d232005585d2e40483d2d3e08cc8f398bb43afedb84343c3ba3835d +0ba82a86dce859cf655f85e63e41365e0dbefcf511b9a27a2b6e66b2ad3a +c657902842287a317e46ceaa93b5088f09d53a65815b44538af90ad3b06b +4e5e2dc509f02e30a01e05201c67d4d39582bbe64e20b669f5fd787909a3 +30fc50a95b31426bbb57a4fbf9feacdc31f98bcf50da7e50c2bfc169c6fd +ccf213cdb878653bcea372968ea6e31fd30dd55434cc91c0af22179ce669 +a05493f195e12432c6173ae2ac3c94fb83f38210014a9f969ea2b44e99f5 +e5a7317e848d429ad62167a4fc5001149676c0c28cdf59b8d1c5a582f516 +3eee855312777fee6dacbf993f5c058f355dbde6552dc960d336eee445dd +11d53fd21b745d1e5ec317efbbef25e070d0a36797a6081c356ac2328e64 +e5c55fbc81dc75d9c1575548ece74b8307eef485aa8e28859a2e0435c831 +23a600efb323c362fe9f02407a5411c41a69566cd50add324b63ab939980 +b9d7a929ae4887163cfa7acbfc9fabaab8987a1f6906b9881491cd055b94 +485c968479dbb05b34ed0cd6844729a692978c6928c3392e33e8324ded88 +814cacdac8128e1425c0091a13558100d7cdbed5992795d94d39c32f32dc +621ab6f3b75187a66741f61d6a9c91d791b1cfc3d0e94d4a76302e0c3f2e +cbdc51f14f3251aa5c8bb989f0e13ee500b7b7f2f1e52ca970ad8a7b4b99 +57e93126254297380d67179deb8ff1e99d5cdf7a35c5bb9fa7b402e63234 +78640344e1f10c378ad23c5cd1aa18e1e0b308db70d3a624a455f8e291a2 +ee102ad10776208c2d546cb76d89ca8103a8b95f8acc2d2bdc9791324915 +6c9e05429091071f0c5b76d82c8d1c8a69d840fd460922cd2090624bc218 +0c9391005926a25042a55e322060807363462e1cdeab309033124ba3a884 +1db13f39edae04ec52cde9dbde64ddda1ad805141d4230ec76bd81fd98d7 +0d90fa1aaa26ea551bf687ddd6cdcf3de5a446b266c68434f07d9c0b382d +5816c4e22f22cc03ff78064c6dffb12315c6bcbbf5dc510f5aaabf23471a +234efceeb4aa2f9af9ea787c014c5587ef162fc5b35e8f4c23b168c6e247 +41d33dcc11d2a56d3ba9d8eed6e79aebf9f0faf1a3aeb89d792d69041f0b +b8fadfc0aa090effc6ae5e2f13cdbf54b5bed69b039eef2627769613b6f1 +aefe9b66747fe8feaf7455796740f411a770d4a1764f0483719584880f45 +430e38d3af184145892a08b2add234a3f3ee4ccfc9f6995c02392adafccd +722f366d748cfe9373fbf5f878ed47e9d221fd156bb28369df9e7d2b79da +76120d135ebaf36cff93beb7e313c2b2de7477176fc19609a1b906c995cd +defef08899265b6b8aefb44da1aadefd1c523dce5ca1b84c0c652b3009fd +057789892d4d31764f181754b2e0a62c465587585509989a219711a5e4e2 +5b3b340ca8fdd3f04fef204b1b722b2f6c2ccb00c3cf1a94ba9bdfbfeda9 +e2a062c6f1ced3b8aae5dae32ade1fca1001f98d0ad0e8b36625cc3647b5 +5459552906d8a788eb8bc734ccb65fe9582c71df94fd95d22c5323de235c +28220fb9a2ccb37362174d8cd5922c9e5a87b51d0668555100a33e33750e +f1f795cbed962494a994be7ce8cf71fc58ff4204551b1615ed27cf088171 +fd000b72462b67935961e7c6c3a05bfd67b9ba094ea2c16fdf486da912e1 +e97bfd1c17934535e551cede20c001b5d2adb2be4cbad7d6ba0bdeae4b1a +a739f90293e67ecbdeea4d35825e092697fb05b215083e3f3d6be260790e +2a175fd44eb1c4c16759504827a6eb58a838c4d65fec6eef108495577019 +15740cac164111892e8d1cc447cd208e243a89ab847d8ebf4fb98bff49e7 +a3453facf3b0e8cb67590f390173ddba68324531d2e426aed152e12301d7 +538c1f3c0048a9cc00c009a1a9138460082123209c1e007266fbf236eb72 +21f87d4ca38a0b699e84ca230ffb5095f90a6528bf2a9118f95ac9ab8d2d +ed9eed9b8b27be894b717469758c8d94fa89acc64f530f432d0e5f16c922 +36d6a63410f099c9e909450fd731d698ef658d8ffc1de14817b850814f68 +1a4a9be5cc7a71c381974c249f0b209bfdc2e97f9540c96f57bb4d283622 +00969b82011315289e6a025b137030a0af3b4b42b00fed7cec49df43c59e +3b2495a036dd1b17a8e6adae63bfbbd48807c44b5bbf71813355e1b0e58e +22b6fb88005fc55565be49c17244901b73ef02fc4eb7669be5af22d89c0d +dff0fc6821d810d13e5821d48d4a71d7e463d5b60bc279d0dcf5f8da3a95 +905b56d6f2be95e6d4243b1048e3b662e62401ffaa3bc3f5f90b0854b8a3 +8c38039f61fcb359b06bbb7d59e3b39a295dccd6db9a8b83a6f64ef8dc94 +a77123dd164cfd1c46f1ee51aa19c3d6e7db92a298d10159f2b5eff2caf9 +dc93a6d267fb65bd900d6adf0c6be598050b6d3a9b3a322ab3c9e880d774 +1f58016ff97e5f606b5dbd72ba99252c669209bb556dd5be84fdd7c1ce92 +8a3b3d3aab8d37e6b740227563bb4d60f6bb04052356e1a48d2079feca44 +7ea17fd06f208426d045dee660d1d6460455f8d20dbc5ae64550bbdf60d7 +27d96cc9afef842a8c8c78ea2257e6c6d0d207c80cfe399e8874c693274e +d2c2022d303ca50a70624b07434fb85040a76a823f446c7454dab4f9c05f +10274eb5ba164aa3649d1bc90694316ba5cb3e7df4442e777124cff7ebef +53df2320a0c441ab61666493cb43da46d5711c21699de85bc74359444da2 +e3e397d4c16234f81531505b621aa242a6698886f82b447104b1f1062f60 +b5c87cea9151bb3c627bfa4532b06fd147c556ed8d61ae30a8719dfb8705 +f8a6c74368381403640cc57026d3790c49e2bbd1c0e48285ec6ba44de678 +e3a1394d659c412f09644b83ee1a333a1f51ad8deb4e6d77b3b226ac2c4f +fe653411a7976ae7c4a3cb7df309788da6b483f8a7bab4a6990db74362f5 +bc41d545a320389b2599fd726e426ed9fa2916ece67b058f6a269544e517 +128bda38d117f402409d0d8f8c88ed509aa2ba882e0c579b45af4be80770 +22d7269684eaf0f9afc3054316da6611e3fd260d67fb6fe52c9ade5dda24 +a0050a819ed21342aac9d25194778beb3145f56a66980f620998923521ea +3f957b6ed0c5470734af9f416a16427dd03eff9a0e023452097d4ef936d5 +49a90823cef6de340a1ee02a52851b310cbcf41ae274947a62f9d1d8702a +669023e3caf967204a340694b45fecbda4bf9552f6bdc62d43b3b2c3d571 +9983c182453e22ee34241ab908e667115f7988174684cd70084aefc55caa +f5352a88e9dac45d1ea0e032af61fe9a9118a3931b2050fc6db66ab96a39 +74353b597f34dfd9f72150de23285eda5e555a607d198c291965a7233715 +3f4946a57af0b440ff8567b01a6f46c6d32fea5f8bf57d89dccbab7da882 +ee6c9260e89443b1d7db099477492bd0468850df3db668d741123e7ebe3d +c21748ab4c5cbeb5de33b8963aecafe76bba0c4f6ed8e8263a116ed85e58 +fb71ec4ab0071301be7c7d3afd5fa6ad46c0232807bb7fe129e44bfd16e9 +fd0c8bb5e7cdd86a78b5fb0669093c22eda9151d85b6f58a9c8ead3727c0 +09850bd31a8b4a873d0a506240bb2aeccb8dcb6369532f21d9b967aa8443 +fd6d77cb2d65c4678a5fad188db85940f0a187aa1031dcf5b8e0d0cbfb6d +b3b96fedec5b249b7a69de9b42dfa605bd622de7a220cce9b66e9f3394d6 +13487dc5e82c1e619079cd057b1e19ac05ebdfd7c8bf01c6c66fab49e0b6 +613df9e42beae2f7b9407a2bff8896d8035cea0fd5c11bc5889cb3d90876 +61766138d2625f42d0244adca65d1bc73989328c0eea0b97c7c766285ab3 +351ce2b183f774488a8806c33178090a3808f0ce5e339b87cf7add933301 +ca486742831ca751f0626864ce13172829a8419af5c78794a0eaa17b5bcd +fcb684f7d4bb7af15deb432e44dc7dedf56eb8bea08b46f1e8123a49a349 +a7cbccf833a528f5e22d2d463040e09b91e543a2f33077b3e7b9ecc64f14 +306186cdae1fc317a6ced7e9b4d51a10bbbcf2fadff876b4d9082e3f4aef +dfef230e4232572f4fa33a6e065f6895aa2ea96c5659cb579b023179f0fe +de7ba64bbd9362a7b2b8c4eaec254915629e81d01c839096339b99bc9e25 +84536955feaa52fa20666f65bafd9b2e69c3e8c15d24fa407e7d881679b1 +789a0e2a695d13553c92c0214c9b7562cd6a9a3d77c8b0c2196cef76dc51 +d855c1dac37f96eae4cc7bf07e17dc7c08333d7af33c8b2965ea1f23446b +3c96c52b30ea628ad572694d145b58a606f90b278290297aa372cff56b6f +56f4aad6612eb7c7bd07db4f7d1a70d8044d16d0b5c1605ee02a852ffdb4 +450147b3f9b87d72dc431b34fcdc899462dcc1b6bb6ab1758b6a589e91e5 +8f5196251d00133b43749b7a11fb67a22664c5e38e336dbdeb5509c2d9d6 +2642c07275949df0e2db59314ae0fb34641fc171d3fe1289f919136d853c +d9048ee9db50c699c49e27a8df199590bbc65b23b55bb387eed0c73f2db5 +1cb091f8c22af83103f214199e371f7de1df23f757817200be30610004df +81fe8ed6eba79e856fca21a126ca326ad2f313c16e15754663ad6a065e08 +4050ff005fc899d6e233691b918a093b5f1ffda8839ab23ae66b1bb7b953 +0a7f896ec55de6fb9faf1b49656ff2e57488cd7f1c44114c75f9d571461f +767a6040ffa14e9fb43096f164d60ca530d7cca76d526d1999ac1b52a793 +28651112a65db1f2564ecf90ea6bf2c9ecf515640719c3fb5e36cfc58591 +e227793f39b9d3a9025cb10f324a95c29c488724aa74812366ff0b118fc7 +19f9fd0f202a040be47ec99b46b4dfc3d2a17902a5779c8d52b27231a1bb +5cd794c838daddc3e6824ca8297ba669a818c239b389400faf17aa04b802 +f763029edb9784dfdc42f223e6496a938e613463bf9bbbd59d63300a9ad7 +4e71865cac4b4e81a5864388c3886e70799c8989188341f7d17cb514cd99 +3b211883f171ec6402cc361885f4f4b110757bb3e52941a94bfaebb2faa0 +3e32eb72e25e31abdde82c2a9015478afa0f434ae3f8b97a4bef598d6eda +44ffe1915c26ee0e8339d2d45a6a080550f538ded5542c8b96ca2f596979 +8bb6223e460e857516ab5a3323136ee8fc4b0556a7c39d0cf7acb45e48be +4ae9db325e4750b73289e36a61b301795bdb2ca2a8b933be1c09fd0cd2cb +8677df171d36ef1519a2269b21e4103b2ee151c513df3e10b2a216d6fb22 +18bf2005fa7e0f0563ad96661a7f55e1b5b991f8ca285651b2683c6a7c9d +2d1941374989b06f2e9b42a6af60193dc758dd8e9fcfc7c1aa06eab47e81 +bd79660666defac0c6b9e484df9c17a61ce7a61ef73150e8cd406af6da17 +4d9c2392cc420eddda40f975ffbeacad8ce1b4e14bee29ba8552ff03376f +c034784b38dc1d0ab7bc53943d2545b03d39797af8d58d6dffce56a353d9 +bebc833f04db321ca8642bbb7fcc63ed2349ffa08a33a5d0d78f4fd2c5ea +4258e4671e362036f1f67fcef9d878ae2c203fd9c05200c59cc98633e65a +99d912ec51d6f74500d5358b70e799a6817f59adfc43365d7bba1fd6766c +5c8e76248daf3f01e7a8950fe875d657397797a45e7f99a92887300b6806 +b86db61e03c4c09d6cf507800aeead874a94e6f665746752937214302045 +0b19cfa8db69230517183a03a16e5503882ea1e419c333d3e3b73cef6762 +873ac06bec34c3f736494483442619f5bbadd86f128a5a40b854051893ea +8d31dd6656777ad4ac2572d17c6fb21385b053495d1270e65d78334a4115 +2787ea89b86f97e72718905a11e9c5664837701a3c1c65ccaf26aebe8dab +c1207d5da2079c37883d9235708f370203b3b2a8ec3a5bb35fab93dae115 +aef626dc44b67ca56fac18caf1c22e6fbab93564829a75776630b9c42513 +721ca0fbb0b402f4d1db8f701d2b29fa60162feaa8a167eb3113c6f57036 +e8361357913eb24dd38dc6d3bf4c3176a07ffc75cecf8e5940a310f79a8e +f590844383d631796ade04a91144d073a9413cff34fb454f1fd75cfbe5e6 +525c3bd36ddab80138f6c19aad7417d47df1f1e0fc958fb190a8205b5321 +7c43a4dcb0599be404473d6faebe7240dc402a0e0caa21b56a601b154524 +f44988e5074c71ae8e1948bb2a2ce72fc24cf3b1813cf7408a6b097aff22 +f9d285134d09b7053464259531eb7b270cd5f39f81bbf41a36420f61e5f6 +b429036bbf20e27af1a437becd74c5bbc25ee2519402454fc94d430636e1 +736fe65a643d9b9d21c9a54eac5a8fed51ff60a47b85a0e9423e330e00cf +220c23e056d20aec2fca3e6bc7a61a8366eb940c9bc99fb90e8704e27655 +20335a983eccc7e20b13745c4b4f30a842f1ba64745718c152697c688c73 +6cffcf5cc8eb5756201560413117a45ad3d264291cd51404f98448d31474 +d47d17d201def12867ba679f0e2605de8f3e8135ed0234890cffa68848f0 +6de427741b34c2ea654251ae8450a152538eb806ace3ecfe86d8c4a137ec +c98c6d6cbdc191a5f8f5b5972c70b4896960037b6d4c7c63586a52d5eb59 +47af8c192eb980d0801fa670bb1d08740819f9da1dd9e153010bf9580a1d +0925d8327ea1b88db8d934f40266ddf93e5ea137f267847d826cd7999b33 +c795d0ac05abe2ec1770dd98eea67912f1939118defc9b379e237d6477bc +91ad08e0046b0836fafa1272b0213dce990c90815f5b30d0eb103ac9539c +2f7bd2280264cd95b4be84cbc5139a7628ed211905dcb92cbc3180ac9e6b +b9ecc3cb08608b2395827d5729781dea49d328ba0c1b4cf2cec9f6bbc822 +1f2bbbb9d88f9e7682b9ecc06b9705faa8a90a51678183db1e24cc2c4307 +e16b3c20f08f179ec69df7a8c4261427f5886f9179c493bf2d0ef36640d7 +79925585724aba69df6d1b4f0bd2a356eedfd74a88bea667b102420c2300 +ec420e99b9ce8be1472b617e1255a7f43a0b52f11657f1a4dbb624a24886 +9604fe2062b98f5787d010723e520a4f42a0c3943e151ee627f3d5db90e0 +7747e1a88a53c4784c8d2b042b9c23c9e436d7d88343171161a364cd8961 +37a19582a00d774ef01c7c3fc9e9c7be5074c858d2bacd707a6a4f322027 +137d6ca0421ed9f9c7e7229e867678e5272cfc7156a419e893404ad7dabf +a5d8b6fd0787cb4fe1a901c34dd931f1b64f0c470ff807005fb66350d0ea +eb84ebef2c2399cd14a4454ea5004bddd99988b39c4134b92121ec77faee +55cc716eecc58b594b39c41dcab308efa4458ed71943ec5805dcd0194ddc +1ba04a5d3d42d07ac62a907ea25cd2a7e77aba470324d41dc1a3fe088388 +787b3312f472cb4f23a414fa5f7c7a0cc5d121d7642b5b3f0cf7ca2173af +3f878f374938251feb3ce5ddd2d7703fc79a130978ac516daf70ae903799 +28bea3a4296f48725d578d2e8fb0f932e398404fa8a242024bc011c0ae81 +7b92bb104712253a5d89c543a744332069e33ca08bd133211d233ef799f2 +fed6a20a9073021e505def8b79e1279dacc062cfd4dddc2e8e0a7fda5dd6 +bb5a745f99cccb7ec1df532308da3da0f236c74639c280ea649b2f7ec27d +24221470b642567f3b2e1cd0b3ffa65c5ac986b557aa9b444bf470380435 +abae9b51c6da7ff753810ca7938d8a1c47d2b41fafd236cb5998f3ef365e +1f700bb257679ba3a82e235a3e97a667a6ad94412839c96dcd49dd86ccbb +6df8ad01756b311e9fd57ccd2eb2f19f035e214804e2b77769319a5389c2 +35f3ca2a73c616c9ef0984abcba167d7d652b330c68f4f6378aba69628b4 +2d59eaa2a7e4c782f6eb96f6758d17d35650b15cb5de9bf973b3b6f67c1d +f3285be8322fc2b44359640a3ba5d6d7b96142583a00a9a0ef84fbf14046 +09ad55b2aefe8c5c8f58ed21623bf765f81dbb6cca6d2a51fb7730a14839 +392cad6b47f5e03448350ab36a37d9ff2b9dab69be5196511072b10cc91f +2e6b5160b2b1bd112e6c02d14063a9bb46977b0d4bc79b921fd942f916c9 +c5708e0d133c8309de2f6ee0b1afc996c889c36de20fbbbfd32878f477cd +7735c7c3fa59e9c46e654ea20b4381d9f6c6431082e6918d532bcd539284 +af0333a783c9e7fd4fa1e4da5ce8fea2ea4037644a24532d65fa5c1ee982 +89e4b9abaf71a35d308a9b8c337f70babc5fc8dbb0327143707ca5b675c5 +2d3cf09f7a4f667fcda03d8c82d157e661517787ce6bfb35ea772de13c66 +2bd24b74ff9ab0fbcf6635d8e06b54b5b3125d17ae13d175cb7922338ec8 +9d1159fea2110995ce48f7d2b094f06d11d59b3a64a44a83d48c78855e47 +21243e82d9858401b094a236fa0a90d61863931c30d13b9bf33a35ac0d11 +a999f2b4dfba6fc187f8c235a5217d777a5a97112e7db6a8a4b06b07d9c9 +f41820e233c8b58b9e47ac56ad1ddcc0b35dd03976bc776c6ac3692ec0ca +f8c75ea7825bc84156468ca7b269d890ec9d4a365b0b31d2f6530185d5e0 +2acc3ce14eea55ebb5667067825a8682e135d23c78863d32065ddcf1a755 +e0de6dea7220d1a28416b96db40b1e9f159aeb070c9a9515f301f162b0cf +e32c6c89287de6e2b40458e3393826189a10af8517ff5a10c41c9d05d999 +aa9305a2ee8e7fe46076bc9c5722ee0a140a144ae383e84a8abe70af5d29 +96a0a896cd499caa0ed7867e7c3aac563763216e7769d12218b584d853ec +01db93ca22d0c8d6b286b20b6b26d6ef19f2cebe7030ecaa68d069fac7a0 +09d61770b5e8f83024a99142f59d88297cb8d093992c3c6c11b043b151e8 +20df640407d8bc829bfc196bf2901e63c6f16102d03ffb7c54a7a560f5f9 +5cf8379f4a2eccdcb604bd553e6157b4381940d1b3c768dbfbf2618812f5 +7fbe744b3d8ad680dd9223d8bf2412ecbb614d05b485e3b4669d22b417f5 +02cce2d705c208b15fa83b5be77ccfc1c840f385a58ae49fbe6ab4e53912 +473630e0cfecefab95ebc632a2b10a2103bfe801ca0302542080cfb4cf4d +4c241b1a6c8d28114516e3f1bf39dc02db73e6d9a797279acfd79b02a71b +ae34860dd0e11b18954129f8dd57c039bb7063a4c92f0f6a1e25f4ae59d6 +6c1cc6b73a79d6a56f7f2a8a64d571caa8a760f4f485d770d000ddf393ba +784bb27b781c47678dd78ae9b5d5e8b57d163c42c7a55e4aae22061686bf +aebcede728ff2f65e75955585208c176d100912836b5200a79062d4f09b1 +ba9465b0e937e289160ec543a4cedbbe0cdb5ecfbb4838138ee9e1ac757d +3c5f04fb6b510b389e2f521759e403bfc8ec6bd79e2d40bdd81901c10dd7 +4620acaac9108940daf03af23f09d3c8b785db562b05e597056406557857 +e96fc8bea53c2c2ccd0ea6572abb0acacfe29e737173d665ab6dc2995f60 +807aaa4073a183aed23c26c67eb137c937999fafc63b66a021125e4ee5c1 +a745ad1fff2bd828dcef392052965ce0e9af7a2c88d730fef69da91083fd +83d9fe9f73d42a8dbdcaba85b0fa93b210dbf49cdcbf5d4b69e07375fab1 +a39038cc51f66f0b10eebe0cc61f697f7025d9755830b2d65f1ad0db91ef +ebbfb578053de329935bb28d6ed6c12f748a2f70458990f04d56c35557e3 +8bc5d2e5de7f52bcf00c3bcce091aaa8852d53ac686f8f407baf3f7c8968 +69f3b62f44a5e2291aff9d30d7b5c663658a41add74562dbb0f1062f564a +9b907846291700151de04c1a55cb945eaa2e7a709218ec56d1becce1c0b7 +dc41d5f016ae8080c3b07311590a0def35337fc3c844c0ccd04926be9fec +509b1255ef12f368d20601b1ac8c68b0a935f987a21de0f8191604e921ea +0c04b00dc188fd73499852dbcccd4119ef799472b353be7f7dcc904ddfdb +920839f3d4a13bb1796f2dc886f31217845f8d7a543aabbc720311fd0e6d +a31ad3daa06d5e7e6270a34304f35ef170a7abe733428e96b0522fddbb5d +eb35aacec147067fe066c9ef145246fa3d444d176c274b91fddb8a7bd7ff +7cc7693c25895bf931eb321dc9d79f662a17691f9bd1662fecbcecf6d1f9 +cd8ddcda56d19811f05fa48bcb492feb355b0ec7c04d6046549c56f7799c +2cd0d9dade8809de7d510702e525ad9cc82c41b4fb36218e3d72e905c507 +159076a9c0e4a008ccca17bd594c69f5eee656426f865fc1988d677b72ce +b710b29a0aa8f8337552ae30e93bf7c6e5d013555872dba4737dc5f08c0f +efd428c66fc8da675373f13f89102688977e18e14dedd7f3b676256b0263 +b66b013617d9a026794b0d6040c23c5506a98530249633a6beec46117c96 +ec036eaf6439e25b8e57754af5ebaaf9b57880ad4fc93f002fb03e9fda21 +df4acb78296b0c49a5a852c134c3b10755177a0dbd6c54ea7a2b9bdac62b +5d7f3da649df856478e4baf97899e0f891a96536c283f5c81200c51c6ab6 +77285450c7f7e96836b6da5660f6cb76782ddfc64b6fc348ebc3ba4a46f7 +19176296d8c5a31132b3fa7d935a5d777c1dc84d669d564cb4fd689a38ce +680d0b3b130caea0be43864826d0d154019fd0d865f1c389cd367cb5248e +24640eb6f66603e50581f6fb5aca6cfec1d6dbf4196da10a5e1ebb14e4ca +0251c4c8412cc1673d6e7a9666b04b090567efa0b830d2362fd384cb0303 +8a40290597bdaffe429bb89fb66b9dfcfa92f39d92a8baba7266d144ac04 +f069093ebb3fcea961ba4497d3628ad207e0c8c4fac0e5f3f2a663a8d05d +b6dc33b890ae13d84dce64b495d24cc749b121659373ca31cee09bff2e9e +e5b62e89d5faa4482a75f341dd172500a54b98fc108a69a3ea94db696513 +d4c7691e0095ed3900cd4489ab008b5460b34ae8dedf3721c60de7086605 +6c391137cf23255c565bf11403bdeecf8bf39ad5e4317a4bb37003b2e7c1 +400c3b8ed7f63719bddf07908dc2decdb0f68e8ef722851c4420303f6de1 +b5efc9b2598732fd1f2cbe45a504bd7fbfdafeade3add7274a1e875aba3c +4e0abfc6444944b79f95b5009560818f7a0599e5bab4405378fadfe084f1 +653e5a0166714047e8bd4e4cb116596d8089bae9147ec1d62cd94491af75 +a1743d58bafa11b63b447c954a8d7fe11d39d969feac8fa93c614f97807d +ac62cb7a84a974a0fa555a2e3f0ef662706efcb828ef72e2ea83b29e212d +f89ffecabcb08dbb7119203c4c5db823bf4e8b698b763fbd4d21e57940d9 +1754959d21f3f649d856ac6615eac692ebcbac555f772eb6ba3cece5ebfb +cfcc2f3d8dcad7edc697df93aef762cd47cc3ba9e2cdd10940be676efe7a +a3749170edb47b7562805e3f8bd978b18057c9110ff8d19b466ea238af32 +993e2d3021745b238021f824d887d2e01a7ff12fc6f084b35292f4864579 +406c0f61d0ac7cdf7e4770b424e2ccc22353e6c82bf8ff172973df267ded +bdaabc2a742beea02e35b9b253f98de9ca131f802deee2905ca1a6dc4608 +19a59b4a4265c723007d0215fc8ac2a91ec5f86cd6aac1e370a297103c3a +3cff58c7ae201cbaaa8a12c93e95e73974f9abcd678451b1db02ebb2e10c +c5abfa573a2ea4219fd1851765649318bb556b728d432ec05a86e9894aad +9cdca63d08642655801bb37f28b6e11b958e8e800c8d521ca4aa045fe9ab +ac02dc015d18b1901d519181ef60227170a07f3328a6d5fe4c5aedb35fc1 +3dbe86564a9b1dd4c7ec648880360cdd1742ed4ac409450f1d9681cb5e46 +5edd1de2a2c7f8ed63436f98e849504ae71bb872683ae107ad5df3ca0b47 +a5b79513e02d7c540257d465ae4521cb3449d79c931e2ce8c5b0a0a4ac88 +cef7b9e5f92bf721ad51682d6b6f6c14747f78eaac1891fe29aed4eaf177 +e3d2fc655ae889c0c30a3575a76c52e95db2f6a4d8ffee9518391954b92d +39dae4e97c4022031f8ab390b66ada6dc9ab2de4d1dddf26ac4032981a69 +08f73d34b4849ae28832cddc0dcd116a47d9262b0f93c24fbfdf8a78e6ae +ae3357f3fb89530854257a9db773a1acf5271fc4ca04a06b46dbe661ca11 +9f45e0080cd129e1a7c23a33f1c48af960761b117d9d91fa5a0ed3e47865 +b774a322f7dddfda2960b91fa7ba20c8f9eb213251299ae328b28ef54b0f +55fd54f8047c555e4045cbd70964e1c953e471408e4f25fe8ca7009bfe44 +0244b1e30dff518ea7ce5078027baba4e07ecf0ebecb497b4bd88f1ff72e +b261f6dffec0ed895e237b5608d31ef479e8c9ae9003039a5fe67252ee39 +774e1501100c0fcf154f5c5c81c70539e03118ab91f4ce247f6132d46346 +bbbb126c09d7459c1977e6e367a0c83d14edf7dea081e5f795a7c831fd1b +325b33674ec9c2b68029a0e600746329ea2e1b9bdd5cb2b140468e53c108 +8e8f2567425443f8146ec37101fa4dfccb0e032fff6cdfd76382463551b1 +ae8ca6cbff0e34a3f75ad400a9573217f8cbb00a6d59ff46e48421e97091 +cb17f53f20ebeb89609ea55ed6ba4101f2f3ceccbc7ade21202439ef91d8 +a9a783c22de7e6601b50c4342e094d0eff223494489fa92150425da1b432 +908423fb3f41e0b115ec1ba592a4f920d15610b9fb33f9912aba67912d05 +1ee00a13282c1909a3a56c4ed06f2f4d1739dc296b7492aad0446f87a416 +c6db4d42b504dec3a6756f3d0845ab2d2e151aa5fde12b31a9c3b5ae1cc9 +d97192bc048f00dead66940004281c4d5a92c20b1f77795cb4f98b8eaa7c +be16f9b9d4a34a1a53e0a0deadb4fb4b20d9e8064d3412ea8d2ebd259b8f +2f04bf4bf11a5ab7883c99943d762549c3d5866bb6ed85a0e862eafbcfc7 +03bf4b77cecc0d65bce4df33e0d65456397f231f8cbf66672457cf539817 +6aa5292fae24695009e55904a04588659a3a23fa11989b925705ab45f954 +6f862b0e176fddf75b70d9ef7389f750becbffae25d58a1252cc04a79e13 +fbb6a666fd87cec5562c3e14fd78ad05be28ff3871d6fceff5aa8965bb65 +67ec76d105a6348e915b27767f5010011e80e0e2f9c34742a4eeba369e66 +8faf086a45ac9bcdd76c758db01a78602412a4244c759ece0b963d9ea58b +0efbf4376bf115288803a54cfcf78584c8af80da2a3324096463e3898285 +57de6c6354444b12a74d5e66053f6907c48522cae9e93bccdb4632131add +52eb374213888125de71994c31dba481b70b2e4c1f10b865d58ef09fc9dd +2ca7f69bd2855895256caa5dd6bf7d4d8b341d677c56ca08fd7ba37485b1 +444af8be0dcdb233a512088936ab4d7fc8c03139df396b7408747b142782 +d9406db0dcd31368d2f23ddef61b0da3c0704e9049ccf7f904548c3ca963 +76eadf1ccf77f94c157f5b84f74b0c43466134876a90c5fdc2c53af70c3f +f5c2d13cb665fed9016454bac1a629361c8ea62f4b2399233e8587db6e75 +a9cde3530f20a68ec155d275a4aa6f63aa5cd115244643b54911c954feca +d57be2a6c40f1bac38e393969617b066f7d94e8b18dd80fccd0168d4a385 +f2f1489d1dd41b68d47e5ec66ec568333d1f584e3dca90f1367a990630d0 +14355be7dc45378aa111c319838edd441f15e125f928e044640f25ffdcc5 +c116c3f6ce0d4d3195187b22200808366eca9b508ec45e664e562186efec +a97b22835d384758849605a01973cd9ffc1657b124950c9d9fa3e18b1a20 +7156c4f96f08b87824373c2865845d17a0dda71b1d69f5331c5676d0648b +ca80a7958a2aa034d7e1e9fafead9248e6e64f9ec327c60ae4f724e1fb95 +8a71e82ac3842768b27b506b5982311557432dc3f270ae6eab23a42fef70 +dd0d407a02cbadeb7b8b74a2523cf46a5f61e52b053c2007f75ae053a96d +e00646662d027d93f950e516cddff40501c76cd0d7cf76c66b7bcd1998d2 +7a19f52635c8e27511324aabbb641dd524d11d48a946937b7fa0d89a5dbc +4b582d921811b3fd84c2a432dacb67d684a77ac08845e078e2417c7d9e08 +bd555c5265024aeb55fef4579b46f8c5e79770432c5349d5a65a47ce9338 +e1b599328bb1dff2a838f732852f3debf4bb9b828f9274d03d7cf813b123 +687c5e78a26310d87870bfcb0a76bf32aa20e46f6b2826912e562f503aed +11e427b7765cd2a68da2ec0609259ff14f57c07963d075e96f8bd2eab9a0 +dc32714dd8905f2627c6d6f33563436bda2d7fa9a976f88947b84c72f454 +bf0b66ca84470375d2ff252b4a2df52ab613d0c8ef0465ff1d809ca82025 +c2122a8f44c56ebfa25690bf6a05675ebb8634ddfd24c3734fe8cb32d6d6 +c69c72a4951cb959175770b4286d383e7a3f158450945c8a2ccf7e54fb19 +aa8d2d98a07f0c55f834f2728d89f82a598269750115a02287c4d415cdaa +14e1d9e7032684002f90603c0108dd26b40fb569bb21cc63d0da7e9e1873 +9df0a9c85bc340d2b0940860d95571dc244628c59bab449f057e409e58ca +cc3369f4baa8e53c6765a55620e78341dae06e5cdf2fa5e5ba58634b29ee +ddfee7f78672e55f18a7debbc30862f278f83f4cc123ab591371f548fbf9 +bd24b3453b9b57051c2e67edff2104f3a05a9f0cb7efd81c1b1b0a2bbe95 +21854902526e5d4fa1b3be270811b972e8726623410cec7911c07f871428 +1caaead97c503714eaadb14ae5923f020093722df1b9d9c055d7d5f95af2 +a9fbc5ab6f6c2bd655f685534d7dc5fbb5ebded6ccdcf369bd83c644dc62 +84c2810495888e9d8f464a42228cdc231d5b561c6b210bc493fc1e7bfd66 +5a6c4055a6a629f571f4f05c15cb2104b4f9d0bd1b1f0ab8252da384eeae +f5fd5c663ad7a2c29f65a48a30ed8de196f9eb8ea314c6e86989298146a5 +589f76f12664c8d008228b33144679d16ff564453b5e4e9f813191b6c99e +2680e20a410949ac30691b1428a255b6185b7e3802e8511192e73c376f3d +eb807ad2727fbb4b27538b3213da0746231b1c1b595a958466155835c537 +e0df4a0ef272d4c3f7f2ef011daed38bc58bb0fd7458e48060db98971bd4 +b24bc7bd0de92573a1c7a80a5fa2b34fbe50271dabeb83aaa4235cb7f63d +6a6b399360df8b1235e4e9ab59698930044a98d5e083b5f5a5772309b390 +9e1ff2a252734b32fee3940f0e1ba61f54dd1d3f6ff0d57c9ae75a302d14 +b9dd9034279aaca80b6bd05c74bf3d968305a5046910871223a3ef8c77d8 +25d7e6d3d2809e76064c473d1cd7c05666040b6eba647e34588f49fd70a0 +3c937933a2272c938d2fd3aa8149f215bb48f3bb45090bcb9a6ace393a44 +f1a9bda2ad09a5f566b2e8887880afa45a603a63ffe7c188e3eae926a903 +4f1803368e773f42c7391dff1b9ce8599161515c549aca46aebae7db23ec +8f09db0e0f590aab75e8eb890df354b37cd886bdc230369783a4f22ab51e +0f623738681b0d3f0099c925b93bbb56411205d63f6c05647b3e460ab354 +1bf98c59f7f6c2ea8f29d8fe08df254d8a16aab686baf6856c4fed3ec96b +0328738183dbc1eebb2a3d301b0390ed8bd128bd8e7801c89941485c3c86 +22b5f223cb07dca74f0e8643240044e8c376abbd8c82ff98c6dba9b6d244 +5b6cf4189d63c6acd6e45f07485a0fa55eff370da7e71c26469740a68627 +a3c297d2bf215121fb67815b7b9403aecca10d21e59fabcbe38f5ca66e7b +551b22e28f2d1fd7303d15a42c45bf54b40ef7fc93060ae5164e54f91c55 +20bd303a98d0667a02a900813b260c0343021ac01872fd62cb6abebc7ad3 +a4456805159839ca4a3e35db586221169ded66f852e8974e3815d4d7659f +6a9bb93585aaf264f06cb6da6a26e51683945224158ea69719b8e4e36eb1 +01333aac974db8f84b051724cf245fe7a4c86582b5dbb9a5d9318180e33b +8d92c22c44b0d18f8ca34dfa4ee9693c1a26fedece01635fc5eac1fefa81 +32458254ad46dfdfd2be12a1e7f32f3728f286f1d5d4394424a073696b65 +e3c459aee9310752231fa703faf35e11796c4eeef698f4109ca8c46ee322 +5dc2e3e04fa787188e583321f8410b68b9624ff60679d3f25c13e5ea7506 +a3ce8d0bebb99d9a959ad92d8cf909988d9250b310629903d6bfcad4581a +504b91b2c91889987f36d6fd0be1d0ee5aac00aa0cb48d78a1f7a64a777f +089573ba79452efcc31c8258fb317369feb0d7ccd48cf13da6d1ccb59a4a +48ea0b398e590c1169113fed81639e13e96aa268d99cfdb7aee977fbe85f +f784853a06642b5521ae0a7f610c9739af31ba7a5157ebbbad999e23794a +d2cf25af987dc85dfa29639957cf28e7f2b7671188045130a6e2785f8d8e +30e91f0f68c1cc9f2de902952730003e816e4f5703db7a97b4c566f80547 +42fa77be563ef681a4513b9a68b2b0956551c74545cc9883428dfa72fd5c +4eee93256b26bc86ea34f7427cb0c0cc22c0cc343f739c6c0c46d0923675 +5e04d70587426ef875f8c89ff8492ea23e4e4d763b84a6437a440e69eb70 +65ab6d8cf5f8444a844e6ef3d158b451d121daea2d0e2b423eea24254226 +7eff1b4224c4e80af2a7becac1649e4bbef09f39415e9b1e3750d7ac47a1 +068a4f5ce30840b00574eb4e683e3ec25f6e690feeb0d354568efbc354ba +813ca1400734a67693af127b0f636d58b83e91548f98e3d87da7fd7cdebf +f3ecb4b9272d1c83d4980170378d32f1d98b87c440881af9ec052510982a +0c02ba6743bdc7691a44bae5e044c25304c1a2525cf2c0694494a2e9aa34 +f36af43ab288807ffa4bd418ad51d98c75f2b2f01abfd834d3305682b6b8 +62ef69d05962aac485bb4f560583a5dbb74e967eaf6d299160753ec32249 +bb1d9851d5441cb0c624208e69dc876cd8841a66976b5d7f9c99be68363b +8112d33d971f2c4f2a1feca88ba1a794ddb725c5e2e2c248082231059aef +729bb5fee5006ab8809f63e162fc0743c047c7984a9e6333b433fa143d73 +72d4a74fe37314508e04f54dc7a1445e2d6178ec9c041d0cd4fda5cae830 +4b16feb21f3222261c293a8b058dc708405c1a97ff34eee4ca69ff4e1ee2 +a03380d52297574e3aa50c8afb826fc94a14e8caa9ba89d6e92913be9e07 +bf7ae011e6bd142d8952d9c2304735e875d1ddcf82fa9fc0c6449df2acf0 +d5f6cff6d21ef6b2d29022ed79c4226c97f163284f2311cf34d5b0524a1a +a446645b9d05554f8b49075075f0734b3d1ea31410759c174fcc7305d2c1 +d7128781043cba326251a3375784a506cf32d6a11a4876f85ffa2606fbdf +27dd16d64b2108d808e33c409dd33f6e0c6079e47e7196016f261e824fba +b0e4f91a189747053e648ad2d942ece8f582f052668b63a23a2fae4c75a5 +180db7811aac654270ec6e341126e3561429f1d41fe7ba3f1de9f8bbb8d9 +fc5cebdef869376a2e42dcaa578c0807835e58d75c39f91a83d5c1eb86a1 +b0f7aab991f65eef030f212d38d10b1913bff71717c06c78d9a1be136f21 +4be157ba11ba309326c55c23ae8512646751fb82ae200c06bd2e644bed38 +c7cee826cb587ee8ff378b7fdc00ec316bd4a9c24e2c250cb3d64f8ecbb8 +7f4d81626d7f1e4491908bf17c48c84bb1736693eb4d0fe634484cdd590f +a40ae94d44f348ba683a43004b487f047745fcdfdee2e913328a11a99530 +9bd117e0e5be4fb25d176d59dc2b1842418141190ed9ae1f33e5354cacfd +a5e4bc186119e1461bcd98517e675276ddf0296d3b3cef617dfa36b4759c +944fd721e1bf63d45cea90b5817a40d153a2f779e03487cad3c1375425ac +8cbabf7f754d16cabe45c65f1be4441908e0969d5a5111c931e724537dea +7cd3fbfec9b2f7d3efa747bf586e9218c3106c49276b89fa28f770fa0644 +fe1f3fe3adf07f59c755a5b39a2ac1d6f23c256a293bf3b31b6b9cf4c622 +b188d6e7401c038657c78bfde9ba09f508f1bbe3ed79793772cfc928c4da +519f7dbf3ff7074284437d2de8d7b7c78829642d924abacf353119e9088d +14739935a23667c432806085c3af71ffb7c5fe6b4412b9b1044c1e62ee0a +a5ce7e0322bc65a8c7d874270d84136526e52d0c7f9f93199c6bb7301216 +a19bebcef3c5633f21d012b448d367157ad928e21f8e471e46982bc46a7f +df1bf816a86dc62657c4ebf286134b327ce363ab6a66634eaa2a42e99034 +069fe1302febf06959eab8e7304da4d94a83ac1650a02c38c1c4b7e65c43 +e3a6fb0213e57ac49e58721a4f36996069caedefeb48f1a59303459d5873 +f3bedcdb9d00c1cf31130c27b60928f210e1aa5e1c8e04b86d2049f31265 +9198fa646c53afa9058eb8ceb41bda65f415c79ac92af5790b176de1d300 +f1c06b782d584f458dbd07d32c427d894f84215a8e7819e295ee98d976d5 +644f11920ff2f49cb1075c3bb42b9fe4b561362902f11a75669b7e7c4475 +b65f1ae48834cd67816eb63b58cda2f50bc22eeb0cc965569b476bedded1 +2701668f609393659b266bb0e37bb27afc90bca271366e34754383363592 +0f9a3b508aabfe8deef585b07a992460c592a150b325b1e50e4214a2f483 +e9dfc826c54b488493a96eaa37276f5a9666f0a5388fe388263d2c0cf614 +c6cd01571da4389f01fcdbd0ade1c435d64c5921b5bf7dbebd5268100a03 +1e1abb8cbd83873089a9e08cf80276c7e30d2bb40280278c29fa818eb079 +87623b1cfe13e0b01e27be0a8320b69b5afee820f4705202158b7f3059b3 +655bc28a754d088fde23d43d6a9389da8bc1cf3e8ea1a6f4328c196e655e +42184444d8c0614c7167c91a492c24c8357794c61f5e47cdaf4b38004a5c +8fceaa8151e929328bce1b8f67b22034f3f75e4d105283337c3d460e7d99 +89920c43f5e1449c74ad6ab5ea029cc6e497ea60068451c4ef2132fb87ae +049077a156c868b768df4a4c475a532e2a22d999931c64f8bcc18f51d25f +0f94fbd3e9e6c094f78da062f80c4aa2b86fa572cc469e629deb4ba0c553 +55e8422b562ed2f694d0e8e5540144e30841d7593b255edd4a61dd345d5a +00e411d2c50d64782a3ebedf945fc31c00d2fe4ca800f5aeeaf12ab399db +956362e979bd7ef0787188e43835e5389ac444d13204af6bf1875622f175 +09f32015c28729cfa3b3cca90308eefaf260e3fd9df10f3e76786b8bc0eb +a30e8cd33689aabc55e3ce387cdb89a30573495852a48009cb58a0fd34bd +da911159ccacc94698ffb94c5f45f15ecc9e82365174cefbe746f95eee44 +7a33b4d823487e203478eeb2d8c4bc7b743427778249c56e48fe17d0a501 +7b693509ddfe1f42bdef97aedcc26ceffa9357dd985cdf2c70bbfc987354 +6f0aa7df227ec42f9ca2482f58809e3f9650444568c54d3520bd0a7301ef +48bfebef1fc4332b5ca851fd786c1ece136fe9e575b69393b5aec2611903 +fae6e7a5046e2ff350becb8700f209b1131044afd32fed1bc1297b6a2f29 +6ec3b87f170e92aabacc8867360e4dbce9ea29f0c1df981f6cecc8986767 +0ccfb4c9faeaad7ca9029b8ff0129fec4a040f80ead041b3bc8af7526675 +ed9e13204e64d76440a097d77c535d34165bfe9ffcade530abcc75ae224e +890d5c110004e218bd827a02ac7340e18bf3684c43e664e0a37d5fd4fd1c +4d4489d25a99d542c16e06685652cfa3567da4eb0cb517be1482939da0cd +d0ea3519ad1e51bd9dc7b9077375a8cd3b5de9888697e853bacddbbdd1a3 +0e442e1d6f2d652046821813d0cc0e8f16c97cdd32daf239f5b2b65ef620 +46f6e9821b2e2ec539302747795fa746318514d38bdf0d0e490c00e114d5 +03e7fc9a8fb83b14337a5bb4d640b52630f5450bb3bfcf7cecfbb1ef5192 +ae401265450db197bcfa07315ff95a809bc5fb4249e3a728a817f2580ae3 +50d8d6577f79c883ab4a3119d9ab98219aed0d1e826023a66da814396058 +d95e52d9af8bdbcb0454721f27855b686d13bdb473f650c9865f3e04f08d +b10f5256a3e59bcf16b12a84bb7ef3b370647cdad5929b722a05f5b3669e +14c232bb82fcb9c1dd8155ff4515f4e83c895cafb86754e896f38e5f3beb +5d29f1bd99cb8a09c5e50f412f6d8a773b79021ab2c4831aa663c5defc4d +553616874dd5bd8b75c7a2af7d029aab5a72528fbc4b5ee3d30d523412c9 +60b432434017c4cd68b2062d28f307fc287e11663511d1a6b52143afac0d +ce0f7ba3f326fb707fb8d2c985dd60090e6664f2344e098a7a1a6448026a +2ee651e8141cd7786b6543f512e4c31d25dcaf6652b1eb52706300b771cc +0c49295067befc044ea46341927123ad4b7d094784bda7fa7b568853d0b6 +1e4cc39e1abcc9479f91a2501009ae34ef7d5ff56205cf5288503591cc55 +c48abcc78daa4804549562afc713a4c11152e6e4331619b2e474a25ffb62 +7c46112fa4259f07871f8d6882e9a7ec62d20a86a0c502815d0a8f3f5ce7 +cb4a6a74b6db8e17d54bc919b82c7c729cc05b98855b9d8a0fabd8a9bdfd +4333f395607631f57c0473be0fb290c4f40a7aa6ac49208570ffa1d0f849 +d4871ebcf9ef6f5106301cf54ff8cc9918d6de74d519fccba58bb1c21543 +f3bca9f43c211b2e5c233ff6dff2c9b56d3f656f6070d13dfd0be04653e4 +98c670770e01c07b731ca0e2eb56e608828fedaf1a31087f2d43cb4c0074 +e576769b0830577c86ad5de48ee216df02d7c4e4ec231afd8e76c608fc9d +06cc86f38cf4d839e0a0829902f56cf2f86f08b975a6bdd0642d6b4c78e2 +57cf9a4f52646a952f6a220c36c91db7f44c7f44bddf33328ea8cc01827b +5f2d79e3ee6c514a4f8597a847ef5f32c6400736e6ade28faa7bc6e9c6ba +e4bbff236fa6dd2b0ed23fc77f92649feba149f82488260b0bea2a4fe1f4 +65d96d8c51719e5e10d4c17d1b67e700aac36b1ed55c93b4b2604e72f51e +b30fbf5b64c6fcaaef764639ebd789f82ed354712c7f9fcd1df257e14c0e +8fd59a0eddab684bb1b4176d79b22ad2605bf534e4b8fac2272fbdeaf210 +0424a2c5cc65f8dd5faa13313dd926128ed466046ee94bd3eb41f3ea5505 +5a70603a2ae1981bfae8e77d850fc5a5bf1bacb3df9b7cbce68ce7979fad +a73c2900526b68236c6d37197b0c521c5b1cf5cbbc89238586eceb99818e +aa47ca94ff615233575fe83d0d50d734351e0363030a12300f7b20450946 +17bb209c346ac1d35402b617d6260fce04ce8b3231ab5c05af30b0f3ccb3 +3616d3df334c8d963279537563222dfbb705c3e14616ad01927f952e6364 +4c4b7fa44ac97616c1521facd066aa33b2296dc03682eb6a3b9dd8e5bf62 +53f10667ecb07bbd50553f1b211067f5cf098b64b84d94ba9ad8b146dc9e +8e9be06bc14cfe0945e22fd819856d6996e857c0bb5f292defeb493589f4 +515700753885d61eee1b8c19e6e94fe2302c07933f949d6bf119d207fb04 +dae7bcff7578bf33d77e29611c7cf03b2df12c242827ec4c4e5b5343ca3e +4f7f38ed337583e30dedd78a082f41d60cbad55d59dbba11af1bd296ed6f +e31d2e10d3a8b5ea698e656ff97755a47ddd862d23309e2e6ed3e3e111c0 +2c3a713d782fe301dbaff0a4225f932576622d1cbae40d20f46958298d01 +783851c894f2712bfc4736d3802e548a704878e2d139348671fb96d0ddbb +f56d9349172caef0dfed4b84d867116d91063dcdf9ec401dfe8abb269ee6 +0d646bd12e0752313e2ddc272d9f4aeb9d940987596ab623f9198765cec4 +62f7b6c540c9a70c9a872bd28ea62e056560b61ec51fc68eafe008f20760 +246e06374ae5a6bd2577217700507978811ec29985ab644e474e41e8a105 +295fa67ae05e0739e8c7fbc51104522934942f53e1e1df1ec2a66f0a74b5 +9885cf2c2fad1cab3e2b609f126ac8b7350d5408a7df9ed5c27a10ef6505 +6f0d877cd7bb902977ba93e6e8520d2d018560ec8143876ad0dcb95b173d +af72c0d413bbb5541f14faa57eedb3ac2430e36911d2f486d9ebf9cb6745 +2ccc763e1e46e7a4b8373e06082176a6c66d045e18f90b4b2ad15802f6ef +cf2130cdc627601ecc19887784b6de7fb6a193bc3d057ace29f74199acae +69526ba6f7a2c669593f9d0849f12e37201c32c88384e4548a6718cbb2ab +714ccc917d93b865ac7d7d4dbd13979843f4f5c1f8b937ef12fcdc9aff50 +f09d2625f4367ee70a98772a273d8919952102aa03297e3cbcd876da5abd +2ceb162b8fe1d9a22ff694495528c09a8819fbfb6946ab205d4b2424f6d5 +6fa1c704065cb64fb2aa0fdf291fd5e7daa38667e6d8e889be7f4c453da0 +59c492cd25fcf4a03a6995897145273a66cd6ba999138bc8e2aa7d080f9d +231497ed28a9a27b6b0d4785bfaee46fee71b26d6839f2549a14e7ab7347 +0b6cf368d2d49e74c78d93477828e4582589cb447d795181d3f13dd8ad52 +3c750df8f19b3260c17a6598b406472a7204dd26c5988911ce9884de9a1d +ce33d834becb1dc80efb07f32d3ed6c2a484c5d53746071576c3f67f25ff +1558986fe2dc2265b4fff79c07e3f4c6c0ce8319e04c14728ed722cf214f +65066148bc817753dfdcc0950bf80dc515002e1a92e7d8936e9b3aa9635a +a6d512c68aebc79a62a6bd17a411bba7684e1f06be9bc3d1aca25d50c8bd +1d75597194cf87c9ffe04ff28bea91b5b9521fd356ed9e036466137586ee +f0a8795486438d0d9707cb2854f12963929edac394c562235ca71376d938 +e4e1518668180b857d75318bc22e9f0683749047e7649f9e20b35204b6ee +60c0d47bebf53179a083f0b4cad5b3327a3faf2cf03753e3e46c05773629 +7e9bb305f603369cbb568350b2b5c6d23a35c551e0ab28b082e321ef4ed0 +e2704d35c75b4750af782160c2f2e9aab0e14e541e95b64ebedd66db2c12 +a8935a60177cab634e20a8871a3a72f4b21c3a34d9dac37176a321c2ce3e +e828d140c8445117e7fe4738000c30ffae8e2a48bd618cc8813e38fa0f86 +92ca634d1e56010987483aa0f08980d91528df3d370ac724acb238e141ab +595dcb3da7a769de170edd5763078d1084e2ebefadf8a50a816b50722617 +c9539dbd68d9062b015639708dd900aecf4f15adb36339c05a9aec7403ed +771f9f28c60e52bda3ba6902e06334036c1dfd66d35ed00e3fc0bebf55da +416093b5cf512217c47f905ccc91fad879d63dd1380519a02025ddf15d70 +eaa1bd8cb6be67608fbc5c94796bd09ba35933f64c5e72a26db1ae40ef49 +af5e972fa44660588292b67ac670bf046cb1f5a7a0d73ffd6df862744786 +4a56393b0f1b4cfcfa362c74634713093161b29c94a2526b7138aa92fdde +b37a8c1f30a6b3837d9500b340515f0412e681f5bf36e7869fa157df18e5 +c79df3e6aca924d7b7dd2e0d5b87682d7ea6913b26397ac180fb75fabc1b +8e156ed542b9d8c83079bccd141c187f90d72694de4f6d08520d11cd454b +bd3c2e6d259694fda0c8decc724bdd650163b7f6ce1181590c06de4c0dd8 +536aba318cabf54782c919e07c2ffa1034143175d05deddfcd7dce6c86a9 +ec9bf6a4437da474aac2dbce2c91aedc20043f179d5c9120f3dfb1cf6906 +c27f2ec68cd75035c283e1672ea90d953a23a1515c420b81c3270fa06573 +4d003eca1bb71a2dacdab67e44f47c266c2ea1776648b62bc110671e6eca +4546d3c72c8acd956e10452c32532ed51bf3d0518467fa829efd9c896e8e +1e5c7ff6da0b51e872e403470affc95f25e1d2b9b59ddb0472705e14fdc8 +fc2af16527188508be10d098372cd7eb7d62a85c8d8dd1d0f55ae3ccd0a6 +5dd6bf776dc187bf4de409d5db3fcc5a6d852848a251f4fb4e01dac5e9b9 +587fa8c46ce03689709008b34dfb3dc105def80a1b515abcbe06e73fdf7e +7136e40cc922fe9a9da1726747e84427f288d934747b6c587490734906b8 +a91144ac82a57957cffab561714e1ff5148a39499dfc8cc96bf5d87ced17 +825e8f80cd943d9a73945fb8bc51cf1f9cb39c605491c1bb8f1c4139974a +59471ead310d041b1ca1ecd5e9f92007cd8243cb3fb1ec5256444699a9fc +ed6cb31eaf0912c16fa480a1cb4a8f4a9cb6a4d9a9903d1e2f674286032b +489b8a23ac4719fe435a9fa2d79abdbaba740e69d5ed611421b1aefcd06a +362ddbb7b79aac41e3e90657afc0b87a6e8c57ceef70a628efe19f568634 +50f47b5c6d95870039caa3d07a54e58df064bb5f59dbe9b9a2c7c84d7e0f +32386309560a0efa2cbfa27f861b208b2df4a062ffe2c59c057296aaf5c2 +0f48ffc9ff0692f8cfbd6fc6ed1f3a14537ba40d7267e6b5f69c997a949b +26577a9a99db3f53167355c4967dabd522292ddaca3c537bcf303ce76add +eb99f6664227a94d6a698dd5a5d40008349376067d057e28e55972264502 +e035b1f5e33d7b3aeae016f9be50f2aa09aa138d15d7af3c1ccb805f2d5b +cd4e9b2b5c288b2af4a25abf0a9093749377c9e8232ba1af17962f85064a +23b0a13f11acbb471cc700f9f1b588f72cb63d3d1a95a93502ef74ed212a +c452f1a84619bbdf61a1dc79c0d9ba29c7f19b400f682cf66f7705849314 +f5c8bbf973f2c53bdb060932156bf2c9cd8d36cf6271075500b0e3e6ad49 +958af46a9dc950f4c29f1ab5dc0a85924f7ffef259f778459c80118b1eb1 +ed29208d1145b21b19d62f755de4972c57a09b3decb0a8096ab025fe6b9d +be49ae35394f0ea40d3693980f97f712b27f0e28d8a549acbf1da63518d0 +374941effacf63ac3de0523cfac0dcaeb690de5836741fe58917c7ecffc1 +95e7b560a3e763aa70fc883751bd60ea0a0f893d8e9fe75a66c67e202c24 +84f66708ae74413c0101fe0b5003be20881345d917203b582a247e6c74a8 +1d0479f317aba7b9dbbc0a92e91c51fbe8775a44c57699acc9da84ad60fb +9629929d1edabbd70b4ef9887ce4ec2469f154fada42de54240cf3302364 +7c492ba17e6936a4d85e0751df0945463368a803fb40d8ded22abe118250 +86cfff1878abe5b100bc08b991cda6fdfd579332360f0c3374842edce6ed +e43649d6702f34668a29bf387e647f96d78f33395e8d4b3521cb4fb0956d +12c924c16eee798cde68e319a358cc3524c753177d976d4e14a2e0cb72a4 +80cd87bfb842060b1266568af298bbec58a717c577be73ad808e004348f1 +6aead32a3d57457376ab57197534d6e469ed24474a83618f3ce21df515a1 +22918f4b62c642de0c8a62315ebe02bcfc529c5b8f7c127085c2d819e29a +f44be20fa077ee01a8d427bbe3d97a9d2bafd77f17835279bf135900aee5 +9bc49582b18d468bf93e47ce0bdd627775264ebe9e4172839a444f928580 +8c95895b7e23592b2dcd41ee82e966c26aa2143e3057161511796e980998 +1f2e4ef5868b3bf4576e3546e6407e35cdf14654bcefa7557d09407545a2 +38173080b4771ea52054736677a8d9749a2b22b46b24fbff93c55aa2274b +8c7ddbd751bcaf1df00ccbe1f24a80622aff192fd6db2238db941ec44ae0 +dd73f6b2f80d89bd0aa30c038583deba14913d38a7b61b54522755e251b2 +aeca62033a39ec1143b2b960f9cb87f748428bec3243b8164f07d5ff72eb +f2ef69347bb933241c2401a96ba5ffa3f9ad060c41f4e6bf7280af65293a +bbae49d723dbc4be61d7e13f7a5931a697e7f2c6582dff416341ccf5a24e +9a53686a1e13bbe0bb480c19a4e72a5e477bd29f39dce1a17f63f1e8c696 +d5f8855cefdbf7ce681c7d6ac46798ca9bbdc01f9ad78ce26011ee4b0a55 +786bb41995e509058610650d4858836fcedfe72b42e1d8ba4d607e7ddbbe +3b0222919c85de3cd428fed182f37f0d38e254378c56358e258f8e336126 +9b1f1acd7f387686e8022326a6bbc1511ed3684e2d2fc9b4e53e83e127e7 +84da13550e593bbad1c87493f27b60240852e7fa24392fbf3f478f411047 +3f00a8fdb6dcb8aae629dc7f055d85341d119f7f6951ae612ffa7df82111 +d1ca48306a57a922cf4c3106f0b5e87efba6815f6de4294c7a0394087067 +677889d22a3fd86b0796200300d2716445078027fe0c0b05c86ac80d2095 +ae874324ee6ea3553bcb92fc1522a6d1524f6fa22b71598fbce784a10b5b +61e50307ef4409ffb7b38f27800f2185140ed08fc4ab396050b068025a9d +e4bddcad201e72ed9b41c4ffd4cee743c9c2345b95c5071442defc8ba5fa +9c63c56e209df41d10d93135a8080f7cccacf67e0b0ddb3e0a31df32b83f +290b3c536e9949973cdc80aa5c8a4feee20290a95f68e59f54050192de42 +f27464ee374e4d2451ee8708933b970402c90ca3070843a449d7c3146347 +1efa666a60fd5cbf55a47e4a3c5c318fc1af944d58d32690a2c7eeef09b2 +d94721896e1e3e76e44a8efd524ed5d6f5eb9da093d277441546c6828745 +ad71b6c13f653dd631bc6fc55d0eb4648b7bd9c0eddb13222542f2b6e8d8 +b80bfab4365f4199a41ac690979285d917de79359a183e6fc254b63e6408 +6d33e3c029f472f40742a99f92999f302f79994ffd615f1a848194cb56c7 +12146850f5e400303bf5bcd4e5fdccd1fe2edf5352d525cb15d8327f45a2 +6e3ac276dc8780c65724d28dc6bf9c7c985840070c35e32859168890d599 +a884dc2a90194cc2e9cc6a20c6c0ee11b20adf3aff01db48eb8dba7b0c81 +7fc10cf5a66e8171a2823a4cd22f0e80c82011ae56dd895ae2d3ebe84ff3 +d521c31453e0909cb9b1cf0b030eb6b7059ec38038cae12d0e1cc4b5b3bf +e6c821faac9b8792441e2612aa1ee9318b71f9966d7d3a64abe349be68b1 +744de7b212f6be73a0e1eb2fa30850acc3d9562f989cb2d4fbfbcd5d3ef7 +ba55717da1cabf197b06ee4d8650e968518b6103fbe68fcd5aab70bdd21d +66f09f96208db67c1b345672486657295a39a7fd689b2c9216c6b46a29dd +1283bdba295dfa839a45b86c14f553ff903a6f7a962f035ce90c241f7cde +13bab01d8b94d89abdf5288288a5b32879f0532148c188d42233613b7a1a +7f68e98e63b44af842b924167da2ab0cab8c470a1696a92a19e190a8e84b +1d307b824506e72e68377107166c9c6b6dc0eed258e71e2c6c7d3e63d921 +39690865d3f347c95070cd9691a025825421be84bd571802c85e2c83ba53 +841223435a9ced5dead103b470a4c6ae9efcc8b53331c61d0e1e6d3246cd +aa1b0da347685121196a07e97d21b10ad34e7031d95c1bafa37b4141bf33 +a6be401129dcd64086885f4b5f1b25bce75a4cc8be60af35479509e64044 +d49c8a0c286e4158a5f346ef5fe93a6d4b0a9372233c7434a7a6f9e7ea21 +30c0b4b9f62e3a74cc5d2916ebdaa51a1ef81fceb6cf221e70002a8a3106 +bfbccc2d1809dde18e9607fcaac008fabb72e8c50244507f4013c5a268a3 +6135ead9cc25362c37aa9511589f18d812e6039490f9c599f44e88754ac1 +4f6c1841d570efde27958c7f1b2c68772584e1d12fea252e3a6ec3b051a7 +6faebbf6f5101978e24a9ca927c02065e8e49150a55c64dd30757e8a33d5 +2a788437a9181efb47414dbc22fdeda203d4122137bd045611f68314e12d +1d6a5ec270c8919562c03e3af7b0e0deceeddbdaf3eab8fb5632e44dc1e8 +d46e2396b0236a46659164e33709415e7b347f7f7b87a9224a189ddf5178 +2cf66c9d385470a51efc88696176f6d3ac3b7b95fa074c981194e22981f5 +1d925f980393b7102f1f836b12855149ef1a20d2949371ddba037b53a389 +7617c257bbdfcd74bc51c2b40f8addfe1b5f8bc45aa4d953c0d1d5f4091c +6af796af6513c820499969593bfd22f8c6dcde1d2ee2c0ceebb5bd6a1ce4 +5fa61094e932b380cee381f4485e39b4b1797f2a7d8d90bcbf89b9cb1006 +2d50fff083743bf318157caac1c0179c87c03a2857fc002979e7cc97feda +966b09ceb761d3f55cf07637256c6aa8b8e5cb6aa9739452a330afbe7082 +975ee39fad5e8106e8ee05771157e92d99003533d922ccc37add065b6236 +7613d039741f99edc77c230fe8d1baba720a185186662376b947bbe1a686 +4b42c61ebe1abd40d890751ab8945c629de3b6d2a49809dc693f9e397097 +cf1e568c258081242460af2de0ca44b7ba2734573967b3bdec0e5e64598c +cbf41e630d821491504f414d9b54a3100dd5105a141cf61bd3ec41b67368 +c8cd366c543754ee800ffee3d19c9cd0d408cc772da10e4d8134964b0a61 +232e2dfbeacd0fdee12792504bb327a2e1fc44127f8577ca51d380a760b3 +740e6be46455cbf3917b90f0dfeadaa25d5d9f66cda43ebf9f75e0191a06 +25ba29666bbe8678822a453d4e876bad4a6b0d4b6cf98feb60339c9eba2a +dce4ef7faba428422c503d0210dcf8d884ca9f5094aab9f3b1a2238b569f +444748902907cb0d9d7ca33fccdd0cd29bc68e44f7bca5092be6272bc949 +baae5af92c302bb21f91b6ea8463265680f7c16f45d8ff35392a10eab87e +296f3af4478032b5b021db8510deb617941130d45c46fb3647d94b162fe2 +2738766fb6d76a06ab6803818b27c5ff4205ba668f95b5ec5ce4ce6da545 +c13ff56f417a4e0b3b8554a1e2a985a167e168adc8c4db28a601a80ab451 +91bf32acfd8d25c39c2f17fb3bca1296d3d160f25b43b4d6b94f20ffe012 +b779339b12860dfc897b366e3d400e756f4f9f4d2c86fb9d94c11ebd1450 +eaf720056e2c39529331bdcb104d113b42c94af2c6a5035750b7ae7fdcba +b6116d74bc07a11d4357ecf73d99221dad5cba4a7136425c2a3ac0e092fd +606a4ab722195e3b7fdfb5a5e3ccbb85fc701c42bec43b54e964dff3fa04 +193043eead7681cedae9cce6919949ea60ef5630c4b9263c8f98b4bc74a1 +63ccf3d0a0bc1deff39b800ac90bd734dda7ecdc73169ad77e129887db80 +7a253f8807a422eda8a16c9ee9bb8fc0942634bfe035dac9f7e36d09844e +39477c043399db4d07b3617da9d6eee76d0fde9201da98b906050748b68d +8c944ace3c96e90a3c2b63eae27b9152cb7274fa336866d71b65a57f1bc2 +bb1f482a67f3993dcb3ff24abb0223f9a026c81b2b33127a1dad8929dec7 +5d46bdd790eb1addd771c5c3965a2f514d3a128117a44560cc10a729bade +4e6c86de7c09a39602235c803902e34f5c176b18e127d71a011dd9a3a61e +ebfaa4a4e2a5651be6f4067e5e09bb4f3514d67c2129e4d3ea9568661138 +1e45af07bd84f883c70577a986416747f3bd8d1bf86d3d7b07e8a350899d +3c2dae237bd5ece45faba7a0ba30fcda7b7eec9fbeaa5a94620686d1e403 +1cd2512e8d89451c7bd8eb432c8862023d66f3f9fcec0d47598e2df59525 +d673a5ff493d458748cd6341f161a0a3e8996ca5b496508578fe4f653924 +2ae28bf4b7397c02b726fd5f9d8b898938bb668a546be6e42865f4f030d9 +5faa289eb24f7b8e249b224a95a2245605d67417a489626df7417855b8d3 +1c0043cadd2b461d32e1b39ccf409757c37b68f84e752bde6b5bbb847bf1 +57ea3434802def983d6ce5ceb3e9fbc4911b5484e99bb94dc3f383e50672 +0e85a91ed378e352838cf02921ee0ea94be01b5a60f9b1f58fcc1b4f527e +43725de9b9dadc3ef462fa279bd7138095d4cff2a0563039f71e383430dc +f628dc9611b2e3db08fb2da1d5383dc1a3c784e1e64541fde1d9d7f42505 +de96d3d0a401099fc2879af0293b0eeb143b78cc221f670c0479bc150047 +0cacb9a282e334e428b527acdfbfc56e6aec8d4d60745c1dc000011b6248 +d9ab4a17dca7cc74e17d33c0641710b02cb1edb0addc6be214b17e9f845b +2d9c8bf03c19e131e00f91f2a393b5f2ae7c3d4ae9021c4d7891d84d5067 +377ce92836e42eacd7e540824f7ac95360ce116d41d17a50748748971c82 +27f089a22ee0d21940de854f737547b73c7517addd9bdaab425a6c2908f6 +87dd990d6cba4d84308bdd4c4435a6480ecfa1a14daabd4d8e2398178e48 +de28b84f7ce4b61d2e6e64fe043c29a941f6de7621ee6f6d8b506221df05 +db238b8fe4323cb5f259d4d3d9c94d4ae1ca37d6c34345489c0284171346 +e9830e2e3c6c167238a7ffe0989d3eac870cd44102cae139469b9d909b5a +9c34792f693ac94ecd35d2277080e30a2d24b50391b6f2a3d3b6c81f7ed1 +a7b218903e7fed7a63269e27d793a2e0b40320ebf447c71f36d40dee002d +7257f43c8add31edf2c571123e46fdb413e007cc89e99b6f98d77ab38bff +cf140f787e45ffb2c7cc4ddbb59a4e32dfc36e2875f204ac851d757c1236 +12deb31324ea4c201d27fdab46e9f3988ad2bcfb8e9cfa8c487831a9b0c6 +60b20fb66b4c77f52359ac96f3b3d189aa0571c1c53db06ddb10f08882db +0b1e93e9478d4c75626c5fbdbc6044c4d82684b310ab2af144d12bf36f1a +c0bf6249d1da9ab319453594cb19d0e93c4e047fb49229c0cce76d0cece4 +2e76fabd2425382afe707db032cf617b046a59a2fc1bb3838d98fd5c8053 +ecb918bc14762e4ca45027623988f434ff4cb08bc9bff5d7de21940e3e03 +1ee042d9c30662aa76f96213fb5a92047af60f320e4660eadd1ec19d0086 +072f2202af5f219725f81882f10d1e065a8035a9946d0ca0e48a5e7dcf61 +0283b834eda01e7d94b3453830daade2aa6c947989b290c02ade0d7b2620 +813ad177ed82813b6a985d5c0a2d42419bda763d409da085936e33c817ae +68e5467eddc30be172de855a0f7f5c527555b3f4d942401b450f08273b1e +c5b5352fdb8562a71f276284cf7c27537e628f94bcbffe8d669ea2645752 +60830f1e65e83a2204cec393f6d92d4f61f317471b4b93039d298ca2cc94 +eeada0140823a2bcd1573e732e7b4bde7368f2ecca5961ad547f554ae989 +98d87b7e5d07a85c382bcea1693a697224f41eb8b406bc6a0c3eddfe8b5c +f25b11c3e4bd91ea7d6274cd6b3ee7b8f18cc3fd502a324c645568dce9e0 +d43caa61f7306fd5488fcfc439d85f8160ebf0ac90fc541f9c74d35d7833 +09309807a639477bb038200738342e50136dc64baa7cc1b879c61f7e1b90 +e1f2bd4f6e54c4dc97b8e4adeb102979203a31fe26a7f58c609915a95abc +4acc263179423f8ab16b04272d5592fc536f29a45cbcdbe15890f119ca9f +c7a52eef41dfa5c4fed087eef8e698ba738e300bd58f2a1a10da1198c1f9 +b60e2032f8384a86aa84027df21cb87977528e3bb9bea1e3a6879c56402e +a29063afc6ac0194f4944433f9a5872cf0a2a741382d7f3c0ca7817d5d7c +4b8bf53af0f18b1eb54480519cebb61d983157e039b13025e7980eb36f54 +3451bbb84e470ffd0f98eba80c74f238729dd6278294388a2e06de68a719 +47b6d478c85f124d14aaa835620e49b7f5a4f21347302c0f0864f7ebaeec +d0831c36187cbe9c848736764a31056d2cef27c07cca00033dcddca9a2f3 +b9ebf28e67257b69cd38bc23c711b6a2f6e4dda9bf5a19da275e6a8d683c +723bfbb95a90a344a6f421f0b67ae84c74652288b0597e4c86c28f73808a +77455f2948e8df634c2d14f221626b019033f9230c9167982cca9ae6dc37 +aecbcb49fd9fc1dbf2d11bba7187888721bc42a7f47c23e07d2fc5a7a91c +0dfe255a7f9d17e69af1618502a6b90b1dd748c7eaca1e1ebe8b861b04ff +e5f628f47eb4e7e65311037d7a5713d7cc3552dc85f452ba74c4f12aecd0 +d72892c940c3325640d62fe3bbbc71361dce6d54766e1fb99dedcb2d19d2 +fa6fa21f9116e03952ebbef659816a62db51a9b5b3916ff818518774ccd6 +79d44100d7236f211f36fa80a4cbafb3db76ba1e7e7f12082b0140eed2cb +5e793e24501715c6c170ad4f856a4bf16bb10210025156e635264d3cf18b +1fc1e8cd2fcfdc2ab1a24af9087975bfcf6fb703fb36e288e58d0d2ffc98 +bb4318001d931ad6161dcdf8984e6690e0f6bb07af81bf07445f8f57b355 +6b960d24e7cd152708489e4d953ab6a155a757e002ead97585e6c5333d7e +5aaab2731f047f3490432e0ebf3d0d628eefa8c1f665b9c86aabb0706639 +5bc372e16378f0d9b439c98e7bf87be73e934995d58e4e70d3ae9a5b54c8 +87a19f2826a772c39d41805c642354d9bec75b065f148f7c1e435dabbeaf +e4a5744e3f2894a928121ab069bffa3218a106a9dbb83971353a7c7a5616 +d9da66fbb908173f9b07aadcbd4d112cc353e7b70476046ce5a92e86eaff +4eec40acc840005f51f55c9f5874216851e9cf3fa431d95d3032e779e356 +4bdce33966a3a798b170a06c4cc9f73700224c858c36bbf2d0326c337ce9 +46f69c19a84187fa50afc5b36010f9a7612e3a25e846d49bb907af9505e7 +d8c78748d7dcb501bbb3d6603e829deee3784f2f3ca583d3738d6d2ecfb8 +eaa887103606211a3c1b5cd74a3e0e96fb57da91baebaecd3669661e7b1d +579ba41928a40a7028acff6cd409e601d23ff66ff2c8acb12e535360d727 +60d2e988d801930e0e9443d60dcb9f378fa75d58d73e6a3b6e5b26407c82 +67d50ad97787f8a9b91765e41552283cb67e43e59bf71cf08b9755c8ce47 +0cf374832c72d1e9702b55bcfc8b5a4e966d5072fb2a72a2108574c58601 +03082ac8c4bba3e7eeb34d6b13181365a0fbd4e0aa25ffded22008d76f67 +d44c3e29741961dbe7cbaae1622a9d2c8bca23056d2a609581d5b5e3d697 +08d7e369b48b08fa69660e0ce3157c24f8d6e59bf2f564ce495d0fca4741 +c3a58ec9f924986399480ee547ad1853288e994940bd1d0a2d2519797bf2 +8f345e1bb9cbf6997dae764e69c64534e7f9dd98f86b5710ff8b500e1c4d +f509da50c64e213ebdf91978553a5d90908eb554f09b8fc2748c9c405903 +e7bfbf0ea7e84254fb6735f09bf865244238e5fed85336c995bc3a3b9948 +947a6eb95db4cd1b64c0fccf82d247a2202e9e7eef5a550557625a0192bc +8bcc9e461e52833f6b8729ccd957d5c4b6e07016e864fc02b792c7400ace +d0a8f43c755f87bba6e5c6e1022416e5454cb34a19865d951f7aea527760 +53658cbf306ead832244f3062c39a0a121a1157a8e47008163c5bfc88197 +be16e9a1ba26a035a16dd38cc28dffb666dd4ba7356c66b7bced9e26e905 +4ce25f6d36607d8f5dda1e21ac96a815bb2989f01130ba1aca9aade554fe +effdfef5d6b0d2a01aad92f599f6a12e121010ae6acc6f150f19e7305271 +97da761b07530ca19b84b119e5edca1fad18462143b8913d6b3f6864b713 +7a93bb9e1bc29c09d660704e8d8292c61072ebfe35c354a2342b2458a353 +31d043874380d439388e46688a53bcfe01bc190ef1a6b5dec9d40aafe822 +261b28bf3e2d76f3dc4302506ce3387b4aa2a51cd4ba1faa2ed1fd7df664 +6772fe9f83d253451eeb0448b444b8ca80cc7cb653c2d1eaa0de6f2b1c72 +47e6d24ae72e620e200aff83a557a1aa7a0ce0a9cfbbeae03c31d8cbf1d8 +20b53b688ed2ffbd83418d743ee31e3d62216ac7be6c12bc1917548cf670 +d69fd2e78d9f7786ada0ea30a6f6d9fbd1f1406337151ffa1d3d40afbe03 +728fd1aa2fa8a4f075796b9de9586b71218b4356fb52daa01d3c18cb75ae +d4d33fc809dcb6e3dcf7aee408a0cef21353d76ed480bf522fdfe86e0e0a +b7d097defcb793057f0ce98ea4989a9b6787b14029a4bf10315a2557149a +fe9c91e7d825f7518b343fb556f0177a8f6ca08fbda9913d52997511590e +b9942c9813b4cf4d4aae4919401f2fc11fef0620eb5c40532cdb22d5fad6 +919a3a710de6c40d54993b5386636499c866938e33bc703a99c73adc228d +95cac73ff4f4a275c04d0d787b62c6a184dacc4024d23f593e7721be232e +9882fb738160e52ab905f0ce2c76ae6ff2c8bbe118a1acdb3b464178cf01 +94bc6a50df1090e9221be11e49f254b06c3236a31569b947ad041d1c6b55 +bfdec3c18c791ace0fe2a59504eef64a4eec4b5c8dd38b092745e0d5ad29 +276bf02c419c546627672a5764a4904635bff86fd0781d36fbdf13485229 +71f355de2b0ad250052f50ad70f61afc870ac7a816561d3232b73360d4ab +2727b2fd045f254c782bb3f1f49d94c6d625047071b7e32da5c6d21a86de +9283fd632074430772bfbd85e0c9ccab1dec16bbc049c3e223bec1b65c8a +9e98cf58b30a74f74f1a842dc91e30c023498e280ac55edd58f4cc731d81 +e443d9b9efdf5fea63c9f357320e01b8740eedaeef2495cd02eb2f338b3e +674fb074cc497d7b1937b188da857c2c230e9a931cbc00c85a7a36fa80b4 +56588e1bbabbe4ef429a6aef9bd4eb89c5752421bd049aa13f4dcf9b51ce +2503e90bc118fac78a25d187353d6f5d496cd6130b337666f49619cea985 +dfbeb7e49c67c1e0f0f8e9ec8ba14624ed0982dcbb69415e4b3c8ddba140 +397eb1fc1ddd36c94c374f018873ba41109e45afa51f0e691157d5958c06 +26fbc0903ae25e47ee372389cf65472a3e4d9769550bdc42c0b72f9a297c +d5d3c16ec67e06036e740ab664abc9f10b9499269b73ad3678daf4474329 +c2c7252c1f0df1e3b5e8f198dfef8325cb1e7e8057897a3d7fb5bb5858e0 +cfc0c115bbd7362d8e8ee41862af6eeda681cabbb06f72ebd2ae0b0be45b +a9e1be83f1da30687a655e5d148fcc17d9f53b760810a565f6d2f4cd5da3 +5434116edef756adb4d3df544a1de593be988f2bb8d36c34deaac7d9dc15 +cba49764f1e03aa09fe21fcd7c74e3d6487ebe219569e019f10dd163046b +c1a3cb2bcbaa8558197cb2c18709a998b4efa8ab8c9a71d2ccf942c17662 +1b88dee6b424165d6ce10ac48375e760983818e0085276b1674dd41042e1 +a01a8de111c903f74834199b3230bd475d92c6226ef74eb1daaec3475a6a +fcb47644a17c7e390ee3b16bef1c1ca6c55eddc44fbefbdde525921b3047 +0d76817bd8ac724739a8e743eb09cf78e88adad527d4f115b8a32ed4898f +45bab3eb802b8168aec061e3ecdb026c056fb9efe7e2df48bd516ccb12ce +00de08ed8be4ee0c41f40f4c8f64483e0ade90a78d6d4fe9203fe0b97c60 +3b2f8882bc15a212453c691c52d00fae8a3a26934ff8acf68d4352eef75a +0b10d938e55b7333dda2db0296a69e9775bf82b1aa6d684fd9080fc1c11f +ab4369c7a95a9504063db900a6e345bf6dd99be041230b2e60cc86b8c345 +1d84a9c2cb4ab6d74d63dd43dc26eb6b384f5222796d4083dcc3e1651548 +d9469f09a33b213a33ac52a6a2e23802d8f8a75c01a607940daab0051410 +73a88130bc192f303616adb113c0051b65e12086cb319c0a5323fa7def40 +402f5f87a3b2c2cf0e92789985f6775ac2743e1ffe2d0668291059740d45 +43bae7a2897e5e658592bf5a72966097742e0702deecb0cb12499eab701d +34ba37a08346217a415e44297a181bbf3744f0a49230ad6f030e11462be9 +afc2ae14e0587bc02311b48b8e2122c28cdf14414f3680fa52dbbb63b17f +6ebe4a1204f3c5d6150cbf89a8023890383153838d4dde77d4c8b1b78823 +8918c564d3babfe58eeb154307dd1997f5ab7105426e35c279008b2677e4 +695c60f956b348799c04b734338018fc27f7de7ad9d73468fdbc5283bd14 +c066ddad9a3562f16baae15d72d7bfcb409e1c874e9db1a8cde233b282b9 +6e76e9c08d85ddfbd3cce7e64104d0b0e95291bd91f405ff82f41601ee20 +8471e613fbbee67f269e4e954c36d1d18ca9880b7cc2b08fc990978efdc5 +1d157deefedaa765c1e26ee125d4a2514a41a3b95e9151a824532d7d6486 +35ad622718fe71219a697e94c2e64f26424cbb767acdef5cda70e179cd29 +b7e318d1c6d3ad26fd5fdcbf2fc221301cc1f10f5ed86b40a1a6bcc01c90 +eafd65183e75609610637b99fea57885efe76437df02a2ffc21223d039b5 +74955d9a54ff41980eddaa8768c5ad883a0c9150877392b990d63c6805db +7b8d6ab1358cbedaedb6feadb0ee4fb8f9c1ca03a3e755a74227a8930bb7 +2ea0a00b48fc626fa14d7d48624aedc31c556f44e982f3ccbde7ee735f73 +629ab1b65bcbcf0a3586a920477e8c960219802fcb1bc3a179032b324f8d +c424899b38275886cb5bc771f26a0880767d49cc23426a40a4b6ff8fe48f +d747565fc537565f6d7fd08706accc60f5fbcb45bc785f45ee9b0812366f +ae71b23ec43f3549c8224d78baf18719f05108d5741e681457ead8abc050 +462481771a8dc6cfeb98956e163981a98c59ab44d90e9c3a946c453b5071 +db0c769f7fb5144c7ab0c9ef1a6db1addcde1d4ae1daee1b4035af256a04 +df53926c7a2dcdb94caaf12f986e20929ba4e396f3aa7c93a7abaef1294f +5f13a0dd3c3aaa8fb38da3e15daa32163b7437af683b4f5e64cb14aebbde +8c69ed2e8cdbfb213fc8129af29ca2c06c8f85a5038d688d1fa5d1b54ebe +4dea81a49ce24131f8e6702e7aa4e2cba078d5dd373f894ccb275f49c690 +1dc772e1d2f5fb3fe15dbfffac62c87110162074eb72ae4e5e446bf7e650 +a554178d0d64d3c07f330f0d99e99f2239cb1597f2e5f443854cdb0f5fab +b28fe62f22e7f3419d017980f325351bb04f8f3c3dc57fee03cc029bd29b +202308d5a800ed2d500d41ace8e54e2557bf25b627883beb8118d800eb94 +f4253f855168f7fc8a2d29c5fcb76bb90a6c4e345722b8991a854047f46e +4e97336be85470b6be2b9ba573dbc4967ddcdbfc3b6fc35b0c7f3f2f570c +55dc3fee6d80bc6f46cc7e4d86a0b86f6fa61d062e213d9e442db63fbf11 +d03165b44572096995ed342893bb672f6bb55ff8fed944667995f0f89a48 +a904c47420f32afd14129c6e2bedffce1f07ea69d550b6909bb5beb4aa08 +b0b44f35e018ba5206fdb4df0228462c1fdbb95a429e53eb27bb1b0490db +f07202c3608d0f4ce08570e3d6aa3d4581c569b57bd8c1ea0e4ed3fc5497 +e316ecec06e6be582d9170d426f6d22d8c7287b8219945c124941ca8812b +e97efd9105eb6999edc0665016633b3b48820df736125b7c76c9f3a67d93 +8a2a0a6b743fd42aebc46a0249be459f16811ac9eba7b63bad7c2e88f175 +0eff8da5faaab5659824f9d19b3225aad2ac17c52c523414d3031d08a926 +30abf474fe02a32b44d3b7d9fe0c19aec16ca6d018b71d9d395ffaea0788 +0d4501d7cdf0f7077a2d63303d09083080d67f1f714a1b271dab9fc9866e +4b0571a171eec8a4e351ba2d02438cd108a33b1106acaad0ccdb051061ea +7f40543748115f29debfb4be4b42cae8762d62114ec6f8ef68c478a8e05d +ecfa18b0368428efec9eafb2353f95e3d71e1636b9d9f94a77e692843255 +698576dce13b2b858d2d15ee47cdba3ed08d64b77ab46dd29bba6aac2106 +ab847de378cccdaf35c64e50840248915f4fc110992c493cb1b9cd0b483f +0f1abf5e9b018210b477fea28234ffbe5e0bbe01338e0842a89f1e00a0ca +7cdde0b2d7c324d5e17d8d3415ccad703507497ac95360ce660b656e5f66 +72a2f50761f3d02ccdc1d5692d7797699b8e2147cfd4817c81a432ff6a5f +39cc54927fa146cbed56a55f85f123c0a94b7553a8819b329d9dd122c502 +94e3f6314d5117db89ae7597c4691b6c542979a1ca3d26a8e23d3eb698c7 +1841651e08ec771cfb974d6613f2143872c739b62796bd0a45172530793c +28d93a65b59f79c245248d2c09428657a35b0c0e367bf7a4a4f0425b3f4b +485d9f402e164328a4b963f456829a39035c00283d2e4fcb71a42da6d42a +d46cb751287de34e6519c60bb3f1a6ba91f7bfa21dca96ee712af5681701 +18ece8a0535d9ba1dd4bd835e004a2f38c5ba43c9b30d17045e5649fbbac +188922e442182d4bdafaefb39e00106a5a7765f3d67850471e3629e526af +8691f935b57bd38465665204a214fef1006ea37dc0781073ced5fc042781 +93650393c3cadfddedcc5550ed483bb6355f54600e9758e647f9c9711f1b +e7df05d0e50a698615307c18f6d4886f50188011ba499d03831185915f3f +77c4b9ce708d78423b110776aaaf90396be0381616d1e9b0c1dcf68b6396 +82399da2a7323bf42ae5347599ef4ae9e5c135522c5ecb87e201853eb899 +db60d24acad17d6b7c2c7ea4dc221f3cb6d6caacd1ac0822ea3242ad9b4d +d15116c3874e3012fad26074a23b3cc7e25d67ef349811dbc6b87b53377f +0cf972040a037ecb91e3406a9bac68c9cab9be9a6bb28e93e3275b177cd5 +0b66935cbe8dd3d6a8365625db936b2cfc87d4d6e7322df3dbe6ccda2421 +a5e5372566f626a5e9d8bc66959e443286f8eb4bcfdeb6c49a799f1efa69 +63260d0ea2d51260baba9207fb246da927fc4c89e9c4dd5848fd4ef6f81a +cd836f5f06ff0fe135cafd7ab512af55a57727dd05a5fe1f7c3c7bbe8ea7 +e6680fcb3bbbee1cf2e2c0bba20185f00e2dc3afd42f22de472cdb3eaa5a +ddf8c6fb3682eea5548c51ddca25ca615221127b4438ea535ab3089c9ed9 +b971f35245cf831d9461a5da9d57bc4e5606d26535a7414cef6aee2a7b95 +bf2276044818ee0f3b0a16532934b8b745d8137b42ec2b28fae7d55fc02c +9ccfa4e0055f8a4be96e1e235c01b8b6ad509b832a3e90161e0a449934e7 +4be973c939b31cbc19dad4c58e9be89d242f0ce200548cdd4fa2081ab3f8 +e01f358d5db24b7a50eb2096d833378921f561f132cd7988708ee10cffb6 +2256201801c667e176b1dfaecde9756d725bef093457805e16f550e8a7de +87ecd46e5b09646b73ee74f890a36867636911e4cda2c46a40e7d57cf297 +9696046614c85b1a47ba55c60544ebd3ad7d750d003bda56dd7eed8c4702 +f8b319aaeef9d3cdc59b3e63ee93c6e1e857af273eb90909ecf36ef4c276 +895c78aa762e5376c5c542f854fba864ebce56e4b0207091139f053c2c08 +3b7ddcd0a9909b52100002bc3f8c47bcb19e7a9cb58b1ac03fee95e81195 +072d3aa7c8079632725f63425a3550a947834d29ac9a26d0774e90248e18 +996731fd9aa53ab62b40ce557d98e874b763d9d629a173f0c7babfc00ae7 +82daef5f00cf3608ebeef403dbbc19e16a1d160b889f4a10359d9eacc19d +7b5f126b31720dce7fc35ec861dfa56ea23fa18423ff4e8fe6e53fc6ba16 +b95a2b5dec00f614e4f835281ee0b4bf549e7e882689e0b445dd46fc40c9 +090e5575fa2c34b02a51ad0bccf6a7bb83ca3b929285e5e9fd054b72c47b +733a66c5abda526b18b2e49d0746e067e63b948a45eab2f4221c5b62ae21 +a5d9d7cd8aa9eeb49588891d22c56b14b55ceb6488f02b73ab3b7f6c5555 +b75452594658255e4cd58ac4815f2e1bc3888c6777f62aac2f0a57d416c3 +765c991f0f9a33d888aeb2d527b482c042ee23783a04a73ad13dfc590a52 +f3116f8296cacc7ab29b7d87e7864561a5d0a12bde2d36ee697064f41d1b +ca6ef2f801caab5295d19bf4c02b10c19f73b44635ba48a0806b967d7dfc +ce9a4850171a78532cb30020c0d66b3b1e7c75eaa7894904c181a022e8bc +9b2b8ef1202f3c7d36bcab4742d4a4761bb55b64da0d99685d319f5da8fa +132be6c0483f50e2657ae8af1e28f969440d6ed43eb00e95fd9e1cd490a4 +8646f6d008598751f7a41b43fbec7770fe591012b6b0c4ae18775ccc7db5 +de0ded2dd53e82c89648d46f0d0cc5d3ac5aa104239608d512a4353b9547 +04fe6eb7e73d718323cf9d748b8ec5da01ec9358267de12cc22b05ef0312 +e4b6ac5dbb6d06d7f2d911f20d527f504d62547aef136834b3695df8044c +383b6145e824d3931a602f081d9d656f84987a1ef121772f1f5b37a116bb +d2e77d4ccda01411545d24e15ce595db4cd62ee876b8754df0b85b44e011 +b82d76ce45795e6c2c58be8690b734a8880a074f303a70da4a1b086a6de6 +56c02cc7a4c25258eff18cb0fd868214bb46f972e26509f868d065b3cb14 +1c316898cf22293391bd7051ac3a6927aada952a8fd0658ce63357c07f34 +acbf8c99a5537da0023e901f0eb5547e1b466b7d982c8c539798b76ee2a2 +252437a81a37c3b63f625172d682eeed0b795860b2755f020ef52a138353 +003c61be2052cdd7d73b2cdcd26b127660a7b22fc51a6a2f6034f37e3e46 +c1d7f83f8b28c7c965993abba1d358362833580d9c63fa85d4cb949f97de +579fb6807b95a58b78f596db50055947dd0d0e597d9687083e9bc0266e86 +90b884b27f4094d8fb82ffdbaac4d580340a9ef8aa242be87e54b601af19 +87a48d267c04e371ae77163ebd0de3f5297b1060442ecdeac38334844e38 +0f294d4be73935fd8a38de7fba6d082c3d9156d7e88f2cfff0459377cbb6 +041f37a7e05010753b98e0b67d5827aa312129bb3c3bd883c12323756406 +d555720da8a0bb30edcfa760c01ecc2ba3b15fecccf5a10e9f358822e0ff +b64178fce2ea6a1105bfb72df0e4bc499b207ae26b8ea960de48e7ee7010 +b4e671dff795e4cdc5b43e81b1604d224f0616ae311f1208859c502c1a10 +940e7b9cd11be728bd3a0c8005ae23aea32c1b642812198a6f1aed32cb75 +97152b1340dd35ada1b81051e393d38f3740fa9523df6a83b8ca7dbceb33 +6e299b54cd998d4dfef804733c76156585e42b7284cbcc4047ba6b290efc +aa60953e98cd2b4bc2893857fa6a339f820142a52ccab0df09a2709df550 +f22e5921cbca408e7998cc1cccb8adf6d8f8b71e6685ae59d290fa33f5cd +664d73e434237424060f634262f04e9a71a977556e93b692ddc3aad26d92 +97dde71e4def64932151ad572af6e681082e9944ddbec6e7a8bdfd534233 +9ca3106ca1ccc80eab14f1655978b137fad8f399df7cbfa2d7d3d9675e0e +9afec37369a8ede2c93145ab3f42a375926946680c215fa16bf7416fc892 +bacd806cd424b9f85b47802c4336918f7486af2a03bf0d39b10169d35494 +419cb1ab7b8f407897f70c18303e91563b497d70b7181ede6aa0c3efe089 +ca6135b34dd1019b298e3677f8da61f864a67023c31eaa716c40cf3d397f +9a1209564c9ec759c37028079661d2a56374203c78b023ec61340bce5d96 +e477a4f77e5c0db7c0d1257b4bbbc6f889b17e6eaab045b8adef6f931e4d +0795583d60a6b7002cf61639c6f930671f3b8ac05a1c4e002f4bfc50d8b2 +3029fc4dce1b602cc3a5533336271bccc226559ffb127e3a562f92f89824 +552b9a70466d5a3c74ae515a222b109d490f26e8fc2d9d72bc8af6d1dcc7 +80463c7af81993bac2ce4aece9d95ab736b1dc73e32d1237bc8ec2b52513 +36dbabb4ecc7ceb5d18b02043281eb9a3bfdf19bc4853c9b1722ef1cdcf4 +fcec534923db2e2653dc48545a9850c0ac2e4594abc9f7d18a0bcf2fadfb +bf085d465a4d10528312f5d790eb9511ca01061c0d94136b99a043bcf278 +c18223b1e0f1cc062b32b79e28dec2dc59a0aaa4b5f3506923c83e6a87fa +08a1d941bb644c994491cf7f3b0e2ccf6c8a8ba89376f76dfdb592374f93 +528e78e31e0b18719346b9f1486f652638e3120687774030444674cb0778 +96385c41f6566819652d825dd58f9a4308ff79b45d7828dcbfebc406e40a +c46e866cb0e3e97d6ce7fcac19a9d0fe39bbde66c5f0cf775eb3b1e6d7e1 +1f67e7edb3d5c4facc85c916bf13322b56a0414ca27d145cb740fa2c37cd +8c142d9301f1ac3704cf6a8e93973a07fde5a331cf0cbb370c7ba555de61 +18a6cea0ecb2c0e37152390cc57e2e4fb3791ddbc383ee26b6f4006d0d68 +4880888011020f856a9de47f45440f127cf27ccaea7d40a3869d39ec7dec +ebc06382d294717644b6118354e15544fd4c6d88df9245c9a83b30e6ce09 +e2498dd1df488a019b179cb859889e6ad2838f749e3b038b280ebc8d5c3a +b03e8f15751214691edf0f86281e612d7ec0773c8a5d2b433266402df62f +fcc06879ca196aaf1fc73a5f01ac46b44d6cbe7743ae9a862c20445ae2be +1544f413d010280cc2941900bf3c42ec088cb21b44a915bb810e7666b545 +5324465c5943eedcef0c09128a995f431382e2062f5e39f4338c8eba1bca +e553cb60bb8f3e5038ac8073398c49f06dc734b18afa7921ea0d455e6e73 +db8ad9f77fb5ba6c28af6b4f18cbe46cf842c82d6c960be1520a5fd929df +ac7e00ede976fb2be0a07f659079a421fca693de89ce9b8fcb42b0176d9d +f3ddd58f921e13e216933d27b49d175b423751c451be7618eaab054d3b8c +23e8dd6fd60182d61e9b5c86b3b764a29a62f913ee7524d8cb33737d7224 +d95dc4bb8c2ad6397604a0ffecc8865adcb540e5da1cd769077838515118 +ebc9f0b988545c1881dd2e7a8fd73e11bd7ae9085fb4d45526b23a346b0f +e4281ee3d588106db5f7c386c488d8f2f4dd02d4c08e74c1034f987a44e5 +d39fd07538de57a42987ce290fb2f6557e8b5cbcaec168f5780927226415 +1e11e3667d33b36a793aa53e9e2d1102c9eb30cb3ba0ebac953e0227fe4a +3d3c0eb57e4390c3d35db0c41946e45be2830a1ae33fa25cf2c7c9cb4550 +ce9ff6c6e3d628fc7284daa6241604c90dde6339b7f7e7df3733416cdac8 +e5291357e4983d74d3582a490438a7fdb0af97001a31990b1de68e6adb48 +917daa387e647f9f13312db57310c7dedc2a2ea80800b4f4bbaa99c6b7b2 +7ac8345cb659489307e2565ebfd17774642c9ae5d3c18068dc35170c7d58 +4cf4173f1baf98137fa249c81f3347e1dadd6b1ba0f50c3b64c1eab183a0 +937b0f7278eff101e5267fa6480da7d602844416490c2c2c7eb0d44ac8f4 +75cfd611db5ec268db07c0b3608825c3e12834a2b2efaf5e2723c5199c42 +6011cf22e64e4c0d31d563f321097935ea0c6fcbf5acd3748d90079f6ab8 +687288dc55df29fe7958f566b27b73e2ea30747247f7a2b2add0602c7d64 +d23f52e7c96748e6a54ee8c4629b2aab8882169653f0ba7f05236bf14364 +244720f3259cbed73a318b29e4a9305deb65a2c9dec8a9d0f9a9f6fae541 +83e0f4b9a9a567057a1794945168dc23cec25d1c02ea9242c9fb6d8fc11e +e8874bd80a5226373ae87cea91853d0625c777ceb1f5a6f3debcf2f75a61 +460c7b4067f568ecd01f62901ade8bf8fbc5db9c6720420496f0cb48a002 +99870773c2e7b12e83987a5d0290d9bbf589ac889bf7d4334a5147187a7f +71008f216ce917ca4cfba5347078f354897fd87ac48af6a6c62711d2eb3a +5882bf3b32c0f1bfda976f850c9dcb97170e78c229a27fd5e292d161ece9 +a8c47a223cbdc28e24f79f6429c72b5752a08f917feda941582c36d9acb5 +748c86072858d053170fdbf708971a0bd5a8d8034ec769cb72ea88eb5cd7 +49f35be6ee5e9b5df6021926cae9dac3f5ec2b33680b12e95fd4ecbf28eb +a0503c10c6f2be6c7c47e9d66a0fae6038441c50e6447892f4aaf0a25ccd +952c2e8b201bb479099f16fc4903993ac18d4667c84c124685ae7648a826 +6bc1701cc600964fdcc01258a72104a0e5e9996b34c2691a66fa20f48d7c +2522333dfdabf3785f37dd9b021e8ee29fa10f76f43d5f935996cbf9d98d +92d0a84ce65613f7c4a5052f4c408bf10679fc28a4a9ff848d9e0c4976bb +dfdfb78bb934cd72434db596cb49e199f386a0bda69449ce2e11e3a4f53d +be134c6d7fe452a0927cf6a9a15b2406f8bd354adcde0ce136378baa565f +b9c51a03b1fbe1e166a1f92af26bd9f072250aaa6596a236ba2d5a200c90 +a760ca050421abc78223b2e8b2eea958ab23084fa1947574e846e48aeb12 +26cebb8b5a92089e9ea771557599e2fff44d75bcf600e76ae7289ba98cf3 +98208c5104562834f568ebd62801b988b0a9fdf132b6564566103b3d2d8e +6a099b7fbad8a13b8cd7f6729bb6651fc1019e66c4bd6ff27410bd5cdae7 +4010bd68b066bffdb4fd5e3dd9cf7e1a1353f7a4c5157e3ad508f4ca0259 +9761b7cdd6a81b3560b8765be3b0432fe4c25dcb4001b00c7fa62874f681 +ed22127dc3974605a05be8d8fcf9701f859ffce4dc598091891ab7596ac3 +4cd851ecfd2dbbaa2f99dac376f7bb40703fd0700d7499a7c24726bdc9bb +3b88c6a82e52686c1ee945d8825092bc81848a08722ac5a1d24353f95ec8 +18f3fa487d9600318091b0ae9874b42bb3cb683a2518b18cc1bd86c6e5e8 +3d37c14ef4fe0c77b03a3314995b1e7c1066b98c4375bd1fc5fadee1b024 +7ece4f95a0f59978d543910deb2e5761632c74c508269c4e4b9e315bda02 +975dc771fc30c8164b9df9172a4e571d8ca578cd2aaeaa0dd083e74cdc2e +d938b984b96d76a64b8c5fd12e63220bbac41e5bcd5ccb6b84bdbf6a02d5 +934ac50c654c0853209a6758bcdf560e53566d78987484bb6672ebe93f22 +dcba14e3acc132a2d9ae837adde04d8b16 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +0000000000000000000000000000000000000000000000000000000000000000 +cleartomark +%%BeginResource: procset Altsys_header 4 0 +userdict begin /AltsysDict 245 dict def end +AltsysDict begin +/bdf{bind def}bind def +/xdf{exch def}bdf +/defed{where{pop true}{false}ifelse}bdf +/ndf{1 index where{pop pop pop}{dup xcheck{bind}if def}ifelse}bdf +/d{setdash}bdf +/h{closepath}bdf +/H{}bdf +/J{setlinecap}bdf +/j{setlinejoin}bdf +/M{setmiterlimit}bdf +/n{newpath}bdf +/N{newpath}bdf +/q{gsave}bdf +/Q{grestore}bdf +/w{setlinewidth}bdf +/sepdef{ + dup where not + { +AltsysSepDict + } + if + 3 1 roll exch put +}bdf +/st{settransfer}bdf +/colorimage defed /_rci xdf +/_NXLevel2 defed { + _NXLevel2 not { +/colorimage where { +userdict eq { +/_rci false def +} if +} if + } if +} if +/md defed{ + md type /dicttype eq { +/colorimage where { +md eq { +/_rci false def +}if +}if +/settransfer where { +md eq { +/st systemdict /settransfer get def +}if +}if + }if +}if +/setstrokeadjust defed +{ + true setstrokeadjust + /C{curveto}bdf + /L{lineto}bdf + /m{moveto}bdf +} +{ + /dr{transform .25 sub round .25 add +exch .25 sub round .25 add exch itransform}bdf + /C{dr curveto}bdf + /L{dr lineto}bdf + /m{dr moveto}bdf + /setstrokeadjust{pop}bdf +}ifelse +/rectstroke defed /xt xdf +xt {/yt save def} if +/privrectpath { + 4 -2 roll m + dtransform round exch round exch idtransform + 2 copy 0 lt exch 0 lt xor + {dup 0 exch rlineto exch 0 rlineto neg 0 exch rlineto} + {exch dup 0 rlineto exch 0 exch rlineto neg 0 rlineto} + ifelse + closepath +}bdf +/rectclip{newpath privrectpath clip newpath}def +/rectfill{gsave newpath privrectpath fill grestore}def +/rectstroke{gsave newpath privrectpath stroke grestore}def +xt {yt restore} if +/_fonthacksave false def +/currentpacking defed +{ + /_bfh {/_fonthacksave currentpacking def false setpacking} bdf + /_efh {_fonthacksave setpacking} bdf +} +{ + /_bfh {} bdf + /_efh {} bdf +}ifelse +/packedarray{array astore readonly}ndf +/` +{ + false setoverprint + + + /-save0- save def + 5 index concat + pop + storerect left bottom width height rectclip + pop + + /dict_count countdictstack def + /op_count count 1 sub def + userdict begin + + /showpage {} def + + 0 setgray 0 setlinecap 1 setlinewidth + 0 setlinejoin 10 setmiterlimit [] 0 setdash newpath + +} bdf +/currentpacking defed{true setpacking}if +/min{2 copy gt{exch}if pop}bdf +/max{2 copy lt{exch}if pop}bdf +/xformfont { currentfont exch makefont setfont } bdf +/fhnumcolors 1 + statusdict begin +/processcolors defed +{ +pop processcolors +} +{ +/deviceinfo defed { +deviceinfo /Colors known { +pop deviceinfo /Colors get +} if +} if +} ifelse + end +def +/printerRes + gsave + matrix defaultmatrix setmatrix + 72 72 dtransform + abs exch abs + max + grestore + def +/graycalcs +[ + {Angle Frequency} + {GrayAngle GrayFrequency} + {0 Width Height matrix defaultmatrix idtransform +dup mul exch dup mul add sqrt 72 exch div} + {0 GrayWidth GrayHeight matrix defaultmatrix idtransform +dup mul exch dup mul add sqrt 72 exch div} +] def +/calcgraysteps { + forcemaxsteps + { +maxsteps + } + { +/currenthalftone defed +{currenthalftone /dicttype eq}{false}ifelse +{ +currenthalftone begin +HalftoneType 4 le +{graycalcs HalftoneType 1 sub get exec} +{ +HalftoneType 5 eq +{ +Default begin +{graycalcs HalftoneType 1 sub get exec} +end +} +{0 60} +ifelse +} +ifelse +end +} +{ +currentscreen pop exch +} +ifelse + +printerRes 300 max exch div exch +2 copy +sin mul round dup mul +3 1 roll +cos mul round dup mul +add 1 add +dup maxsteps gt {pop maxsteps} if + } + ifelse +} bdf +/nextrelease defed { + /languagelevel defed not { +/framebuffer defed { +0 40 string framebuffer 9 1 roll 8 {pop} repeat +dup 516 eq exch 520 eq or +{ +/fhnumcolors 3 def +/currentscreen {60 0 {pop pop 1}}bdf +/calcgraysteps {maxsteps} bdf +}if +}if + }if +}if +fhnumcolors 1 ne { + /calcgraysteps {maxsteps} bdf +} if +/currentpagedevice defed { + + + currentpagedevice /PreRenderingEnhance known + { +currentpagedevice /PreRenderingEnhance get +{ +/calcgraysteps +{ +forcemaxsteps +{maxsteps} +{256 maxsteps min} +ifelse +} def +} if + } if +} if +/gradfrequency 144 def +printerRes 1000 lt { + /gradfrequency 72 def +} if +/adjnumsteps { + + dup dtransform abs exch abs max + + printerRes div + + gradfrequency mul + round + 5 max + min +}bdf +/goodsep { + spots exch get 4 get dup sepname eq exch (_vc_Registration) eq or +}bdf +/BeginGradation defed +{/bb{BeginGradation}bdf} +{/bb{}bdf} +ifelse +/EndGradation defed +{/eb{EndGradation}bdf} +{/eb{}bdf} +ifelse +/bottom -0 def +/delta -0 def +/frac -0 def +/height -0 def +/left -0 def +/numsteps1 -0 def +/radius -0 def +/right -0 def +/top -0 def +/width -0 def +/xt -0 def +/yt -0 def +/df currentflat def +/tempstr 1 string def +/clipflatness currentflat def +/inverted? + 0 currenttransfer exec .5 ge def +/tc1 [0 0 0 1] def +/tc2 [0 0 0 1] def +/storerect{/top xdf /right xdf /bottom xdf /left xdf +/width right left sub def /height top bottom sub def}bdf +/concatprocs{ + systemdict /packedarray known + {dup type /packedarraytype eq 2 index type /packedarraytype eq or}{false}ifelse + { +/proc2 exch cvlit def /proc1 exch cvlit def +proc1 aload pop proc2 aload pop +proc1 length proc2 length add packedarray cvx + } + { +/proc2 exch cvlit def /proc1 exch cvlit def +/newproc proc1 length proc2 length add array def +newproc 0 proc1 putinterval newproc proc1 length proc2 putinterval +newproc cvx + }ifelse +}bdf +/i{dup 0 eq + {pop df dup} + {dup} ifelse + /clipflatness xdf setflat +}bdf +version cvr 38.0 le +{/setrgbcolor{ +currenttransfer exec 3 1 roll +currenttransfer exec 3 1 roll +currenttransfer exec 3 1 roll +setrgbcolor}bdf}if +/vms {/vmsv save def} bdf +/vmr {vmsv restore} bdf +/vmrs{vmsv restore /vmsv save def}bdf +/eomode{ + {/filler /eofill load def /clipper /eoclip load def} + {/filler /fill load def /clipper /clip load def} + ifelse +}bdf +/normtaper{}bdf +/logtaper{9 mul 1 add log}bdf +/CD{ + /NF exch def + { +exch dup +/FID ne 1 index/UniqueID ne and +{exch NF 3 1 roll put} +{pop pop} +ifelse + }forall + NF +}bdf +/MN{ + 1 index length + /Len exch def + dup length Len add + string dup + Len + 4 -1 roll + putinterval + dup + 0 + 4 -1 roll + putinterval +}bdf +/RC{4 -1 roll /ourvec xdf 256 string cvs(|______)anchorsearch + {1 index MN cvn/NewN exch def cvn + findfont dup maxlength dict CD dup/FontName NewN put dup + /Encoding ourvec put NewN exch definefont pop}{pop}ifelse}bdf +/RF{ + dup + FontDirectory exch + known + {pop 3 -1 roll pop} + {RC} + ifelse +}bdf +/FF{dup 256 string cvs(|______)exch MN cvn dup FontDirectory exch known + {exch pop findfont 3 -1 roll pop} + {pop dup findfont dup maxlength dict CD dup dup + /Encoding exch /Encoding get 256 array copy 7 -1 roll + {3 -1 roll dup 4 -2 roll put}forall put definefont} + ifelse}bdf +/RFJ{ + dup + FontDirectory exch + known + {pop 3 -1 roll pop + FontDirectory /Ryumin-Light-83pv-RKSJ-H known + {pop pop /Ryumin-Light-83pv-RKSJ-H dup}if + } + {RC} + ifelse +}bdf +/FFJ{dup 256 string cvs(|______)exch MN cvn dup FontDirectory exch known + {exch pop findfont 3 -1 roll pop} + {pop +dup FontDirectory exch known not + {FontDirectory /Ryumin-Light-83pv-RKSJ-H known +{pop /Ryumin-Light-83pv-RKSJ-H}if + }if + dup findfont dup maxlength dict CD dup dup + /Encoding exch /Encoding get 256 array copy 7 -1 roll + {3 -1 roll dup 4 -2 roll put}forall put definefont} + ifelse}bdf +/fps{ + currentflat + exch + dup 0 le{pop 1}if + { +dup setflat 3 index stopped +{1.3 mul dup 3 index gt{pop setflat pop pop stop}if} +{exit} +ifelse + }loop + pop setflat pop pop +}bdf +/fp{100 currentflat fps}bdf +/clipper{clip}bdf +/W{/clipper load 100 clipflatness dup setflat fps}bdf +userdict begin /BDFontDict 29 dict def end +BDFontDict begin +/bu{}def +/bn{}def +/setTxMode{av 70 ge{pop}if pop}def +/gm{m}def +/show{pop}def +/gr{pop}def +/fnt{pop pop pop}def +/fs{pop}def +/fz{pop}def +/lin{pop pop}def +/:M {pop pop} def +/sf {pop} def +/S {pop} def +/@b {pop pop pop pop pop pop pop pop} def +/_bdsave /save load def +/_bdrestore /restore load def +/save { dup /fontsave eq {null} {_bdsave} ifelse } def +/restore { dup null eq { pop } { _bdrestore } ifelse } def +/fontsave null def +end +/MacVec 256 array def +MacVec 0 /Helvetica findfont +/Encoding get 0 128 getinterval putinterval +MacVec 127 /DEL put MacVec 16#27 /quotesingle put MacVec 16#60 /grave put +/NUL/SOH/STX/ETX/EOT/ENQ/ACK/BEL/BS/HT/LF/VT/FF/CR/SO/SI +/DLE/DC1/DC2/DC3/DC4/NAK/SYN/ETB/CAN/EM/SUB/ESC/FS/GS/RS/US +MacVec 0 32 getinterval astore pop +/Adieresis/Aring/Ccedilla/Eacute/Ntilde/Odieresis/Udieresis/aacute +/agrave/acircumflex/adieresis/atilde/aring/ccedilla/eacute/egrave +/ecircumflex/edieresis/iacute/igrave/icircumflex/idieresis/ntilde/oacute +/ograve/ocircumflex/odieresis/otilde/uacute/ugrave/ucircumflex/udieresis +/dagger/degree/cent/sterling/section/bullet/paragraph/germandbls +/registered/copyright/trademark/acute/dieresis/notequal/AE/Oslash +/infinity/plusminus/lessequal/greaterequal/yen/mu/partialdiff/summation +/product/pi/integral/ordfeminine/ordmasculine/Omega/ae/oslash +/questiondown/exclamdown/logicalnot/radical/florin/approxequal/Delta/guillemotleft +/guillemotright/ellipsis/nbspace/Agrave/Atilde/Otilde/OE/oe +/endash/emdash/quotedblleft/quotedblright/quoteleft/quoteright/divide/lozenge +/ydieresis/Ydieresis/fraction/currency/guilsinglleft/guilsinglright/fi/fl +/daggerdbl/periodcentered/quotesinglbase/quotedblbase +/perthousand/Acircumflex/Ecircumflex/Aacute +/Edieresis/Egrave/Iacute/Icircumflex/Idieresis/Igrave/Oacute/Ocircumflex +/apple/Ograve/Uacute/Ucircumflex/Ugrave/dotlessi/circumflex/tilde +/macron/breve/dotaccent/ring/cedilla/hungarumlaut/ogonek/caron +MacVec 128 128 getinterval astore pop +end %. AltsysDict +%%EndResource +%%EndProlog +%%BeginSetup +AltsysDict begin +_bfh +%%IncludeResource: font Symbol +_efh +0 dict dup begin +end +/f0 /Symbol FF def +_bfh +%%IncludeResource: font ZapfHumanist601BT-Bold +_efh +0 dict dup begin +end +/f1 /ZapfHumanist601BT-Bold FF def +end %. AltsysDict +%%EndSetup +AltsysDict begin +/onlyk4{false}ndf +/ccmyk{dup 5 -1 roll sub 0 max exch}ndf +/cmyk2gray{ + 4 -1 roll 0.3 mul 4 -1 roll 0.59 mul 4 -1 roll 0.11 mul + add add add 1 min neg 1 add +}bdf +/setcmykcolor{1 exch sub ccmyk ccmyk ccmyk pop setrgbcolor}ndf +/maxcolor { + max max max +} ndf +/maxspot { + pop +} ndf +/setcmykcoloroverprint{4{dup -1 eq{pop 0}if 4 1 roll}repeat setcmykcolor}ndf +/findcmykcustomcolor{5 packedarray}ndf +/setcustomcolor{exch aload pop pop 4{4 index mul 4 1 roll}repeat setcmykcolor pop}ndf +/setseparationgray{setgray}ndf +/setoverprint{pop}ndf +/currentoverprint false ndf +/cmykbufs2gray{ + 0 1 2 index length 1 sub + { +4 index 1 index get 0.3 mul +4 index 2 index get 0.59 mul +4 index 3 index get 0.11 mul +4 index 4 index get +add add add cvi 255 min +255 exch sub +2 index 3 1 roll put + }for + 4 1 roll pop pop pop +}bdf +/colorimage{ + pop pop + [ +5 -1 roll/exec cvx +6 -1 roll/exec cvx +7 -1 roll/exec cvx +8 -1 roll/exec cvx +/cmykbufs2gray cvx + ]cvx + image +} +%. version 47.1 on Linotronic of Postscript defines colorimage incorrectly (rgb model only) +version cvr 47.1 le +statusdict /product get (Lino) anchorsearch{pop pop true}{pop false}ifelse +and{userdict begin bdf end}{ndf}ifelse +fhnumcolors 1 ne {/yt save def} if +/customcolorimage{ + aload pop + (_vc_Registration) eq + { +pop pop pop pop separationimage + } + { +/ik xdf /iy xdf /im xdf /ic xdf +ic im iy ik cmyk2gray /xt xdf +currenttransfer +{dup 1.0 exch sub xt mul add}concatprocs +st +image + } + ifelse +}ndf +fhnumcolors 1 ne {yt restore} if +fhnumcolors 3 ne {/yt save def} if +/customcolorimage{ + aload pop + (_vc_Registration) eq + { +pop pop pop pop separationimage + } + { +/ik xdf /iy xdf /im xdf /ic xdf +1.0 dup ic ik add min sub +1.0 dup im ik add min sub +1.0 dup iy ik add min sub +/ic xdf /iy xdf /im xdf +currentcolortransfer +4 1 roll +{dup 1.0 exch sub ic mul add}concatprocs 4 1 roll +{dup 1.0 exch sub iy mul add}concatprocs 4 1 roll +{dup 1.0 exch sub im mul add}concatprocs 4 1 roll +setcolortransfer +{/dummy xdf dummy}concatprocs{dummy}{dummy}true 3 colorimage + } + ifelse +}ndf +fhnumcolors 3 ne {yt restore} if +fhnumcolors 4 ne {/yt save def} if +/customcolorimage{ + aload pop + (_vc_Registration) eq + { +pop pop pop pop separationimage + } + { +/ik xdf /iy xdf /im xdf /ic xdf +currentcolortransfer +{1.0 exch sub ik mul ik sub 1 add}concatprocs 4 1 roll +{1.0 exch sub iy mul iy sub 1 add}concatprocs 4 1 roll +{1.0 exch sub im mul im sub 1 add}concatprocs 4 1 roll +{1.0 exch sub ic mul ic sub 1 add}concatprocs 4 1 roll +setcolortransfer +{/dummy xdf dummy}concatprocs{dummy}{dummy}{dummy} +true 4 colorimage + } + ifelse +}ndf +fhnumcolors 4 ne {yt restore} if +/separationimage{image}ndf +/newcmykcustomcolor{6 packedarray}ndf +/inkoverprint false ndf +/setinkoverprint{pop}ndf +/setspotcolor { + spots exch get + dup 4 get (_vc_Registration) eq + {pop 1 exch sub setseparationgray} + {0 5 getinterval exch setcustomcolor} + ifelse +}ndf +/currentcolortransfer{currenttransfer dup dup dup}ndf +/setcolortransfer{st pop pop pop}ndf +/fas{}ndf +/sas{}ndf +/fhsetspreadsize{pop}ndf +/filler{fill}bdf +/F{gsave {filler}fp grestore}bdf +/f{closepath F}bdf +/S{gsave {stroke}fp grestore}bdf +/s{closepath S}bdf +/bc4 [0 0 0 0] def +/_lfp4 { + /iosv inkoverprint def + /cosv currentoverprint def + /yt xdf + /xt xdf + /ang xdf + storerect + /taperfcn xdf + /k2 xdf /y2 xdf /m2 xdf /c2 xdf + /k1 xdf /y1 xdf /m1 xdf /c1 xdf + c1 c2 sub abs + m1 m2 sub abs + y1 y2 sub abs + k1 k2 sub abs + maxcolor + calcgraysteps mul abs round + height abs adjnumsteps + dup 2 lt {pop 1} if + 1 sub /numsteps1 xdf + currentflat mark + currentflat clipflatness + /delta top bottom sub numsteps1 1 add div def + /right right left sub def + /botsv top delta sub def + { +{ +W +xt yt translate +ang rotate +xt neg yt neg translate +dup setflat +/bottom botsv def +0 1 numsteps1 +{ +numsteps1 dup 0 eq {pop 0.5 } { div } ifelse +taperfcn /frac xdf +bc4 0 c2 c1 sub frac mul c1 add put +bc4 1 m2 m1 sub frac mul m1 add put +bc4 2 y2 y1 sub frac mul y1 add put +bc4 3 k2 k1 sub frac mul k1 add put +bc4 vc +1 index setflat +{ +mark {newpath left bottom right delta rectfill}stopped +{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} +{cleartomark exit}ifelse +}loop +/bottom bottom delta sub def +}for +} +gsave stopped grestore +{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} +{exit}ifelse + }loop + cleartomark setflat + iosv setinkoverprint + cosv setoverprint +}bdf +/bcs [0 0] def +/_lfs4 { + /iosv inkoverprint def + /cosv currentoverprint def + /yt xdf + /xt xdf + /ang xdf + storerect + /taperfcn xdf + /tint2 xdf + /tint1 xdf + bcs exch 1 exch put + tint1 tint2 sub abs + bcs 1 get maxspot + calcgraysteps mul abs round + height abs adjnumsteps + dup 2 lt {pop 2} if + 1 sub /numsteps1 xdf + currentflat mark + currentflat clipflatness + /delta top bottom sub numsteps1 1 add div def + /right right left sub def + /botsv top delta sub def + { +{ +W +xt yt translate +ang rotate +xt neg yt neg translate +dup setflat +/bottom botsv def +0 1 numsteps1 +{ +numsteps1 div taperfcn /frac xdf +bcs 0 +1.0 tint2 tint1 sub frac mul tint1 add sub +put bcs vc +1 index setflat +{ +mark {newpath left bottom right delta rectfill}stopped +{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} +{cleartomark exit}ifelse +}loop +/bottom bottom delta sub def +}for +} +gsave stopped grestore +{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} +{exit}ifelse + }loop + cleartomark setflat + iosv setinkoverprint + cosv setoverprint +}bdf +/_rfs4 { + /iosv inkoverprint def + /cosv currentoverprint def + /tint2 xdf + /tint1 xdf + bcs exch 1 exch put + /radius xdf + /yt xdf + /xt xdf + tint1 tint2 sub abs + bcs 1 get maxspot + calcgraysteps mul abs round + radius abs adjnumsteps + dup 2 lt {pop 2} if + 1 sub /numsteps1 xdf + radius numsteps1 div 2 div /halfstep xdf + currentflat mark + currentflat clipflatness + { +{ +dup setflat +W +0 1 numsteps1 +{ +dup /radindex xdf +numsteps1 div /frac xdf +bcs 0 +tint2 tint1 sub frac mul tint1 add +put bcs vc +1 index setflat +{ +newpath mark xt yt radius 1 frac sub mul halfstep add 0 360 +{ arc +radindex numsteps1 ne +{ +xt yt +radindex 1 add numsteps1 +div 1 exch sub +radius mul halfstep add +dup xt add yt moveto +360 0 arcn +} if +fill +}stopped +{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} +{cleartomark exit}ifelse +}loop +}for +} +gsave stopped grestore +{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} +{exit}ifelse + }loop + cleartomark setflat + iosv setinkoverprint + cosv setoverprint +}bdf +/_rfp4 { + /iosv inkoverprint def + /cosv currentoverprint def + /k2 xdf /y2 xdf /m2 xdf /c2 xdf + /k1 xdf /y1 xdf /m1 xdf /c1 xdf + /radius xdf + /yt xdf + /xt xdf + c1 c2 sub abs + m1 m2 sub abs + y1 y2 sub abs + k1 k2 sub abs + maxcolor + calcgraysteps mul abs round + radius abs adjnumsteps + dup 2 lt {pop 1} if + 1 sub /numsteps1 xdf + radius numsteps1 dup 0 eq {pop} {div} ifelse + 2 div /halfstep xdf + currentflat mark + currentflat clipflatness + { +{ +dup setflat +W +0 1 numsteps1 +{ +dup /radindex xdf +numsteps1 dup 0 eq {pop 0.5 } { div } ifelse +/frac xdf +bc4 0 c2 c1 sub frac mul c1 add put +bc4 1 m2 m1 sub frac mul m1 add put +bc4 2 y2 y1 sub frac mul y1 add put +bc4 3 k2 k1 sub frac mul k1 add put +bc4 vc +1 index setflat +{ +newpath mark xt yt radius 1 frac sub mul halfstep add 0 360 +{ arc +radindex numsteps1 ne +{ +xt yt +radindex 1 add +numsteps1 dup 0 eq {pop} {div} ifelse +1 exch sub +radius mul halfstep add +dup xt add yt moveto +360 0 arcn +} if +fill +}stopped +{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} +{cleartomark exit}ifelse +}loop +}for +} +gsave stopped grestore +{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} +{exit}ifelse + }loop + cleartomark setflat + iosv setinkoverprint + cosv setoverprint +}bdf +/lfp4{_lfp4}ndf +/lfs4{_lfs4}ndf +/rfs4{_rfs4}ndf +/rfp4{_rfp4}ndf +/cvc [0 0 0 1] def +/vc{ + AltsysDict /cvc 2 index put + aload length 4 eq + {setcmykcolor} + {setspotcolor} + ifelse +}bdf +/origmtx matrix currentmatrix def +/ImMatrix matrix currentmatrix def +0 setseparationgray +/imgr {1692 1570.1102 2287.2756 2412 } def +/bleed 0 def +/clpr {1692 1570.1102 2287.2756 2412 } def +/xs 1 def +/ys 1 def +/botx 0 def +/overlap 0 def +/wdist 18 def +0 2 mul fhsetspreadsize +0 0 ne {/df 0 def /clipflatness 0 def} if +/maxsteps 256 def +/forcemaxsteps false def +vms +-1845 -1956 translate +/currentpacking defed{false setpacking}if +/spots[ +1 0 0 0 (Process Cyan) false newcmykcustomcolor +0 1 0 0 (Process Magenta) false newcmykcustomcolor +0 0 1 0 (Process Yellow) false newcmykcustomcolor +0 0 0 1 (Process Black) false newcmykcustomcolor +]def +/textopf false def +/curtextmtx{}def +/otw .25 def +/msf{dup/curtextmtx xdf makefont setfont}bdf +/makesetfont/msf load def +/curtextheight{.707104 .707104 curtextmtx dtransform + dup mul exch dup mul add sqrt}bdf +/ta2{ +tempstr 2 index gsave exec grestore +cwidth cheight rmoveto +4 index eq{5 index 5 index rmoveto}if +2 index 2 index rmoveto +}bdf +/ta{exch systemdict/cshow known +{{/cheight xdf/cwidth xdf tempstr 0 2 index put ta2}exch cshow} +{{tempstr 0 2 index put tempstr stringwidth/cheight xdf/cwidth xdf ta2}forall} +ifelse 6{pop}repeat}bdf +/sts{/textopf currentoverprint def vc setoverprint +/ts{awidthshow}def exec textopf setoverprint}bdf +/stol{/xt currentlinewidth def + setlinewidth vc newpath + /ts{{false charpath stroke}ta}def exec + xt setlinewidth}bdf + +/strk{/textopf currentoverprint def vc setoverprint + /ts{{false charpath stroke}ta}def exec + textopf setoverprint + }bdf +n +[] 0 d +3.863708 M +1 w +0 j +0 J +false setoverprint +0 i +false eomode +[0 0 0 1] vc +vms +%white border -- disabled +%1845.2293 2127.8588 m +%2045.9437 2127.8588 L +%2045.9437 1956.1412 L +%1845.2293 1956.1412 L +%1845.2293 2127.8588 L +%0.1417 w +%2 J +%2 M +%[0 0 0 0] vc +%s +n +1950.8 2097.2 m +1958.8 2092.5 1967.3 2089 1975.5 2084.9 C +1976.7 2083.5 1976.1 2081.5 1976.7 2079.9 C +1979.6 2081.1 1981.6 2086.8 1985.3 2084 C +1993.4 2079.3 2001.8 2075.8 2010 2071.7 C +2010.5 2071.5 2010.5 2071.1 2010.8 2070.8 C +2011.2 2064.3 2010.9 2057.5 2011 2050.8 C +2015.8 2046.9 2022.2 2046.2 2026.6 2041.7 C +2026.5 2032.5 2026.8 2022.9 2026.4 2014.1 C +2020.4 2008.3 2015 2002.4 2008.8 1997.1 C +2003.8 1996.8 2000.7 2001.2 1996.1 2002.1 C +1995.2 1996.4 1996.9 1990.5 1995.6 1984.8 C +1989.9 1979 1984.5 1973.9 1978.8 1967.8 C +1977.7 1968.6 1976 1967.6 1974.5 1968.3 C +1967.4 1972.5 1960.1 1976.1 1952.7 1979.3 C +1946.8 1976.3 1943.4 1970.7 1938.5 1966.1 C +1933.9 1966.5 1929.4 1968.8 1925.1 1970.7 C +1917.2 1978.2 1906 1977.9 1897.2 1983.4 C +1893.2 1985.6 1889.4 1988.6 1885 1990.1 C +1884.6 1990.6 1883.9 1991 1883.8 1991.6 C +1883.7 2000.4 1884 2009.9 1883.6 2018.9 C +1887.7 2024 1893.2 2028.8 1898 2033.8 C +1899.1 2035.5 1900.9 2036.8 1902.5 2037.9 C +1903.9 2037.3 1905.2 2036.6 1906.4 2035.5 C +1906.3 2039.7 1906.5 2044.6 1906.1 2048.9 C +1906.3 2049.6 1906.7 2050.2 1907.1 2050.8 C +1913.4 2056 1918.5 2062.7 1924.8 2068.1 C +1926.6 2067.9 1928 2066.9 1929.4 2066 C +1930.2 2071 1927.7 2077.1 1930.6 2081.6 C +1936.6 2086.9 1941.5 2092.9 1947.9 2097.9 C +1949 2098.1 1949.9 2097.5 1950.8 2097.2 C +[0 0 0 0.18] vc +f +0.4 w +S +n +1975.2 2084.7 m +1976.6 2083.4 1975.7 2081.1 1976 2079.4 C +1979.3 2079.5 1980.9 2086.2 1984.8 2084 C +1992.9 2078.9 2001.7 2075.6 2010 2071.2 C +2011 2064.6 2010.2 2057.3 2010.8 2050.6 C +2015.4 2046.9 2021.1 2045.9 2025.9 2042.4 C +2026.5 2033.2 2026.8 2022.9 2025.6 2013.9 C +2020.5 2008.1 2014.5 2003.1 2009.3 1997.6 C +2004.1 1996.7 2000.7 2001.6 1995.9 2002.6 C +1995.2 1996.7 1996.3 1990.2 1994.9 1984.6 C +1989.8 1978.7 1983.6 1973.7 1978.4 1968 C +1977.3 1969.3 1976 1967.6 1974.8 1968.5 C +1967.7 1972.7 1960.4 1976.3 1952.9 1979.6 C +1946.5 1976.9 1943.1 1970.5 1937.8 1966.1 C +1928.3 1968.2 1920.6 1974.8 1911.6 1978.4 C +1901.9 1979.7 1893.9 1986.6 1885 1990.6 C +1884.3 1991 1884.3 1991.7 1884 1992.3 C +1884.5 2001 1884.2 2011 1884.3 2019.9 C +1890.9 2025.3 1895.9 2031.9 1902.3 2037.4 C +1904.2 2037.9 1905.6 2034.2 1906.8 2035.7 C +1907.4 2040.9 1905.7 2046.1 1907.3 2050.8 C +1913.6 2056.2 1919.2 2062.6 1925.1 2067.9 C +1926.9 2067.8 1928 2066.3 1929.6 2065.7 C +1929.9 2070.5 1929.2 2076 1930.1 2080.8 C +1936.5 2086.1 1941.6 2092.8 1948.4 2097.6 C +1957.3 2093.3 1966.2 2088.8 1975.2 2084.7 C +[0 0 0 0] vc +f +S +n +1954.8 2093.8 m +1961.6 2090.5 1968.2 2087 1975 2084 C +1975 2082.8 1975.6 2080.9 1974.8 2080.6 C +1974.3 2075.2 1974.6 2069.6 1974.5 2064 C +1977.5 2059.7 1984.5 2060 1988.9 2056.4 C +1989.5 2055.5 1990.5 2055.3 1990.8 2054.4 C +1991.1 2045.7 1991.4 2036.1 1990.6 2027.8 C +1990.7 2026.6 1992 2027.3 1992.8 2027.1 C +1997 2032.4 2002.6 2037.8 2007.6 2042.2 C +2008.7 2042.3 2007.8 2040.6 2007.4 2040 C +2002.3 2035.6 1997.5 2030 1992.8 2025.2 C +1991.6 2024.7 1990.8 2024.9 1990.1 2025.4 C +1989.4 2024.9 1988.1 2025.2 1987.2 2024.4 C +1987.1 2025.8 1988.3 2026.5 1989.4 2026.8 C +1989.4 2026.6 1989.3 2026.2 1989.6 2026.1 C +1989.9 2026.2 1989.9 2026.6 1989.9 2026.8 C +1989.8 2026.6 1990 2026.5 1990.1 2026.4 C +1990.2 2027 1991.1 2028.3 1990.1 2028 C +1989.9 2037.9 1990.5 2044.1 1989.6 2054.2 C +1985.9 2058 1979.7 2057.4 1976 2061.2 C +1974.5 2061.6 1975.2 2059.9 1974.5 2059.5 C +1973.9 2058 1975.6 2057.8 1975 2056.6 C +1974.5 2057.1 1974.6 2055.3 1973.6 2055.9 C +1971.9 2059.3 1974.7 2062.1 1973.1 2065.5 C +1973.1 2071.2 1972.9 2077 1973.3 2082.5 C +1967.7 2085.6 1962 2088 1956.3 2090.7 C +1953.9 2092.4 1951 2093 1948.6 2094.8 C +1943.7 2089.9 1937.9 2084.3 1933 2079.6 C +1931.3 2076.1 1933.2 2071.3 1932.3 2067.2 C +1931.3 2062.9 1933.3 2060.6 1932 2057.6 C +1932.7 2056.5 1930.9 2053.3 1933.2 2051.8 C +1936.8 2050.1 1940.1 2046.9 1944 2046.8 C +1946.3 2049.7 1949.3 2051.9 1952 2054.4 C +1954.5 2054.2 1956.4 2052.3 1958.7 2051.3 C +1960.8 2050 1963.2 2049 1965.6 2048.4 C +1968.3 2050.8 1970.7 2054.3 1973.6 2055.4 C +1973 2052.2 1969.7 2050.4 1967.6 2048.2 C +1967.1 2046.7 1968.8 2046.6 1969.5 2045.8 C +1972.8 2043.3 1980.6 2043.4 1979.3 2038.4 C +1979.4 2038.6 1979.2 2038.7 1979.1 2038.8 C +1978.7 2038.6 1978.9 2038.1 1978.8 2037.6 C +1978.9 2037.9 1978.7 2038 1978.6 2038.1 C +1978.2 2032.7 1978.4 2027.1 1978.4 2021.6 C +1979.3 2021.1 1980 2020.2 1981.5 2020.1 C +1983.5 2020.5 1984 2021.8 1985.1 2023.5 C +1985.7 2024 1987.4 2023.7 1986 2022.8 C +1984.7 2021.7 1983.3 2020.8 1983.9 2018.7 C +1987.2 2015.9 1993 2015.4 1994.9 2011.5 C +1992.2 2004.9 1999.3 2005.2 2002.1 2002.4 C +2005.9 2002.7 2004.8 1997.4 2009.1 1999 C +2011 1999.3 2010 2002.9 2012.7 2002.4 C +2010.2 2000.7 2009.4 1996.1 2005.5 1998.5 C +2002.1 2000.3 1999 2002.5 1995.4 2003.8 C +1995.2 2003.6 1994.9 2003.3 1994.7 2003.1 C +1994.3 1997 1995.6 1991.1 1994.4 1985.3 C +1994.3 1986 1993.8 1985 1994 1985.6 C +1993.8 1995.4 1994.4 2001.6 1993.5 2011.7 C +1989.7 2015.5 1983.6 2014.9 1979.8 2018.7 C +1978.3 2019.1 1979.1 2017.4 1978.4 2017 C +1977.8 2015.5 1979.4 2015.3 1978.8 2014.1 C +1978.4 2014.6 1978.5 2012.8 1977.4 2013.4 C +1975.8 2016.8 1978.5 2019.6 1976.9 2023 C +1977 2028.7 1976.7 2034.5 1977.2 2040 C +1971.6 2043.1 1965.8 2045.6 1960.1 2048.2 C +1957.7 2049.9 1954.8 2050.5 1952.4 2052.3 C +1947.6 2047.4 1941.8 2041.8 1936.8 2037.2 C +1935.2 2033.6 1937.1 2028.8 1936.1 2024.7 C +1935.1 2020.4 1937.1 2018.1 1935.9 2015.1 C +1936.5 2014.1 1934.7 2010.8 1937.1 2009.3 C +1944.4 2004.8 1952 2000.9 1959.9 1997.8 C +1963.9 1997 1963.9 2001.9 1966.8 2003.3 C +1970.3 2006.9 1973.7 2009.9 1976.9 2012.9 C +1977.9 2013 1977.1 2011.4 1976.7 2010.8 C +1971.6 2006.3 1966.8 2000.7 1962 1995.9 C +1960 1995.2 1960.1 1996.6 1958.2 1995.6 C +1957 1997 1955.1 1998.8 1953.2 1998 C +1951.7 1994.5 1954.1 1993.4 1952.9 1991.1 C +1952.1 1990.5 1953.3 1990.2 1953.2 1989.6 C +1954.2 1986.8 1950.9 1981.4 1954.4 1981.2 C +1954.7 1981.6 1954.7 1981.7 1955.1 1982 C +1961.9 1979.1 1967.6 1975 1974.3 1971.6 C +1974.7 1969.8 1976.7 1969.5 1978.4 1969.7 C +1980.3 1970 1979.3 1973.6 1982 1973.1 C +1975.8 1962.2 1968 1975.8 1960.8 1976.7 C +1956.9 1977.4 1953.3 1982.4 1949.1 1978.8 C +1946 1975.8 1941.2 1971 1939.5 1969.2 C +1938.5 1968.6 1938.9 1967.4 1937.8 1966.8 C +1928.7 1969.4 1920.6 1974.5 1912.4 1979.1 C +1904 1980 1896.6 1985 1889.3 1989.4 C +1887.9 1990.4 1885.1 1990.3 1885 1992.5 C +1885.4 2000.6 1885.2 2012.9 1885.2 2019.9 C +1886.1 2022 1889.7 2019.5 1888.4 2022.8 C +1889 2023.3 1889.8 2024.4 1890.3 2024 C +1891.2 2023.5 1891.8 2028.2 1893.4 2026.6 C +1894.2 2026.3 1893.9 2027.3 1894.4 2027.6 C +1893.4 2027.6 1894.7 2028.3 1894.1 2028.5 C +1894.4 2029.6 1896 2030 1896 2029.2 C +1896.2 2029 1896.3 2029 1896.5 2029.2 C +1896.8 2029.8 1897.3 2030 1897 2030.7 C +1896.5 2030.7 1896.9 2031.5 1897.2 2031.6 C +1898.3 2034 1899.5 2030.6 1899.6 2033.3 C +1898.5 2033 1899.6 2034.4 1900.1 2034.8 C +1901.3 2035.8 1903.2 2034.6 1902.5 2036.7 C +1904.4 2036.9 1906.1 2032.2 1907.6 2035.5 C +1907.5 2040.1 1907.7 2044.9 1907.3 2049.4 C +1908 2050.2 1908.3 2051.4 1909.5 2051.6 C +1910.1 2051.1 1911.6 2051.1 1911.4 2052.3 C +1909.7 2052.8 1912.4 2054 1912.6 2054.7 C +1913.4 2055.2 1913 2053.7 1913.6 2054.4 C +1913.6 2054.5 1913.6 2055.3 1913.6 2054.7 C +1913.7 2054.4 1913.9 2054.4 1914 2054.7 C +1914 2054.9 1914.1 2055.3 1913.8 2055.4 C +1913.7 2056 1915.2 2057.6 1916 2057.6 C +1915.9 2057.3 1916.1 2057.2 1916.2 2057.1 C +1917 2056.8 1916.7 2057.7 1917.2 2058 C +1917 2058.3 1916.7 2058.3 1916.4 2058.3 C +1917.1 2059 1917.3 2060.1 1918.4 2060.4 C +1918.1 2059.2 1919.1 2060.6 1919.1 2059.5 C +1919 2060.6 1920.6 2060.1 1919.8 2061.2 C +1919.6 2061.2 1919.3 2061.2 1919.1 2061.2 C +1919.6 2061.9 1921.4 2064.2 1921.5 2062.6 C +1922.4 2062.1 1921.6 2063.9 1922.2 2064.3 C +1922.9 2067.3 1926.1 2064.3 1925.6 2067.2 C +1927.2 2066.8 1928.4 2064.6 1930.1 2065.2 C +1931.8 2067.8 1931 2071.8 1930.8 2074.8 C +1930.6 2076.4 1930.1 2078.6 1930.6 2080.4 C +1936.6 2085.4 1941.8 2091.6 1948.1 2096.9 C +1950.7 2096.7 1952.6 2094.8 1954.8 2093.8 C +[0 0.33 0.33 0.99] vc +f +S +n +1989.4 2080.6 m +1996.1 2077.3 2002.7 2073.8 2009.6 2070.8 C +2009.6 2069.6 2010.2 2067.7 2009.3 2067.4 C +2008.9 2062 2009.1 2056.4 2009.1 2050.8 C +2012.3 2046.6 2019 2046.6 2023.5 2043.2 C +2024 2042.3 2025.1 2042.1 2025.4 2041.2 C +2025.3 2032.7 2025.6 2023.1 2025.2 2014.6 C +2025 2015.3 2024.6 2014.2 2024.7 2014.8 C +2024.5 2024.7 2025.1 2030.9 2024.2 2041 C +2020.4 2044.8 2014.3 2044.2 2010.5 2048 C +2009 2048.4 2009.8 2046.7 2009.1 2046.3 C +2008.5 2044.8 2010.2 2044.6 2009.6 2043.4 C +2009.1 2043.9 2009.2 2042.1 2008.1 2042.7 C +2006.5 2046.1 2009.3 2048.9 2007.6 2052.3 C +2007.7 2058 2007.5 2063.8 2007.9 2069.3 C +2002.3 2072.4 1996.5 2074.8 1990.8 2077.5 C +1988.4 2079.2 1985.6 2079.8 1983.2 2081.6 C +1980.5 2079 1977.9 2076.5 1975.5 2074.1 C +1975.5 2075.1 1975.5 2076.2 1975.5 2077.2 C +1977.8 2079.3 1980.3 2081.6 1982.7 2083.7 C +1985.3 2083.5 1987.1 2081.6 1989.4 2080.6 C +f +S +n +1930.1 2079.9 m +1931.1 2075.6 1929.2 2071.1 1930.8 2067.2 C +1930.3 2066.3 1930.1 2064.6 1928.7 2065.5 C +1927.7 2066.4 1926.5 2067 1925.3 2067.4 C +1924.5 2066.9 1925.6 2065.7 1924.4 2066 C +1924.2 2067.2 1923.6 2065.5 1923.2 2065.7 C +1922.3 2063.6 1917.8 2062.1 1919.6 2060.4 C +1919.3 2060.5 1919.2 2060.3 1919.1 2060.2 C +1919.7 2060.9 1918.2 2061 1917.6 2060.2 C +1917 2059.6 1916.1 2058.8 1916.4 2058 C +1915.5 2058 1917.4 2057.1 1915.7 2057.8 C +1914.8 2057.1 1913.4 2056.2 1913.3 2054.9 C +1913.1 2055.4 1911.3 2054.3 1910.9 2053.2 C +1910.7 2052.9 1910.2 2052.5 1910.7 2052.3 C +1911.1 2052.5 1910.9 2052 1910.9 2051.8 C +1910.5 2051.2 1909.9 2052.6 1909.2 2051.8 C +1908.2 2051.4 1907.8 2050.2 1907.1 2049.4 C +1907.5 2044.8 1907.3 2040 1907.3 2035.2 C +1905.3 2033 1902.8 2039.3 1902.3 2035.7 C +1899.6 2036 1898.4 2032.5 1896.3 2030.7 C +1895.7 2030.1 1897.5 2030 1896.3 2029.7 C +1896.3 2030.6 1895 2029.7 1894.4 2029.2 C +1892.9 2028.1 1894.2 2027.4 1893.6 2027.1 C +1892.1 2027.9 1891.7 2025.6 1890.8 2024.9 C +1891.1 2024.6 1889.1 2024.3 1888.4 2023 C +1887.5 2022.6 1888.2 2021.9 1888.1 2021.3 C +1886.7 2022 1885.2 2020.4 1884.8 2019.2 C +1884.8 2010 1884.6 2000.2 1885 1991.8 C +1886.9 1989.6 1889.9 1989.3 1892.2 1987.5 C +1898.3 1982.7 1905.6 1980.1 1912.8 1978.6 C +1921 1974.2 1928.8 1968.9 1937.8 1966.6 C +1939.8 1968.3 1938.8 1968.3 1940.4 1970 C +1945.4 1972.5 1947.6 1981.5 1954.6 1979.3 C +1952.3 1981 1950.4 1978.4 1948.6 1977.9 C +1945.1 1973.9 1941.1 1970.6 1938 1966.6 C +1928.4 1968.5 1920.6 1974.8 1911.9 1978.8 C +1907.1 1979.2 1902.6 1981.7 1898.2 1983.6 C +1893.9 1986 1889.9 1989 1885.5 1990.8 C +1884.9 1991.2 1884.8 1991.8 1884.5 1992.3 C +1884.9 2001.3 1884.7 2011.1 1884.8 2019.6 C +1890.6 2025 1896.5 2031.2 1902.3 2036.9 C +1904.6 2037.6 1905 2033 1907.3 2035.5 C +1907.2 2040.2 1907 2044.8 1907.1 2049.6 C +1913.6 2055.3 1918.4 2061.5 1925.1 2067.4 C +1927.3 2068.2 1929.6 2062.5 1930.6 2066.9 C +1929.7 2070.7 1930.3 2076 1930.1 2080.1 C +1935.6 2085.7 1941.9 2090.7 1947.2 2096.7 C +1942.2 2091.1 1935.5 2085.2 1930.1 2079.9 C +[0.18 0.18 0 0.78] vc +f +S +n +1930.8 2061.9 m +1930.3 2057.8 1931.8 2053.4 1931.1 2050.4 C +1931.3 2050.3 1931.7 2050.5 1931.6 2050.1 C +1933 2051.1 1934.4 2049.5 1935.9 2048.7 C +1937 2046.5 1939.5 2047.1 1941.2 2045.1 C +1939.7 2042.6 1937.3 2041.2 1935.4 2039.3 C +1934 2039.7 1934.5 2038.1 1933.7 2037.6 C +1934 2033.3 1933.1 2027.9 1934.4 2024.4 C +1934.3 2023.8 1933.9 2022.8 1933 2022.8 C +1931.6 2023.1 1930.5 2024.4 1929.2 2024.9 C +1928.4 2024.5 1929.8 2023.5 1928.7 2023.5 C +1927.7 2024.1 1926.2 2022.6 1925.6 2021.6 C +1926.9 2021.6 1924.8 2020.6 1925.6 2020.4 C +1924.7 2021.7 1923.9 2019.6 1923.2 2019.2 C +1923.3 2018.3 1923.8 2018.1 1923.2 2018 C +1922.9 2017.8 1922.9 2017.5 1922.9 2017.2 C +1922.8 2018.3 1921.3 2017.3 1920.3 2018 C +1916.6 2019.7 1913 2022.1 1910 2024.7 C +1910 2032.9 1910 2041.2 1910 2049.4 C +1915.4 2055.2 1920 2058.7 1925.3 2064.8 C +1927.2 2064 1929 2061.4 1930.8 2061.9 C +[0 0 0 0] vc +f +S +n +1907.6 2030.4 m +1907.5 2027.1 1906.4 2021.7 1908.5 2019.9 C +1908.8 2020.1 1908.9 2019 1909.2 2019.6 C +1910 2019.6 1912 2019.2 1913.1 2018.2 C +1913.7 2016.5 1920.2 2015.7 1917.4 2012.7 C +1918.2 2011.2 1917 2013.8 1917.2 2012 C +1916.9 2012.3 1916 2012.4 1915.2 2012 C +1912.5 2010.5 1916.6 2008.8 1913.6 2009.6 C +1912.6 2009.2 1911.1 2009 1910.9 2007.6 C +1911 1999.2 1911.8 1989.8 1911.2 1982.2 C +1910.1 1981.1 1908.8 1982.2 1907.6 1982.2 C +1900.8 1986.5 1893.2 1988.8 1887.2 1994.2 C +1887.2 2002.4 1887.2 2010.7 1887.2 2018.9 C +1892.6 2024.7 1897.2 2028.2 1902.5 2034.3 C +1904.3 2033.3 1906.2 2032.1 1907.6 2030.4 C +f +S +n +1910.7 2025.4 m +1912.7 2022.4 1916.7 2020.8 1919.8 2018.9 C +1920.2 2018.7 1920.6 2018.6 1921 2018.4 C +1925 2020 1927.4 2028.5 1932 2024.2 C +1932.3 2025 1932.5 2023.7 1932.8 2024.4 C +1932.8 2028 1932.8 2031.5 1932.8 2035 C +1931.9 2033.9 1932.5 2036.3 1932.3 2036.9 C +1933.2 2036.4 1932.5 2038.5 1933 2038.4 C +1933.1 2040.5 1935.6 2042.2 1936.6 2043.2 C +1936.2 2042.4 1935.1 2040.8 1933.7 2040.3 C +1932.2 2034.4 1933.8 2029.8 1933 2023.2 C +1931.1 2024.9 1928.4 2026.4 1926.5 2023.5 C +1925.1 2021.6 1923 2019.8 1921.5 2018.2 C +1917.8 2018.9 1915.2 2022.5 1911.6 2023.5 C +1910.8 2023.8 1911.2 2024.7 1910.4 2025.2 C +1910.9 2031.8 1910.6 2039.1 1910.7 2045.6 C +1910.1 2048 1910.7 2045.9 1911.2 2044.8 C +1910.6 2038.5 1911.2 2031.8 1910.7 2025.4 C +[0.07 0.06 0 0.58] vc +f +S +n +1910.7 2048.9 m +1910.3 2047.4 1911.3 2046.5 1911.6 2045.3 C +1912.9 2045.3 1913.9 2047.1 1915.2 2045.8 C +1915.2 2044.9 1916.6 2043.3 1917.2 2042.9 C +1918.7 2042.9 1919.4 2044.4 1920.5 2043.2 C +1921.2 2042.2 1921.4 2040.9 1922.4 2040.3 C +1924.5 2040.3 1925.7 2040.9 1926.8 2039.6 C +1927.1 2037.9 1926.8 2038.1 1927.7 2037.6 C +1929 2037.5 1930.4 2037 1931.6 2037.2 C +1932.3 2038.2 1933.1 2038.7 1932.8 2040.3 C +1935 2041.8 1935.9 2043.8 1938.5 2044.8 C +1938.6 2045 1938.3 2045.5 1938.8 2045.3 C +1939.1 2042.9 1935.4 2044.2 1935.4 2042.2 C +1932.1 2040.8 1932.8 2037.2 1932 2034.8 C +1932.3 2034 1932.7 2035.4 1932.5 2034.8 C +1931.3 2031.8 1935.5 2020.1 1928.9 2025.9 C +1924.6 2024.7 1922.6 2014.5 1917.4 2020.4 C +1915.5 2022.8 1912 2022.6 1910.9 2025.4 C +1911.5 2031.9 1910.9 2038.8 1911.4 2045.3 C +1911.1 2046.5 1910 2047.4 1910.4 2048.9 C +1915.1 2054.4 1920.4 2058.3 1925.1 2063.8 C +1920.8 2058.6 1914.9 2054.3 1910.7 2048.9 C +[0.4 0.4 0 0] vc +f +S +n +1934.7 2031.9 m +1934.6 2030.7 1934.9 2029.5 1934.4 2028.5 C +1934 2029.5 1934.3 2031.2 1934.2 2032.6 C +1933.8 2031.7 1934.9 2031.6 1934.7 2031.9 C +[0.92 0.92 0 0.67] vc +f +S +n +vmrs +1934.7 2019.4 m +1934.1 2015.3 1935.6 2010.9 1934.9 2007.9 C +1935.1 2007.8 1935.6 2008.1 1935.4 2007.6 C +1936.8 2008.6 1938.2 2007 1939.7 2006.2 C +1940.1 2004.3 1942.7 2005 1943.6 2003.8 C +1945.1 2000.3 1954 2000.8 1950 1996.6 C +1952.1 1993.3 1948.2 1989.2 1951.2 1985.6 C +1953 1981.4 1948.4 1982.3 1947.9 1979.8 C +1945.4 1979.6 1945.1 1975.5 1942.4 1975 C +1942.4 1972.3 1938 1973.6 1938.5 1970.4 C +1937.4 1969 1935.6 1970.1 1934.2 1970.2 C +1927.5 1974.5 1919.8 1976.8 1913.8 1982.2 C +1913.8 1990.4 1913.8 1998.7 1913.8 2006.9 C +1919.3 2012.7 1923.8 2016.2 1929.2 2022.3 C +1931.1 2021.6 1932.8 2018.9 1934.7 2019.4 C +[0 0 0 0] vc +f +0.4 w +2 J +2 M +S +n +2024.2 2038.1 m +2024.1 2029.3 2024.4 2021.7 2024.7 2014.4 C +2024.4 2013.6 2020.6 2013.4 2021.3 2011.2 C +2020.5 2010.3 2018.4 2010.6 2018.9 2008.6 C +2019 2008.8 2018.8 2009 2018.7 2009.1 C +2018.2 2006.7 2015.2 2007.9 2015.3 2005.5 C +2014.7 2004.8 2012.4 2005.1 2013.2 2003.6 C +2012.3 2004.2 2012.8 2002.4 2012.7 2002.6 C +2009.4 2003.3 2011.2 1998.6 2008.4 1999.2 C +2007 1999.1 2006.1 1999.4 2005.7 2000.4 C +2006.9 1998.5 2007.7 2000.5 2009.3 2000.2 C +2009.2 2003.7 2012.4 2002.1 2012.9 2005.2 C +2015.9 2005.6 2015.2 2008.6 2017.7 2008.8 C +2018.4 2009.6 2018.3 2011.4 2019.6 2011 C +2021.1 2011.7 2021.4 2014.8 2023.7 2015.1 C +2023.7 2023.5 2023.9 2031.6 2023.5 2040.5 C +2021.8 2041.7 2020.7 2043.6 2018.4 2043.9 C +2020.8 2042.7 2025.5 2041.8 2024.2 2038.1 C +[0 0.87 0.91 0.83] vc +f +S +n +2023.5 2040 m +2023.5 2031.1 2023.5 2023.4 2023.5 2015.1 C +2020.2 2015 2021.8 2010.3 2018.4 2011 C +2018.6 2007.5 2014.7 2009.3 2014.8 2006.4 C +2011.8 2006.3 2012.2 2002.3 2009.8 2002.4 C +2009.7 2001.5 2009.2 2000.1 2008.4 2000.2 C +2008.7 2000.9 2009.7 2001.2 2009.3 2002.4 C +2008.4 2004.2 2007.5 2003.1 2007.9 2005.5 C +2007.9 2010.8 2007.7 2018.7 2008.1 2023.2 C +2009 2024.3 2007.3 2023.4 2007.9 2024 C +2007.7 2024.6 2007.3 2026.3 2008.6 2027.1 C +2009.7 2026.8 2010 2027.6 2010.5 2028 C +2010.5 2028.2 2010.5 2029.1 2010.5 2028.5 C +2011.5 2028 2010.5 2030 2011.5 2030 C +2014.2 2029.7 2012.9 2032.2 2014.8 2032.6 C +2015.1 2033.6 2015.3 2033 2016 2033.3 C +2017 2033.9 2016.6 2035.4 2017.2 2036.2 C +2018.7 2036.4 2019.2 2039 2021.3 2038.4 C +2021.6 2035.4 2019.7 2029.5 2021.1 2027.3 C +2020.9 2023.5 2021.5 2018.5 2020.6 2016 C +2020.9 2013.9 2021.5 2015.4 2022.3 2014.4 C +2022.2 2015.1 2023.3 2014.8 2023.2 2015.6 C +2022.7 2019.8 2023.3 2024.3 2022.8 2028.5 C +2022.3 2028.2 2022.6 2027.6 2022.5 2027.1 C +2022.5 2027.8 2022.5 2029.2 2022.5 2029.2 C +2022.6 2029.2 2022.7 2029.1 2022.8 2029 C +2023.9 2032.8 2022.6 2037 2023 2040.8 C +2022.3 2041.2 2021.6 2041.5 2021.1 2042.2 C +2022 2041.2 2022.9 2041.4 2023.5 2040 C +[0 1 1 0.23] vc +f +S +n +2009.1 1997.8 m +2003.8 1997.7 2000.1 2002.4 1995.4 2003.1 C +1995 1999.5 1995.2 1995 1995.2 1992 C +1995.2 1995.8 1995 1999.7 1995.4 2003.3 C +2000.3 2002.2 2003.8 1997.9 2009.1 1997.8 C +2012.3 2001.2 2015.6 2004.8 2018.7 2008.1 C +2021.6 2011.2 2027.5 2013.9 2025.9 2019.9 C +2026.1 2017.9 2025.6 2016.2 2025.4 2014.4 C +2020.2 2008.4 2014 2003.6 2009.1 1997.8 C +[0.18 0.18 0 0.78] vc +f +S +n +2009.3 1997.8 m +2008.7 1997.4 2007.9 1997.6 2007.2 1997.6 C +2007.9 1997.6 2008.9 1997.4 2009.6 1997.8 C +2014.7 2003.6 2020.8 2008.8 2025.9 2014.8 C +2025.8 2017.7 2026.1 2014.8 2025.6 2014.1 C +2020.4 2008.8 2014.8 2003.3 2009.3 1997.8 C +[0.07 0.06 0 0.58] vc +f +S +n +2009.6 1997.6 m +2009 1997.1 2008.1 1997.4 2007.4 1997.3 C +2008.1 1997.4 2009 1997.1 2009.6 1997.6 C +2014.8 2003.7 2021.1 2008.3 2025.9 2014.4 C +2021.1 2008.3 2014.7 2003.5 2009.6 1997.6 C +[0.4 0.4 0 0] vc +f +S +n +2021.8 2011.5 m +2021.9 2012.2 2022.3 2013.5 2023.7 2013.6 C +2023.4 2012.7 2022.8 2011.8 2021.8 2011.5 C +[0 0.33 0.33 0.99] vc +f +S +n +2021.1 2042 m +2022.1 2041.1 2020.9 2040.2 2020.6 2039.6 C +2018.4 2039.5 2018.1 2036.9 2016.3 2036.4 C +2015.8 2035.5 2015.3 2033.8 2014.8 2033.6 C +2012.4 2033.8 2013 2030.4 2010.5 2030.2 C +2009.6 2028.9 2009.6 2028.3 2008.4 2028 C +2006.9 2026.7 2007.5 2024.3 2006 2023.2 C +2006.6 2023.2 2005.7 2023.3 2005.7 2023 C +2006.4 2022.5 2006.3 2021.1 2006.7 2020.6 C +2006.6 2015 2006.9 2009 2006.4 2003.8 C +2006.9 2002.5 2007.6 2001.1 2006.9 2000.7 C +2004.6 2003.6 2003 2002.9 2000.2 2004.3 C +1999.3 2005.8 1997.9 2006.3 1996.1 2006.7 C +1995.7 2008.9 1996 2011.1 1995.9 2012.9 C +1993.4 2015.1 1990.5 2016.2 1987.7 2017.7 C +1987.1 2019.3 1991.1 2019.4 1990.4 2021.3 C +1990.5 2021.5 1991.9 2022.3 1992 2023 C +1994.8 2024.4 1996.2 2027.5 1998.5 2030 C +2002.4 2033 2005.2 2037.2 2008.8 2041 C +2010.2 2041.3 2011.6 2042 2011 2043.9 C +2011.2 2044.8 2010.1 2045.3 2010.5 2046.3 C +2013.8 2044.8 2017.5 2043.4 2021.1 2042 C +[0 0.5 0.5 0.2] vc +f +S +n +2019.4 2008.8 m +2018.9 2009.2 2019.3 2009.9 2019.6 2010.3 C +2022.2 2011.5 2020.3 2009.1 2019.4 2008.8 C +[0 0.33 0.33 0.99] vc +f +S +n +2018 2007.4 m +2015.7 2006.7 2015.3 2003.6 2012.9 2002.8 C +2013.5 2003.7 2013.5 2005.1 2015.6 2005.2 C +2016.4 2006.1 2015.7 2007.7 2018 2007.4 C +f +S +n +vmrs +1993.5 2008.8 m +1993.4 2000 1993.7 1992.5 1994 1985.1 C +1993.7 1984.3 1989.9 1984.1 1990.6 1982 C +1989.8 1981.1 1987.7 1981.4 1988.2 1979.3 C +1988.3 1979.6 1988.1 1979.7 1988 1979.8 C +1987.5 1977.5 1984.5 1978.6 1984.6 1976.2 C +1983.9 1975.5 1981.7 1975.8 1982.4 1974.3 C +1981.6 1974.9 1982.1 1973.1 1982 1973.3 C +1979 1973.7 1980 1968.8 1976.9 1969.7 C +1975.9 1969.8 1975.3 1970.3 1975 1971.2 C +1976.2 1969.2 1977 1971.2 1978.6 1970.9 C +1978.5 1974.4 1981.7 1972.8 1982.2 1976 C +1985.2 1976.3 1984.5 1979.3 1987 1979.6 C +1987.7 1980.3 1987.5 1982.1 1988.9 1981.7 C +1990.4 1982.4 1990.7 1985.5 1993 1985.8 C +1992.9 1994.3 1993.2 2002.3 1992.8 2011.2 C +1991.1 2012.4 1990 2014.4 1987.7 2014.6 C +1990.1 2013.4 1994.7 2012.6 1993.5 2008.8 C +[0 0.87 0.91 0.83] vc +f +0.4 w +2 J +2 M +S +n +1992.8 2010.8 m +1992.8 2001.8 1992.8 1994.1 1992.8 1985.8 C +1989.5 1985.7 1991.1 1981.1 1987.7 1981.7 C +1987.9 1978.2 1983.9 1980 1984.1 1977.2 C +1981.1 1977 1981.5 1973 1979.1 1973.1 C +1979 1972.2 1978.5 1970.9 1977.6 1970.9 C +1977.9 1971.6 1979 1971.9 1978.6 1973.1 C +1977.6 1974.9 1976.8 1973.9 1977.2 1976.2 C +1977.2 1981.5 1977 1989.4 1977.4 1994 C +1978.3 1995 1976.6 1994.1 1977.2 1994.7 C +1977 1995.3 1976.6 1997 1977.9 1997.8 C +1979 1997.5 1979.3 1998.3 1979.8 1998.8 C +1979.8 1998.9 1979.8 1999.8 1979.8 1999.2 C +1980.8 1998.7 1979.7 2000.7 1980.8 2000.7 C +1983.5 2000.4 1982.1 2003 1984.1 2003.3 C +1984.4 2004.3 1984.5 2003.7 1985.3 2004 C +1986.3 2004.6 1985.9 2006.1 1986.5 2006.9 C +1988 2007.1 1988.4 2009.7 1990.6 2009.1 C +1990.9 2006.1 1989 2000.2 1990.4 1998 C +1990.2 1994.3 1990.8 1989.2 1989.9 1986.8 C +1990.2 1984.7 1990.8 1986.2 1991.6 1985.1 C +1991.5 1985.9 1992.6 1985.5 1992.5 1986.3 C +1992 1990.5 1992.6 1995 1992 1999.2 C +1991.6 1998.9 1991.9 1998.3 1991.8 1997.8 C +1991.8 1998.5 1991.8 2000 1991.8 2000 C +1991.9 1999.9 1992 1999.8 1992 1999.7 C +1993.2 2003.5 1991.9 2007.7 1992.3 2011.5 C +1991.6 2012 1990.9 2012.2 1990.4 2012.9 C +1991.3 2011.9 1992.2 2012.1 1992.8 2010.8 C +[0 1 1 0.23] vc +f +S +n +1978.4 1968.5 m +1977 1969.2 1975.8 1968.2 1974.5 1969 C +1968.3 1973 1961.6 1976 1955.1 1979.1 C +1962 1975.9 1968.8 1972.5 1975.5 1968.8 C +1976.5 1968.8 1977.6 1968.8 1978.6 1968.8 C +1981.7 1972.1 1984.8 1975.7 1988 1978.8 C +1990.9 1981.9 1996.8 1984.6 1995.2 1990.6 C +1995.3 1988.6 1994.9 1986.9 1994.7 1985.1 C +1989.5 1979.1 1983.3 1974.3 1978.4 1968.5 C +[0.18 0.18 0 0.78] vc +f +S +n +1978.4 1968.3 m +1977.9 1968.7 1977.1 1968.5 1976.4 1968.5 C +1977.3 1968.8 1978.1 1967.9 1978.8 1968.5 C +1984 1974.3 1990.1 1979.5 1995.2 1985.6 C +1995.1 1988.4 1995.3 1985.6 1994.9 1984.8 C +1989.5 1979.4 1983.9 1973.8 1978.4 1968.3 C +[0.07 0.06 0 0.58] vc +f +S +n +1978.6 1968 m +1977.9 1968 1977.4 1968.6 1978.4 1968 C +1983.9 1973.9 1990.1 1979.1 1995.2 1985.1 C +1990.2 1979 1983.8 1974.1 1978.6 1968 C +[0.4 0.4 0 0] vc +f +S +n +1991.1 1982.2 m +1991.2 1982.9 1991.6 1984.2 1993 1984.4 C +1992.6 1983.5 1992.1 1982.5 1991.1 1982.2 C +[0 0.33 0.33 0.99] vc +f +S +n +1990.4 2012.7 m +1991.4 2011.8 1990.2 2010.9 1989.9 2010.3 C +1987.7 2010.2 1987.4 2007.6 1985.6 2007.2 C +1985.1 2006.2 1984.6 2004.5 1984.1 2004.3 C +1981.7 2004.5 1982.3 2001.2 1979.8 2000.9 C +1978.8 1999.6 1978.8 1999.1 1977.6 1998.8 C +1976.1 1997.4 1976.7 1995 1975.2 1994 C +1975.8 1994 1975 1994 1975 1993.7 C +1975.7 1993.2 1975.6 1991.8 1976 1991.3 C +1975.9 1985.7 1976.1 1979.7 1975.7 1974.5 C +1976.2 1973.3 1976.9 1971.8 1976.2 1971.4 C +1973.9 1974.3 1972.2 1973.6 1969.5 1975 C +1967.9 1977.5 1963.8 1977.1 1961.8 1980 C +1959 1980 1957.6 1983 1954.8 1982.9 C +1953.8 1984.2 1954.8 1985.7 1955.1 1987.2 C +1956.2 1989.5 1959.7 1990.1 1959.9 1991.8 C +1965.9 1998 1971.8 2005.2 1978.1 2011.7 C +1979.5 2012 1980.9 2012.7 1980.3 2014.6 C +1980.5 2015.6 1979.4 2016 1979.8 2017 C +1983 2015.6 1986.8 2014.1 1990.4 2012.7 C +[0 0.5 0.5 0.2] vc +f +S +n +1988.7 1979.6 m +1988.2 1979.9 1988.6 1980.6 1988.9 1981 C +1991.4 1982.2 1989.6 1979.9 1988.7 1979.6 C +[0 0.33 0.33 0.99] vc +f +S +n +1987.2 1978.1 m +1985 1977.5 1984.6 1974.3 1982.2 1973.6 C +1982.7 1974.5 1982.8 1975.8 1984.8 1976 C +1985.7 1976.9 1985 1978.4 1987.2 1978.1 C +f +S +n +1975.5 2084 m +1975.5 2082 1975.3 2080 1975.7 2078.2 C +1978.8 2079 1980.9 2085.5 1984.8 2083.5 C +1993 2078.7 2001.6 2075 2010 2070.8 C +2010.1 2064 2009.9 2057.2 2010.3 2050.6 C +2014.8 2046.2 2020.9 2045.7 2025.6 2042 C +2026.1 2035.1 2025.8 2028 2025.9 2021.1 C +2025.8 2027.8 2026.1 2034.6 2025.6 2041.2 C +2022.2 2044.9 2017.6 2046.8 2012.9 2048 C +2012.5 2049.5 2010.4 2049.4 2009.8 2051.1 C +2009.9 2057.6 2009.6 2064.2 2010 2070.5 C +2001.2 2075.4 1992 2079.1 1983.2 2084 C +1980.3 2082.3 1977.8 2079.2 1975.2 2077.5 C +1974.9 2079.9 1977.2 2084.6 1973.3 2085.2 C +1964.7 2088.6 1956.8 2093.7 1948.1 2097.2 C +1949 2097.3 1949.6 2096.9 1950.3 2096.7 C +1958.4 2091.9 1967.1 2088.2 1975.5 2084 C +[0.18 0.18 0 0.78] vc +f +S +n +vmrs +1948.6 2094.5 m +1950.2 2093.7 1951.8 2092.9 1953.4 2092.1 C +1951.8 2092.9 1950.2 2093.7 1948.6 2094.5 C +[0 0.87 0.91 0.83] vc +f +0.4 w +2 J +2 M +S +n +1971.6 2082.3 m +1971.6 2081.9 1970.7 2081.1 1970.9 2081.3 C +1970.7 2081.6 1970.6 2081.6 1970.4 2081.3 C +1970.8 2080.1 1968.7 2081.7 1968.3 2080.8 C +1966.6 2080.9 1966.7 2078 1964.2 2078.2 C +1964.8 2075 1960.1 2075.8 1960.1 2072.9 C +1958 2072.3 1957.5 2069.3 1955.3 2069.3 C +1953.9 2070.9 1948.8 2067.8 1950 2072 C +1949 2074 1943.2 2070.6 1944 2074.8 C +1942.2 2076.6 1937.6 2073.9 1938 2078.2 C +1936.7 2078.6 1935 2078.6 1933.7 2078.2 C +1933.5 2080 1936.8 2080.7 1937.3 2082.8 C +1939.9 2083.5 1940.6 2086.4 1942.6 2088 C +1945.2 2089.2 1946 2091.3 1948.4 2093.6 C +1956 2089.5 1963.9 2086.1 1971.6 2082.3 C +[0 0.01 1 0] vc +f +S +n +1958.2 2089.7 m +1956.4 2090 1955.6 2091.3 1953.9 2091.9 C +1955.6 2091.9 1956.5 2089.7 1958.2 2089.7 C +[0 0.87 0.91 0.83] vc +f +S +n +1929.9 2080.4 m +1929.5 2077.3 1929.7 2073.9 1929.6 2070.8 C +1929.8 2074.1 1929.2 2077.8 1930.1 2080.8 C +1935.8 2085.9 1941.4 2091.3 1946.9 2096.9 C +1941.2 2091 1935.7 2086 1929.9 2080.4 C +[0.4 0.4 0 0] vc +f +S +n +1930.1 2080.4 m +1935.8 2086 1941.5 2090.7 1946.9 2096.7 C +1941.5 2090.9 1935.7 2085.8 1930.1 2080.4 C +[0.07 0.06 0 0.58] vc +f +S +n +1940.9 2087.1 m +1941.7 2088 1944.8 2090.6 1943.6 2089.2 C +1942.5 2089 1941.6 2087.7 1940.9 2087.1 C +[0 0.87 0.91 0.83] vc +f +S +n +1972.8 2082.8 m +1973 2075.3 1972.4 2066.9 1973.3 2059.5 C +1972.5 2058.9 1972.8 2057.3 1973.1 2056.4 C +1974.8 2055.2 1973.4 2055.5 1972.4 2055.4 C +1970.1 2053.2 1967.9 2050.9 1965.6 2048.7 C +1960.9 2049.9 1956.9 2052.7 1952.4 2054.7 C +1949.3 2052.5 1946.3 2049.5 1943.6 2046.8 C +1939.9 2047.7 1936.8 2050.1 1933.5 2051.8 C +1930.9 2054.9 1933.5 2056.2 1932.3 2059.7 C +1933.2 2059.7 1932.2 2060.5 1932.5 2060.2 C +1933.2 2062.5 1931.6 2064.6 1932.5 2067.4 C +1932.9 2069.7 1932.7 2072.2 1932.8 2074.6 C +1933.6 2070.6 1932.2 2066.3 1933 2062.6 C +1934.4 2058.2 1929.8 2053.5 1935.2 2051.1 C +1937.7 2049.7 1940.2 2048 1942.8 2046.8 C +1945.9 2049.2 1948.8 2052 1951.7 2054.7 C +1952.7 2054.7 1953.6 2054.6 1954.4 2054.2 C +1958.1 2052.5 1961.7 2049.3 1965.9 2049.2 C +1968.2 2052.8 1975.2 2055 1972.6 2060.9 C +1973.3 2062.4 1972.2 2065.2 1972.6 2067.6 C +1972.7 2072.6 1972.4 2077.7 1972.8 2082.5 C +1968.1 2084.9 1963.5 2087.5 1958.7 2089.5 C +1963.5 2087.4 1968.2 2085 1972.8 2082.8 C +f +S +n +1935.2 2081.1 m +1936.8 2083.4 1938.6 2084.6 1940.4 2086.6 C +1938.8 2084.4 1936.7 2083.4 1935.2 2081.1 C +f +S +n +1983.2 2081.3 m +1984.8 2080.5 1986.3 2079.7 1988 2078.9 C +1986.3 2079.7 1984.8 2080.5 1983.2 2081.3 C +f +S +n +2006.2 2069.1 m +2006.2 2068.7 2005.2 2067.9 2005.5 2068.1 C +2005.3 2068.4 2005.2 2068.4 2005 2068.1 C +2005.4 2066.9 2003.3 2068.5 2002.8 2067.6 C +2001.2 2067.7 2001.2 2064.8 1998.8 2065 C +1999.4 2061.8 1994.7 2062.6 1994.7 2059.7 C +1992.4 2059.5 1992.4 2055.8 1990.1 2056.8 C +1985.9 2059.5 1981.1 2061 1976.9 2063.8 C +1977.2 2067.6 1974.9 2074.2 1978.8 2075.8 C +1979.6 2077.8 1981.7 2078.4 1982.9 2080.4 C +1990.6 2076.3 1998.5 2072.9 2006.2 2069.1 C +[0 0.01 1 0] vc +f +S +n +vmrs +1992.8 2076.5 m +1991 2076.8 1990.2 2078.1 1988.4 2078.7 C +1990.2 2078.7 1991 2076.5 1992.8 2076.5 C +[0 0.87 0.91 0.83] vc +f +0.4 w +2 J +2 M +S +n +1975.5 2073.4 m +1976.1 2069.7 1973.9 2064.6 1977.4 2062.4 C +1973.9 2064.5 1976.1 2069.9 1975.5 2073.6 C +1976 2074.8 1979.3 2077.4 1978.1 2076 C +1977 2075.7 1975.8 2074.5 1975.5 2073.4 C +f +S +n +2007.4 2069.6 m +2007.6 2062.1 2007 2053.7 2007.9 2046.3 C +2007.1 2045.7 2007.3 2044.1 2007.6 2043.2 C +2009.4 2042 2007.9 2042.3 2006.9 2042.2 C +2002.2 2037.4 1996.7 2032.4 1992.5 2027.3 C +1992 2027.3 1991.6 2027.3 1991.1 2027.3 C +1991.4 2035.6 1991.4 2045.6 1991.1 2054.4 C +1990.5 2055.5 1988.4 2056.6 1990.6 2055.4 C +1991.6 2055.4 1991.6 2054.1 1991.6 2053.2 C +1990.8 2044.7 1991.9 2035.4 1991.6 2027.6 C +1991.8 2027.6 1992 2027.6 1992.3 2027.6 C +1997 2032.8 2002.5 2037.7 2007.2 2042.9 C +2007.3 2044.8 2006.7 2047.4 2007.6 2048.4 C +2006.9 2055.1 2007.1 2062.5 2007.4 2069.3 C +2002.7 2071.7 1998.1 2074.3 1993.2 2076.3 C +1998 2074.2 2002.7 2071.8 2007.4 2069.6 C +f +S +n +2006.7 2069.1 m +2006.3 2068.6 2005.9 2067.7 2005.7 2066.9 C +2005.7 2059.7 2005.9 2051.4 2005.5 2045.1 C +2004.9 2045.3 2004.7 2044.5 2004.3 2045.3 C +2005.1 2045.3 2004.2 2045.8 2004.8 2046 C +2004.8 2052.2 2004.8 2059.2 2004.8 2064.5 C +2005.7 2065.7 2005.1 2065.7 2005 2066.7 C +2003.8 2067 2002.7 2067.2 2001.9 2066.4 C +2001.3 2064.6 1998 2063.1 1998 2061.9 C +1996.1 2062.3 1996.6 2058.3 1994.2 2058.8 C +1992.6 2057.7 1992.7 2054.8 1989.9 2056.6 C +1985.6 2059.3 1980.9 2060.8 1976.7 2063.6 C +1976 2066.9 1976 2071.2 1976.7 2074.6 C +1977.6 2070.8 1973.1 2062.1 1980.5 2061.2 C +1984.3 2060.3 1987.5 2058.2 1990.8 2056.4 C +1991.7 2056.8 1992.9 2057.2 1993.5 2059.2 C +1994.3 2058.6 1994.4 2060.6 1994.7 2059.2 C +1995.3 2062.7 1999.2 2061.4 1998.8 2064.8 C +2001.8 2065.4 2002.5 2068.4 2005.2 2067.4 C +2004.9 2067.9 2006 2068 2006.4 2069.1 C +2001.8 2071.1 1997.4 2073.9 1992.8 2075.8 C +1997.5 2073.8 2002 2071.2 2006.7 2069.1 C +[0 0.2 1 0] vc +f +S +n +1988.7 2056.6 m +1985.1 2058.7 1981.1 2060.1 1977.6 2061.9 C +1981.3 2060.5 1985.6 2058.1 1988.7 2056.6 C +[0 0.87 0.91 0.83] vc +f +S +n +1977.9 2059.5 m +1975.7 2064.5 1973.7 2054.7 1975.2 2060.9 C +1976 2060.6 1977.6 2059.7 1977.9 2059.5 C +f +S +n +1989.6 2051.3 m +1990.1 2042.3 1989.8 2036.6 1989.9 2028 C +1989.8 2027 1990.8 2028.3 1990.1 2027.3 C +1988.9 2026.7 1986.7 2026.9 1986.8 2024.7 C +1987.4 2023 1985.9 2024.6 1985.1 2023.7 C +1984.1 2021.4 1982.5 2020.5 1980.3 2020.6 C +1979.9 2020.8 1979.5 2021.1 1979.3 2021.6 C +1979.7 2025.8 1978.4 2033 1979.6 2038.1 C +1983.7 2042.9 1968.8 2044.6 1978.8 2042.7 C +1979.3 2042.3 1979.6 2041.9 1980 2041.5 C +1980 2034.8 1980 2027 1980 2021.6 C +1981.3 2020.5 1981.7 2021.5 1982.9 2021.8 C +1983.6 2024.7 1986.1 2023.8 1986.8 2026.4 C +1987.1 2027.7 1988.6 2027.1 1989.2 2028.3 C +1989.1 2036.7 1989.3 2044.8 1988.9 2053.7 C +1987.2 2054.9 1986.2 2056.8 1983.9 2057.1 C +1986.3 2055.9 1990.9 2055 1989.6 2051.3 C +f +S +n +1971.6 2078.9 m +1971.4 2070.5 1972.1 2062.2 1971.6 2055.9 C +1969.9 2053.7 1967.6 2051.7 1965.6 2049.6 C +1961.4 2050.4 1957.6 2053.6 1953.4 2055.2 C +1949.8 2055.6 1948.2 2051.2 1945.5 2049.6 C +1945.1 2048.8 1944.5 2047.9 1943.6 2047.5 C +1940.1 2047.8 1937.3 2051 1934 2052.3 C +1933.7 2052.6 1933.7 2053 1933.2 2053.2 C +1933.7 2060.8 1933.4 2067.2 1933.5 2074.6 C +1933.8 2068.1 1934 2060.9 1933.2 2054 C +1935.3 2050.9 1939.3 2049.6 1942.4 2047.5 C +1942.8 2047.5 1943.4 2047.4 1943.8 2047.7 C +1947.1 2050.2 1950.3 2057.9 1955.3 2054.4 C +1955.4 2054.4 1955.5 2054.3 1955.6 2054.2 C +1955.9 2057.6 1956.1 2061.8 1955.3 2064.8 C +1955.4 2064.3 1955.1 2063.8 1955.6 2063.6 C +1956 2066.6 1955.3 2068.7 1958.7 2069.8 C +1959.2 2071.7 1961.4 2071.7 1962 2074.1 C +1964.4 2074.2 1964 2077.7 1967.3 2078.4 C +1967 2079.7 1968.1 2079.9 1969 2080.1 C +1971.1 2079.9 1970 2079.2 1970.4 2078 C +1969.5 2077.2 1970.3 2075.9 1969.7 2075.1 C +1970.1 2069.8 1970.1 2063.6 1969.7 2058.8 C +1969.2 2058.5 1970 2058.1 1970.2 2057.8 C +1970.4 2058.3 1971.2 2057.7 1971.4 2058.3 C +1971.5 2065.3 1971.2 2073.6 1971.6 2081.1 C +1974.1 2081.4 1969.8 2084.3 1972.4 2082.5 C +1971.9 2081.4 1971.6 2080.2 1971.6 2078.9 C +[0 0.4 1 0] vc +f +S +n +1952.4 2052 m +1954.1 2051.3 1955.6 2050.4 1957.2 2049.6 C +1955.6 2050.4 1954.1 2051.3 1952.4 2052 C +[0 0.87 0.91 0.83] vc +f +S +n +1975.5 2039.8 m +1975.5 2039.4 1974.5 2038.7 1974.8 2038.8 C +1974.6 2039.1 1974.5 2039.1 1974.3 2038.8 C +1974.6 2037.6 1972.5 2039.3 1972.1 2038.4 C +1970.4 2038.4 1970.5 2035.5 1968 2035.7 C +1968.6 2032.5 1964 2033.3 1964 2030.4 C +1961.9 2029.8 1961.4 2026.8 1959.2 2026.8 C +1957.7 2028.5 1952.6 2025.3 1953.9 2029.5 C +1952.9 2031.5 1947 2028.2 1947.9 2032.4 C +1946 2034.2 1941.5 2031.5 1941.9 2035.7 C +1940.6 2036.1 1938.9 2036.1 1937.6 2035.7 C +1937.3 2037.5 1940.7 2038.2 1941.2 2040.3 C +1943.7 2041.1 1944.4 2043.9 1946.4 2045.6 C +1949.1 2046.7 1949.9 2048.8 1952.2 2051.1 C +1959.9 2047.1 1967.7 2043.6 1975.5 2039.8 C +[0 0.01 1 0] vc +f +S +n +vmrs +1962 2047.2 m +1960.2 2047.5 1959.5 2048.9 1957.7 2049.4 C +1959.5 2049.5 1960.3 2047.2 1962 2047.2 C +[0 0.87 0.91 0.83] vc +f +0.4 w +2 J +2 M +S +n +2012.4 2046.3 m +2010.3 2051.3 2008.3 2041.5 2009.8 2047.7 C +2010.5 2047.4 2012.2 2046.5 2012.4 2046.3 C +f +S +n +1944.8 2044.6 m +1945.5 2045.6 1948.6 2048.1 1947.4 2046.8 C +1946.3 2046.5 1945.5 2045.2 1944.8 2044.6 C +f +S +n +1987.2 2054.9 m +1983.7 2057.3 1979.6 2058 1976 2060.2 C +1974.7 2058.2 1977.2 2055.8 1974.3 2054.9 C +1973.1 2052 1970.4 2050.2 1968 2048 C +1968 2047.7 1968 2047.4 1968.3 2047.2 C +1969.5 2046.1 1983 2040.8 1972.4 2044.8 C +1971.2 2046.6 1967.9 2046 1968 2048.2 C +1970.5 2050.7 1973.8 2052.6 1974.3 2055.6 C +1975.1 2055 1975.7 2056.7 1975.7 2057.1 C +1975.7 2058.2 1974.8 2059.3 1975.5 2060.4 C +1979.3 2058.2 1983.9 2057.7 1987.2 2054.9 C +[0.18 0.18 0 0.78] vc +f +S +n +1967.8 2047.5 m +1968.5 2047 1969.1 2046.5 1969.7 2046 C +1969.1 2046.5 1968.5 2047 1967.8 2047.5 C +[0 0.87 0.91 0.83] vc +f +S +n +1976.7 2040.3 m +1976.9 2032.8 1976.3 2024.4 1977.2 2017 C +1976.4 2016.5 1976.6 2014.8 1976.9 2013.9 C +1978.7 2012.7 1977.2 2013 1976.2 2012.9 C +1971.5 2008.1 1965.9 2003.1 1961.8 1998 C +1960.9 1998 1960.1 1998 1959.2 1998 C +1951.5 2001.1 1944.3 2005.5 1937.1 2009.6 C +1935 2012.9 1937 2013.6 1936.1 2017.2 C +1937.1 2017.2 1936 2018 1936.4 2017.7 C +1937 2020.1 1935.5 2022.1 1936.4 2024.9 C +1936.8 2027.2 1936.5 2029.7 1936.6 2032.1 C +1937.4 2028.2 1936 2023.8 1936.8 2020.1 C +1938.3 2015.7 1933.6 2011 1939 2008.6 C +1945.9 2004.5 1953.1 2000.3 1960.6 1998.3 C +1960.9 1998.3 1961.3 1998.3 1961.6 1998.3 C +1966.2 2003.5 1971.8 2008.4 1976.4 2013.6 C +1976.6 2015.5 1976 2018.1 1976.9 2019.2 C +1976.1 2025.8 1976.4 2033.2 1976.7 2040 C +1971.9 2042.4 1967.4 2045 1962.5 2047 C +1967.3 2044.9 1972 2042.6 1976.7 2040.3 C +f +S +n +1939 2038.6 m +1940.6 2040.9 1942.5 2042.1 1944.3 2044.1 C +1942.7 2041.9 1940.6 2040.9 1939 2038.6 C +f +S +n +2006.2 2065.7 m +2006 2057.3 2006.7 2049 2006.2 2042.7 C +2002.1 2038.4 1997.7 2033.4 1993 2030 C +1992.9 2029.3 1992.5 2028.6 1992 2028.3 C +1992.1 2036.6 1991.9 2046.2 1992.3 2054.9 C +1990.8 2056.2 1989 2056.7 1987.5 2058 C +1988.7 2057.7 1990.7 2054.4 1993 2056.4 C +1993.4 2058.8 1996 2058.2 1996.6 2060.9 C +1999 2061 1998.5 2064.5 2001.9 2065.2 C +2001.5 2066.5 2002.7 2066.7 2003.6 2066.9 C +2005.7 2066.7 2004.6 2066 2005 2064.8 C +2004 2064 2004.8 2062.7 2004.3 2061.9 C +2004.6 2056.6 2004.6 2050.4 2004.3 2045.6 C +2003.7 2045.3 2004.6 2044.9 2004.8 2044.6 C +2005 2045.1 2005.7 2044.5 2006 2045.1 C +2006 2052.1 2005.8 2060.4 2006.2 2067.9 C +2008.7 2068.2 2004.4 2071.1 2006.9 2069.3 C +2006.4 2068.2 2006.2 2067 2006.2 2065.7 C +[0 0.4 1 0] vc +f +S +n +2021.8 2041.7 m +2018.3 2044.1 2014.1 2044.8 2010.5 2047 C +2009.3 2045 2011.7 2042.6 2008.8 2041.7 C +2004.3 2035.1 1997.6 2030.9 1993 2024.4 C +1992.1 2024 1991.5 2024.3 1990.8 2024 C +1993.2 2023.9 1995.3 2027.1 1996.8 2029 C +2000.4 2032.6 2004.9 2036.9 2008.4 2040.8 C +2008.2 2043.1 2011.4 2042.8 2009.8 2045.8 C +2009.8 2046.3 2009.7 2046.9 2010 2047.2 C +2013.8 2045 2018.5 2044.5 2021.8 2041.7 C +[0.18 0.18 0 0.78] vc +f +S +n +2001.6 2034 m +2000.7 2033.1 1999.9 2032.3 1999 2031.4 C +1999.9 2032.3 2000.7 2033.1 2001.6 2034 C +[0 0.87 0.91 0.83] vc +f +S +n +vmrs +1989.4 2024.4 m +1989.5 2025.4 1988.6 2024.3 1988.9 2024.7 C +1990.5 2025.8 1990.7 2024.2 1992.8 2024.9 C +1993.8 2025.9 1995 2027.1 1995.9 2028 C +1994.3 2026 1991.9 2023.4 1989.4 2024.4 C +[0 0.87 0.91 0.83] vc +f +0.4 w +2 J +2 M +S +n +1984.8 2019.9 m +1984.6 2018.6 1986.3 2017.2 1987.7 2016.8 C +1987.2 2017.5 1982.9 2017.9 1984.4 2020.6 C +1984.1 2019.9 1984.9 2020 1984.8 2019.9 C +f +S +n +1981.7 2017 m +1979.6 2022 1977.6 2012.3 1979.1 2018.4 C +1979.8 2018.1 1981.5 2017.2 1981.7 2017 C +f +S +n +1884.3 2019.2 m +1884.7 2010.5 1884.5 2000.6 1884.5 1991.8 C +1886.6 1989.3 1889.9 1988.9 1892.4 1987 C +1890.8 1988.7 1886 1989.1 1884.3 1992.3 C +1884.7 2001 1884.5 2011.3 1884.5 2019.9 C +1891 2025.1 1895.7 2031.5 1902 2036.9 C +1896.1 2031 1890 2024.9 1884.3 2019.2 C +[0.07 0.06 0 0.58] vc +f +S +n +1884 2019.4 m +1884.5 2010.6 1884.2 2000.4 1884.3 1991.8 C +1884.8 1990.4 1887.8 1989 1884.8 1990.8 C +1884.3 1991.3 1884.3 1992 1884 1992.5 C +1884.5 2001.2 1884.2 2011.1 1884.3 2019.9 C +1887.9 2023.1 1891.1 2026.4 1894.4 2030 C +1891.7 2026.1 1887.1 2022.9 1884 2019.4 C +[0.4 0.4 0 0] vc +f +S +n +1885 2011.7 m +1885 2006.9 1885 2001.9 1885 1997.1 C +1885 2001.9 1885 2006.9 1885 2011.7 C +[0 0.87 0.91 0.83] vc +f +S +n +1975.5 2036.4 m +1975.2 2028 1976 2019.7 1975.5 2013.4 C +1971.1 2008.5 1965.6 2003.6 1961.6 1999 C +1958.8 1998 1956 2000 1953.6 2001.2 C +1948.2 2004.7 1941.9 2006.5 1937.1 2010.8 C +1937.5 2018.3 1937.3 2024.7 1937.3 2032.1 C +1937.6 2025.6 1937.9 2018.4 1937.1 2011.5 C +1937.3 2011 1937.6 2010.5 1937.8 2010 C +1944.6 2005.7 1951.9 2002.3 1959.2 1999 C +1960.1 1998.5 1960.1 1999.8 1960.4 2000.4 C +1959.7 2006.9 1959.7 2014.2 1959.4 2021.1 C +1959 2021.1 1959.2 2021.9 1959.2 2022.3 C +1959.2 2021.9 1959 2021.3 1959.4 2021.1 C +1959.8 2024.1 1959.2 2026.2 1962.5 2027.3 C +1963 2029.2 1965.3 2029.2 1965.9 2031.6 C +1968.3 2031.8 1967.8 2035.2 1971.2 2036 C +1970.8 2037.2 1971.9 2037.5 1972.8 2037.6 C +1974.9 2037.4 1973.9 2036.7 1974.3 2035.5 C +1973.3 2034.7 1974.1 2033.4 1973.6 2032.6 C +1973.9 2027.3 1973.9 2021.1 1973.6 2016.3 C +1973 2016 1973.9 2015.6 1974 2015.3 C +1974.3 2015.9 1975 2015.3 1975.2 2015.8 C +1975.3 2022.8 1975.1 2031.2 1975.5 2038.6 C +1977.9 2039 1973.7 2041.8 1976.2 2040 C +1975.7 2039 1975.5 2037.8 1975.5 2036.4 C +[0 0.4 1 0] vc +f +S +n +1991.1 2012.4 m +1987.5 2014.8 1983.4 2015.6 1979.8 2017.7 C +1978.5 2015.7 1981 2013.3 1978.1 2012.4 C +1973.6 2005.8 1966.8 2001.6 1962.3 1995.2 C +1961.4 1994.7 1960.8 1995 1960.1 1994.7 C +1962.5 1994.6 1964.6 1997.8 1966.1 1999.7 C +1969.7 2003.3 1974.2 2007.6 1977.6 2011.5 C +1977.5 2013.8 1980.6 2013.5 1979.1 2016.5 C +1979.1 2017 1979 2017.6 1979.3 2018 C +1983.1 2015.7 1987.8 2015.2 1991.1 2012.4 C +[0.18 0.18 0 0.78] vc +f +S +n +1970.9 2004.8 m +1970 2003.9 1969.2 2003 1968.3 2002.1 C +1969.2 2003 1970 2003.9 1970.9 2004.8 C +[0 0.87 0.91 0.83] vc +f +S +n +1887.9 1994.9 m +1888.5 1992.3 1891.4 1992.2 1893.2 1990.8 C +1898.4 1987.5 1904 1984.8 1909.5 1982.2 C +1909.7 1982.7 1910.3 1982.1 1910.4 1982.7 C +1909.5 1990.5 1910.1 1996.4 1910 2004.5 C +1909.1 2003.4 1909.7 2005.8 1909.5 2006.4 C +1910.4 2006 1909.7 2008 1910.2 2007.9 C +1911.3 2010.6 1912.5 2012.6 1915.7 2013.4 C +1915.8 2013.7 1915.5 2014.4 1916 2014.4 C +1916.3 2015 1915.4 2016 1915.2 2016 C +1916.1 2015.5 1916.5 2014.5 1916 2013.6 C +1913.4 2013.3 1913.1 2010.5 1910.9 2009.8 C +1910.7 2008.8 1910.4 2007.9 1910.2 2006.9 C +1911.1 1998.8 1909.4 1990.7 1910.7 1982.4 C +1910 1982.1 1908.9 1982.1 1908.3 1982.4 C +1901.9 1986.1 1895 1988.7 1888.8 1993 C +1888 1993.4 1888.4 1994.3 1887.6 1994.7 C +1888.1 2001.3 1887.8 2008.6 1887.9 2015.1 C +1887.3 2017.5 1887.9 2015.4 1888.4 2014.4 C +1887.8 2008 1888.4 2001.3 1887.9 1994.9 C +[0.07 0.06 0 0.58] vc +f +S +n +vmrs +1887.9 2018.4 m +1887.5 2016.9 1888.5 2016 1888.8 2014.8 C +1890.1 2014.8 1891.1 2016.6 1892.4 2015.3 C +1892.4 2014.4 1893.8 2012.9 1894.4 2012.4 C +1895.9 2012.4 1896.6 2013.9 1897.7 2012.7 C +1898.4 2011.7 1898.6 2010.4 1899.6 2009.8 C +1901.7 2009.9 1902.9 2010.4 1904 2009.1 C +1904.3 2007.4 1904 2007.6 1904.9 2007.2 C +1906.2 2007 1907.6 2006.5 1908.8 2006.7 C +1910.6 2008.2 1909.8 2011.5 1912.6 2012 C +1912.4 2013 1913.8 2012.7 1914 2013.2 C +1911.5 2011.1 1909.1 2007.9 1909.2 2004.3 C +1909.5 2003.5 1909.9 2004.9 1909.7 2004.3 C +1909.9 1996.2 1909.3 1990.5 1910.2 1982.7 C +1909.5 1982.6 1909.5 1982.6 1908.8 1982.7 C +1903.1 1985.7 1897 1987.9 1891.7 1992 C +1890.5 1993 1888.2 1992.9 1888.1 1994.9 C +1888.7 2001.4 1888.1 2008.4 1888.6 2014.8 C +1888.3 2016 1887.2 2016.9 1887.6 2018.4 C +1892.3 2023.9 1897.6 2027.9 1902.3 2033.3 C +1898 2028.2 1892.1 2023.8 1887.9 2018.4 C +[0.4 0.4 0 0] vc +f +0.4 w +2 J +2 M +S +n +1910.9 1995.2 m +1910.4 1999.8 1911 2003.3 1910.9 2008.1 C +1910.9 2003.8 1910.9 1999.2 1910.9 1995.2 C +[0.18 0.18 0 0.78] vc +f +S +n +1911.2 2004.3 m +1911.2 2001.9 1911.2 1999.7 1911.2 1997.3 C +1911.2 1999.7 1911.2 2001.9 1911.2 2004.3 C +[0 0.87 0.91 0.83] vc +f +S +n +1958.7 1995.2 m +1959 1995.6 1956.2 1995 1956.5 1996.8 C +1955.8 1997.6 1954.2 1998.5 1953.6 1997.3 C +1953.6 1990.8 1954.9 1989.6 1953.4 1983.9 C +1953.4 1983.3 1953.3 1982.1 1954.4 1982 C +1955.5 1982.6 1956.5 1981.3 1957.5 1981 C +1956.3 1981.8 1954.7 1982.6 1953.9 1981.5 C +1951.4 1983 1954.7 1988.8 1952.9 1990.6 C +1953.8 1990.6 1953.2 1992.7 1953.4 1993.7 C +1953.8 1994.5 1952.3 1996.1 1953.2 1997.8 C +1956.3 1999.4 1957.5 1994 1959.9 1995.6 C +1962 1994.4 1963.7 1997.7 1965.2 1998.8 C +1963.5 1996.7 1961.2 1994.1 1958.7 1995.2 C +f +S +n +1945 2000.7 m +1945.4 1998.7 1945.4 1997.9 1945 1995.9 C +1944.5 1995.3 1944.2 1992.6 1945.7 1993.2 C +1946 1992.2 1948.7 1992.5 1948.4 1990.6 C +1947.5 1990.3 1948.1 1988.7 1947.9 1988.2 C +1948.9 1987.8 1950.5 1986.8 1950.5 1984.6 C +1951.5 1980.9 1946.7 1983 1947.2 1979.8 C +1944.5 1979.9 1945.2 1976.6 1943.1 1976.7 C +1941.8 1975.7 1942.1 1972.7 1939.2 1973.8 C +1938.2 1974.6 1939.3 1971.6 1938.3 1970.9 C +1938.8 1969.2 1933.4 1970.3 1937.3 1970 C +1939.4 1971.2 1937.2 1973 1937.6 1974.3 C +1937.2 1976.3 1937.1 1981.2 1937.8 1984.1 C +1938.8 1982.3 1937.9 1976.6 1938.5 1973.1 C +1938.9 1975 1938.5 1976.4 1939.7 1977.2 C +1939.5 1983.5 1938.9 1991.3 1940.2 1997.3 C +1939.4 1999.1 1938.6 1997.1 1937.8 1997.1 C +1937.4 1996.7 1937.6 1996.1 1937.6 1995.6 C +1936.5 1998.5 1940.1 1998.4 1940.9 2000.7 C +1942.1 2000.4 1943.2 2001.3 1943.1 2002.4 C +1943.6 2003.1 1941.1 2004.6 1942.8 2003.8 C +1943.9 2002.5 1942.6 2000.6 1945 2000.7 C +[0.65 0.65 0 0.42] vc +f +S +n +1914.5 2006.4 m +1914.1 2004.9 1915.2 2004 1915.5 2002.8 C +1916.7 2002.8 1917.8 2004.6 1919.1 2003.3 C +1919 2002.4 1920.4 2000.9 1921 2000.4 C +1922.5 2000.4 1923.2 2001.9 1924.4 2000.7 C +1925 1999.7 1925.3 1998.4 1926.3 1997.8 C +1928.4 1997.9 1929.5 1998.4 1930.6 1997.1 C +1930.9 1995.4 1930.7 1995.6 1931.6 1995.2 C +1932.8 1995 1934.3 1994.5 1935.4 1994.7 C +1936.1 1995.8 1936.9 1996.2 1936.6 1997.8 C +1938.9 1999.4 1939.7 2001.3 1942.4 2002.4 C +1942.4 2002.5 1942.2 2003 1942.6 2002.8 C +1942.9 2000.4 1939.2 2001.8 1939.2 1999.7 C +1936.2 1998.6 1937 1995.3 1935.9 1993.5 C +1937.1 1986.5 1935.2 1977.9 1937.6 1971.2 C +1937.6 1970.3 1936.6 1971 1936.4 1970.4 C +1930.2 1973.4 1924 1976 1918.4 1980 C +1917.2 1981 1914.9 1980.9 1914.8 1982.9 C +1915.3 1989.4 1914.7 1996.4 1915.2 2002.8 C +1914.9 2004 1913.9 2004.9 1914.3 2006.4 C +1919 2011.9 1924.2 2015.9 1928.9 2021.3 C +1924.6 2016.2 1918.7 2011.8 1914.5 2006.4 C +[0.4 0.4 0 0] vc +f +S +n +1914.5 1982.9 m +1915.1 1980.3 1918 1980.2 1919.8 1978.8 C +1925 1975.5 1930.6 1972.8 1936.1 1970.2 C +1939.4 1970.6 1936.1 1974.2 1936.6 1976.4 C +1936.5 1981.9 1936.8 1987.5 1936.4 1992.8 C +1935.9 1992.8 1936.2 1993.5 1936.1 1994 C +1937.1 1993.6 1936.2 1995.9 1936.8 1995.9 C +1937 1998 1939.5 1999.7 1940.4 2000.7 C +1940.1 1998.6 1935 1997.2 1937.6 1993.7 C +1938.3 1985.7 1935.9 1976.8 1937.8 1970.7 C +1936.9 1969.8 1935.4 1970.3 1934.4 1970.7 C +1928.3 1974.4 1921.4 1976.7 1915.5 1981 C +1914.6 1981.4 1915.1 1982.3 1914.3 1982.7 C +1914.7 1989.3 1914.5 1996.6 1914.5 2003.1 C +1913.9 2005.5 1914.5 2003.4 1915 2002.4 C +1914.5 1996 1915.1 1989.3 1914.5 1982.9 C +[0.07 0.06 0 0.58] vc +f +S +n +1939.2 1994.9 m +1939.3 1995 1939.4 1995.1 1939.5 1995.2 C +1939.1 1989 1939.3 1981.6 1939 1976.7 C +1938.6 1976.3 1938.6 1974.6 1938.5 1973.3 C +1938.7 1976.1 1938.1 1979.4 1939 1981.7 C +1937.3 1986 1937.7 1991.6 1938 1996.4 C +1937.3 1994.3 1939.6 1996.2 1939.2 1994.9 C +[0.18 0.18 0 0.78] vc +f +S +n +1938.3 1988.4 m +1938.5 1990.5 1937.9 1994.1 1938.8 1994.7 C +1937.9 1992.6 1939 1990.6 1938.3 1988.4 C +[0 0.87 0.91 0.83] vc +f +S +n +1938.8 1985.8 m +1938.5 1985.9 1938.4 1985.7 1938.3 1985.6 C +1938.4 1986.2 1938 1989.5 1938.8 1987.2 C +1938.8 1986.8 1938.8 1986.3 1938.8 1985.8 C +f +S +n +vmrs +1972.8 2062.1 m +1971.9 2061 1972.5 2059.4 1972.4 2058 C +1972.2 2063.8 1971.9 2073.7 1972.4 2081.3 C +1972.5 2074.9 1971.9 2067.9 1972.8 2062.1 C +[0 1 1 0.36] vc +f +0.4 w +2 J +2 M +S +n +1940.2 2071.7 m +1941.3 2072 1943.1 2072.3 1944 2071.5 C +1943.6 2069.9 1945.2 2069.1 1946 2068.8 C +1950 2071.1 1948.7 2065.9 1951.7 2066.2 C +1953.5 2063.9 1956.9 2069.4 1955.6 2063.8 C +1955.5 2064.2 1955.7 2064.8 1955.3 2065 C +1954.3 2063.7 1956.2 2063.6 1955.6 2062.1 C +1954.5 2060 1958.3 2050.3 1952.2 2055.6 C +1949.1 2053.8 1946 2051 1943.8 2048 C +1940.3 2048 1937.5 2051.3 1934.2 2052.5 C +1933.1 2054.6 1934.4 2057.3 1934 2060 C +1934 2065.1 1934 2069.7 1934 2074.6 C +1934.4 2069 1934.1 2061.5 1934.2 2054.9 C +1934.6 2054.5 1935.3 2054.7 1935.9 2054.7 C +1937 2055.3 1935.9 2056.1 1935.9 2056.8 C +1936.5 2063 1935.6 2070.5 1935.9 2074.6 C +1936.7 2074.4 1937.3 2075.2 1938 2074.6 C +1937.9 2073.6 1939.1 2072.1 1940.2 2071.7 C +[0 0.2 1 0] vc +f +S +n +1933.2 2074.1 m +1933.2 2071.5 1933.2 2069 1933.2 2066.4 C +1933.2 2069 1933.2 2071.5 1933.2 2074.1 C +[0 1 1 0.36] vc +f +S +n +2007.4 2048.9 m +2006.5 2047.8 2007.1 2046.2 2006.9 2044.8 C +2006.7 2050.6 2006.5 2060.5 2006.9 2068.1 C +2007.1 2061.7 2006.5 2054.7 2007.4 2048.9 C +f +S +n +1927.2 2062.4 m +1925.8 2060.1 1928.1 2058.2 1927 2056.4 C +1927.3 2055.5 1926.5 2053.5 1926.8 2051.8 C +1926.8 2052.8 1926 2052.5 1925.3 2052.5 C +1924.1 2052.8 1925 2050.5 1924.4 2050.1 C +1925.3 2050.2 1925.4 2048.8 1926.3 2049.4 C +1926.5 2052.3 1928.4 2047.2 1928.4 2051.1 C +1928.9 2050.5 1929 2051.4 1928.9 2051.8 C +1928.9 2052 1928.9 2052.3 1928.9 2052.5 C +1929.4 2051.4 1928.9 2049 1930.1 2048.2 C +1928.9 2047.1 1930.5 2047.1 1930.4 2046.5 C +1931.9 2046.2 1933.1 2046.1 1934.7 2046.5 C +1934.6 2046.9 1935.2 2047.9 1934.4 2048.4 C +1936.9 2048.1 1933.6 2043.8 1935.9 2043.9 C +1935.7 2043.9 1934.8 2041.3 1933.2 2041.7 C +1932.5 2041.6 1932.4 2039.6 1932.3 2041 C +1930.8 2042.6 1929 2040.6 1927.7 2042 C +1927.5 2041.4 1927.1 2040.9 1927.2 2040.3 C +1927.8 2040.6 1927.4 2039.1 1928.2 2038.6 C +1929.4 2038 1930.5 2038.8 1931.3 2037.9 C +1931.7 2039 1932.5 2038.6 1931.8 2037.6 C +1930.9 2037 1928.7 2037.8 1928.2 2037.9 C +1926.7 2037.8 1928 2039 1927 2038.8 C +1927.4 2040.4 1925.6 2040.8 1925.1 2041 C +1924.3 2040.4 1923.2 2040.5 1922.2 2040.5 C +1921.4 2041.7 1921 2043.9 1919.3 2043.9 C +1918.8 2043.4 1917.2 2043.3 1916.4 2043.4 C +1915.9 2044.4 1915.7 2046 1914.3 2046.5 C +1913.1 2046.6 1912 2044.5 1911.4 2046.3 C +1912.8 2046.5 1913.8 2047.4 1915.7 2047 C +1916.9 2047.7 1915.6 2048.8 1916 2049.4 C +1915.4 2049.3 1913.9 2050.3 1913.3 2051.1 C +1913.9 2054.1 1916 2050.2 1916.7 2053 C +1916.9 2053.8 1915.5 2054.1 1916.7 2054.4 C +1917 2054.7 1920.2 2054.3 1919.3 2056.6 C +1918.8 2056.1 1920.2 2058.6 1920.3 2057.6 C +1921.2 2057.9 1922.1 2057.5 1922.4 2059 C +1922.3 2059.1 1922.2 2059.3 1922 2059.2 C +1922.1 2059.7 1922.4 2060.3 1922.9 2060.7 C +1923.2 2060.1 1923.8 2060.4 1924.6 2060.7 C +1925.9 2062.6 1923.2 2062 1925.6 2063.6 C +1926.1 2063.1 1927.3 2062.5 1927.2 2062.4 C +[0.21 0.21 0 0] vc +f +S +n +1933.2 2063.3 m +1933.2 2060.7 1933.2 2058.2 1933.2 2055.6 C +1933.2 2058.2 1933.2 2060.7 1933.2 2063.3 C +[0 1 1 0.36] vc +f +S +n +1965.2 2049.2 m +1967.1 2050.1 1969.9 2053.7 1972.1 2056.4 C +1970.5 2054 1967.6 2051.3 1965.2 2049.2 C +f +S +n +1991.8 2034.8 m +1991.7 2041.5 1992 2048.5 1991.6 2055.2 C +1990.5 2056.4 1991.9 2054.9 1991.8 2054.4 C +1991.8 2047.9 1991.8 2041.3 1991.8 2034.8 C +f +S +n +1988.9 2053.2 m +1988.9 2044.3 1988.9 2036.6 1988.9 2028.3 C +1985.7 2028.2 1987.2 2023.5 1983.9 2024.2 C +1983.9 2022.4 1982 2021.6 1981 2021.3 C +1980.6 2021.1 1980.6 2021.7 1980.3 2021.6 C +1980.3 2027 1980.3 2034.8 1980.3 2041.5 C +1979.3 2043.2 1977.6 2043 1976.2 2043.6 C +1977.1 2043.8 1978.5 2043.2 1978.8 2044.1 C +1978.5 2045.3 1979.9 2045.3 1980.3 2045.8 C +1980.5 2046.8 1980.7 2046.2 1981.5 2046.5 C +1982.4 2047.1 1982 2048.6 1982.7 2049.4 C +1984.2 2049.6 1984.6 2052.2 1986.8 2051.6 C +1987.1 2048.6 1985.1 2042.7 1986.5 2040.5 C +1986.3 2036.7 1986.9 2031.7 1986 2029.2 C +1986.3 2027.1 1986.9 2028.6 1987.7 2027.6 C +1987.7 2028.3 1988.7 2028 1988.7 2028.8 C +1988.1 2033 1988.7 2037.5 1988.2 2041.7 C +1987.8 2041.4 1988 2040.8 1988 2040.3 C +1988 2041 1988 2042.4 1988 2042.4 C +1988 2042.4 1988.1 2042.3 1988.2 2042.2 C +1989.3 2046 1988 2050.2 1988.4 2054 C +1987.8 2054.4 1987.1 2054.7 1986.5 2055.4 C +1987.4 2054.4 1988.4 2054.6 1988.9 2053.2 C +[0 1 1 0.23] vc +f +S +n +1950.8 2054.4 m +1949.7 2053.4 1948.7 2052.3 1947.6 2051.3 C +1948.7 2052.3 1949.7 2053.4 1950.8 2054.4 C +[0 1 1 0.36] vc +f +S +n +vmrs +2006.7 2043.2 m +2004.5 2040.8 2002.4 2038.4 2000.2 2036 C +2002.4 2038.4 2004.5 2040.8 2006.7 2043.2 C +[0 1 1 0.36] vc +f +0.4 w +2 J +2 M +S +n +1976.7 2019.6 m +1975.8 2018.6 1976.4 2016.9 1976.2 2015.6 C +1976 2021.3 1975.8 2031.2 1976.2 2038.8 C +1976.4 2032.4 1975.8 2025.5 1976.7 2019.6 C +f +S +n +1988.4 2053.5 m +1988.6 2049.2 1988.1 2042.8 1988 2040 C +1988.4 2040.4 1988.1 2041 1988.2 2041.5 C +1988.3 2037.2 1988 2032.7 1988.4 2028.5 C +1987.6 2027.1 1987.2 2028.6 1986.8 2028 C +1985.9 2028.5 1986.5 2029.7 1986.3 2030.4 C +1986.9 2029.8 1986.6 2031 1987 2031.2 C +1987.4 2039.6 1985 2043 1987.2 2050.4 C +1987.2 2051.6 1985.9 2052.3 1984.6 2051.3 C +1981.9 2049.7 1982.9 2047 1980.3 2046.5 C +1980.3 2045.2 1978.1 2046.2 1978.6 2043.9 C +1975.6 2043.3 1979.3 2045.6 1979.6 2046.5 C +1980.8 2046.6 1981.5 2048.5 1982.2 2049.9 C +1983.7 2050.8 1984.8 2052.8 1986.5 2053 C +1986.7 2053.5 1987.5 2054.1 1987 2054.7 C +1987.4 2053.9 1988.3 2054.3 1988.4 2053.5 C +[0 1 1 0.23] vc +f +S +n +1988 2038.1 m +1988 2036.7 1988 2035.4 1988 2034 C +1988 2035.4 1988 2036.7 1988 2038.1 C +[0 1 1 0.36] vc +f +S +n +1999.7 2035.7 m +1997.6 2033.5 1995.4 2031.2 1993.2 2029 C +1995.4 2031.2 1997.6 2033.5 1999.7 2035.7 C +f +S +n +1944 2029.2 m +1945.2 2029.5 1946.9 2029.8 1947.9 2029 C +1947.4 2027.4 1949 2026.7 1949.8 2026.4 C +1953.9 2028.6 1952.6 2023.4 1955.6 2023.7 C +1957.4 2021.4 1960.7 2027 1959.4 2021.3 C +1959.3 2021.7 1959.6 2022.3 1959.2 2022.5 C +1958.1 2021.2 1960.1 2021.1 1959.4 2019.6 C +1959.1 2012.7 1959.9 2005.1 1959.6 1999.2 C +1955.3 2000.1 1951.3 2003.1 1947.2 2005 C +1943.9 2006 1941.2 2008.7 1938 2010 C +1936.9 2012.1 1938.2 2014.8 1937.8 2017.5 C +1937.8 2022.6 1937.8 2027.3 1937.8 2032.1 C +1938.2 2026.5 1938 2019 1938 2012.4 C +1938.5 2012 1939.2 2012.3 1939.7 2012.2 C +1940.8 2012.8 1939.7 2013.6 1939.7 2014.4 C +1940.4 2020.5 1939.4 2028 1939.7 2032.1 C +1940.6 2031.9 1941.2 2032.7 1941.9 2032.1 C +1941.7 2031.2 1943 2029.7 1944 2029.2 C +[0 0.2 1 0] vc +f +S +n +1937.1 2031.6 m +1937.1 2029.1 1937.1 2026.5 1937.1 2024 C +1937.1 2026.5 1937.1 2029.1 1937.1 2031.6 C +[0 1 1 0.36] vc +f +S +n +1991.8 2028 m +1992.5 2027.8 1993.2 2029.9 1994 2030.2 C +1992.9 2029.6 1993.1 2028.1 1991.8 2028 C +[0 1 1 0.23] vc +f +S +n +1991.8 2027.8 m +1992.4 2027.6 1992.6 2028.3 1993 2028.5 C +1992.6 2028.2 1992.2 2027.6 1991.6 2027.8 C +1991.6 2028.5 1991.6 2029.1 1991.6 2029.7 C +1991.6 2029.1 1991.4 2028.3 1991.8 2027.8 C +[0 1 1 0.36] vc +f +S +n +1985.8 2025.4 m +1985.3 2025.2 1984.8 2024.7 1984.1 2024.9 C +1983.3 2025.3 1983.6 2027.3 1983.9 2027.6 C +1985 2028 1986.9 2026.9 1985.8 2025.4 C +[0 1 1 0.23] vc +f +S +n +vmrs +1993.5 2024.4 m +1992.4 2023.7 1991.3 2022.9 1990.1 2023.2 C +1990.7 2023.7 1989.8 2023.8 1989.4 2023.7 C +1989.1 2023.7 1988.6 2023.9 1988.4 2023.5 C +1988.5 2023.2 1988.3 2022.7 1988.7 2022.5 C +1989 2022.6 1988.9 2023 1988.9 2023.2 C +1989.1 2022.8 1990.4 2022.3 1990.6 2021.3 C +1990.4 2021.8 1990 2021.3 1990.1 2021.1 C +1990.1 2020.9 1990.1 2020.1 1990.1 2020.6 C +1989.9 2021.1 1989.5 2020.6 1989.6 2020.4 C +1989.6 2019.8 1988.7 2019.6 1988.2 2019.2 C +1987.5 2018.7 1987.7 2020.2 1987 2019.4 C +1987.5 2020.4 1986 2021.1 1987.5 2021.8 C +1986.8 2023.1 1986.6 2021.1 1986 2021.1 C +1986.1 2020.1 1985.9 2019 1986.3 2018.2 C +1986.7 2018.4 1986.5 2019 1986.5 2019.4 C +1986.5 2018.7 1986.4 2017.8 1987.2 2017.7 C +1986.5 2017.2 1985.5 2019.3 1985.3 2020.4 C +1986.2 2022 1987.3 2023.5 1989.2 2024.2 C +1990.8 2024.3 1991.6 2022.9 1993.2 2024.4 C +1993.8 2025.4 1995 2026.6 1995.9 2027.1 C +1995 2026.5 1994.1 2025.5 1993.5 2024.4 C +[0 1 1 0.36] vc +f +0.4 w +2 J +2 M +[0 0.5 0.5 0.2] vc +S +n +2023 2040.3 m +2023.2 2036 2022.7 2029.6 2022.5 2026.8 C +2022.9 2027.2 2022.7 2027.8 2022.8 2028.3 C +2022.8 2024 2022.6 2019.5 2023 2015.3 C +2022.2 2013.9 2021.7 2015.4 2021.3 2014.8 C +2020.4 2015.3 2021 2016.5 2020.8 2017.2 C +2021.4 2016.6 2021.1 2017.8 2021.6 2018 C +2022 2026.4 2019.6 2029.8 2021.8 2037.2 C +2021.7 2038.4 2020.5 2039.1 2019.2 2038.1 C +2016.5 2036.5 2017.5 2033.8 2014.8 2033.3 C +2014.9 2032 2012.6 2033 2013.2 2030.7 C +2011.9 2030.8 2011.2 2030.1 2010.8 2029.2 C +2010.8 2029.1 2010.8 2028.2 2010.8 2028.8 C +2010 2028.8 2010.4 2026.5 2008.6 2027.3 C +2007.9 2026.6 2007.3 2025.9 2007.9 2027.1 C +2009.7 2028 2010 2030.1 2012.2 2030.9 C +2012.9 2032.1 2013.7 2033.6 2015.1 2033.6 C +2015.7 2035.1 2016.9 2036.7 2018.4 2038.4 C +2019.8 2039.3 2022 2039.4 2021.6 2041.5 C +2021.9 2040.7 2022.9 2041.1 2023 2040.3 C +[0 1 1 0.23] vc +f +S +n +2022.5 2024.9 m +2022.5 2023.5 2022.5 2022.2 2022.5 2020.8 C +2022.5 2022.2 2022.5 2023.5 2022.5 2024.9 C +[0 1 1 0.36] vc +f +S +n +1983.2 2022.8 m +1982.4 2022.5 1982.1 2021.6 1981.2 2022.3 C +1981.1 2022.9 1980.5 2024 1981 2024.2 C +1981.8 2024.6 1982.9 2024.4 1983.2 2022.8 C +[0 1 1 0.23] vc +f +S +n +1931.1 2019.9 m +1929.6 2017.7 1932 2015.7 1930.8 2013.9 C +1931.1 2013 1930.3 2011 1930.6 2009.3 C +1930.6 2010.3 1929.8 2010 1929.2 2010 C +1928 2010.3 1928.8 2008.1 1928.2 2007.6 C +1929.1 2007.8 1929.3 2006.3 1930.1 2006.9 C +1930.3 2009.8 1932.2 2004.8 1932.3 2008.6 C +1932.7 2008 1932.8 2009 1932.8 2009.3 C +1932.8 2009.6 1932.8 2009.8 1932.8 2010 C +1933.2 2009 1932.7 2006.6 1934 2005.7 C +1932.7 2004.6 1934.3 2004.6 1934.2 2004 C +1935.8 2003.7 1937 2003.6 1938.5 2004 C +1938.5 2004.5 1939.1 2005.4 1938.3 2006 C +1940.7 2005.7 1937.4 2001.3 1939.7 2001.4 C +1939.5 2001.4 1938.6 1998.8 1937.1 1999.2 C +1936.3 1999.1 1936.2 1997.1 1936.1 1998.5 C +1934.7 2000.1 1932.9 1998.2 1931.6 1999.5 C +1931.3 1998.9 1930.9 1998.5 1931.1 1997.8 C +1931.6 1998.2 1931.3 1996.6 1932 1996.1 C +1933.2 1995.5 1934.3 1996.4 1935.2 1995.4 C +1935.5 1996.5 1936.3 1996.1 1935.6 1995.2 C +1934.7 1994.5 1932.5 1995.3 1932 1995.4 C +1930.5 1995.3 1931.9 1996.5 1930.8 1996.4 C +1931.2 1997.9 1929.5 1998.3 1928.9 1998.5 C +1928.1 1997.9 1927.1 1998 1926 1998 C +1925.3 1999.2 1924.8 2001.4 1923.2 2001.4 C +1922.6 2000.9 1921 2000.9 1920.3 2000.9 C +1919.7 2001.9 1919.6 2003.5 1918.1 2004 C +1916.9 2004.1 1915.8 2002 1915.2 2003.8 C +1916.7 2004 1917.6 2004.9 1919.6 2004.5 C +1920.7 2005.2 1919.4 2006.3 1919.8 2006.9 C +1919.2 2006.9 1917.7 2007.8 1917.2 2008.6 C +1917.8 2011.6 1919.8 2007.8 1920.5 2010.5 C +1920.8 2011.3 1919.3 2011.6 1920.5 2012 C +1920.8 2012.3 1924 2011.8 1923.2 2014.1 C +1922.6 2013.6 1924.1 2016.1 1924.1 2015.1 C +1925.1 2015.4 1925.9 2015 1926.3 2016.5 C +1926.2 2016.6 1926 2016.8 1925.8 2016.8 C +1925.9 2017.2 1926.2 2017.8 1926.8 2018.2 C +1927.1 2017.6 1927.7 2018 1928.4 2018.2 C +1929.7 2020.1 1927.1 2019.5 1929.4 2021.1 C +1929.9 2020.7 1931.1 2020 1931.1 2019.9 C +[0.21 0.21 0 0] vc +f +S +n +1937.1 2020.8 m +1937.1 2018.3 1937.1 2015.7 1937.1 2013.2 C +1937.1 2015.7 1937.1 2018.3 1937.1 2020.8 C +[0 1 1 0.36] vc +f +S +n +2020.4 2012.2 m +2019.8 2012 2019.3 2011.5 2018.7 2011.7 C +2017.9 2012.1 2018.1 2014.1 2018.4 2014.4 C +2019.6 2014.8 2021.4 2013.7 2020.4 2012.2 C +[0 1 1 0.23] vc +f +S +n +1976 2013.9 m +1973.8 2011.5 1971.6 2009.1 1969.5 2006.7 C +1971.6 2009.1 1973.8 2011.5 1976 2013.9 C +[0 1 1 0.36] vc +f +S +n +1995.4 2012.7 m +1996.1 2010.3 1993.8 2006.2 1997.3 2005.7 C +1998.9 2005.4 2000 2003.7 2001.4 2003.1 C +2003.9 2003.1 2005.3 2001.3 2006.9 1999.7 C +2004.5 2003.5 2000 2002.2 1997.6 2005.7 C +1996.5 2005.9 1994.8 2006.1 1995.2 2007.6 C +1995.7 2009.4 1995.2 2011.6 1994.7 2012.9 C +1992 2015.8 1987.8 2015.7 1985.3 2018.7 C +1988.3 2016.3 1992.3 2015.3 1995.4 2012.7 C +[0.18 0.18 0 0.78] vc +f +S +n +1995.6 2012.4 m +1995.6 2011.2 1995.6 2010 1995.6 2008.8 C +1995.6 2010 1995.6 2011.2 1995.6 2012.4 C +[0 1 1 0.36] vc +f +S +n +vmrs +2017.7 2009.6 m +2016.9 2009.3 2016.7 2008.4 2015.8 2009.1 C +2014.2 2010.6 2016 2010.6 2016.5 2011.5 C +2017.2 2010.9 2018.1 2010.8 2017.7 2009.6 C +[0 1 1 0.23] vc +f +0.4 w +2 J +2 M +S +n +2014.4 2006.4 m +2013.5 2006.8 2012.1 2005.6 2012 2006.7 C +2013 2007.3 2011.9 2009.2 2012.9 2008.4 C +2014.2 2008.3 2014.6 2007.8 2014.4 2006.4 C +f +S +n +1969 2006.4 m +1966.5 2003.8 1964 2001.2 1961.6 1998.5 C +1964 2001.2 1966.5 2003.8 1969 2006.4 C +[0 1 1 0.36] vc +f +S +n +2012 2005.2 m +2012.2 2004.2 2011.4 2003.3 2010.3 2003.3 C +2009 2003.6 2010 2004.7 2009.6 2004.8 C +2009.3 2005.7 2011.4 2006.7 2012 2005.2 C +[0 1 1 0.23] vc +f +S +n +1962.8 1995.2 m +1961.7 1994.4 1960.6 1993.7 1959.4 1994 C +1959.5 1994.9 1957.5 1994.1 1956.8 1994.7 C +1955.9 1995.5 1956.7 1997 1955.1 1997.3 C +1956.9 1996.7 1956.8 1994 1959.2 1994.7 C +1961.1 1991 1968.9 2003.2 1962.8 1995.2 C +[0 1 1 0.36] vc +f +S +n +1954.6 1995.6 m +1955.9 1994.7 1955.1 1989.8 1955.3 1988 C +1954.5 1988.3 1954.9 1986.6 1954.4 1986 C +1955.7 1989.2 1953.9 1991.1 1954.8 1994.2 C +1954.5 1995.9 1953.5 1995.3 1953.9 1997.3 C +1955.3 1998.3 1953.2 1995.5 1954.6 1995.6 C +f +S +n +1992.3 2011 m +1992.5 2006.7 1992 2000.3 1991.8 1997.6 C +1992.2 1997.9 1992 1998.5 1992 1999 C +1992.1 1994.7 1991.9 1990.2 1992.3 1986 C +1991.4 1984.6 1991 1986.1 1990.6 1985.6 C +1989.7 1986 1990.3 1987.2 1990.1 1988 C +1990.7 1987.4 1990.4 1988.5 1990.8 1988.7 C +1991.3 1997.1 1988.9 2000.6 1991.1 2007.9 C +1991 2009.1 1989.8 2009.9 1988.4 2008.8 C +1985.7 2007.2 1986.8 2004.5 1984.1 2004 C +1984.2 2002.7 1981.9 2003.7 1982.4 2001.4 C +1981.2 2001.5 1980.5 2000.8 1980 2000 C +1980 1999.8 1980 1998.9 1980 1999.5 C +1979.3 1999.5 1979.7 1997.2 1977.9 1998 C +1977.2 1997.3 1976.6 1996.7 1977.2 1997.8 C +1979 1998.7 1979.3 2000.8 1981.5 2001.6 C +1982.2 2002.8 1983 2004.3 1984.4 2004.3 C +1985 2005.8 1986.2 2007.5 1987.7 2009.1 C +1989 2010 1991.3 2010.2 1990.8 2012.2 C +1991.2 2011.4 1992.2 2011.8 1992.3 2011 C +[0 1 1 0.23] vc +f +S +n +1991.8 1995.6 m +1991.8 1994.3 1991.8 1992.9 1991.8 1991.6 C +1991.8 1992.9 1991.8 1994.3 1991.8 1995.6 C +[0 1 1 0.36] vc +f +S +n +1959.2 1994.2 m +1958.8 1993.3 1960.7 1993.9 1961.1 1993.7 C +1961.5 1993.9 1961.2 1994.4 1961.8 1994.2 C +1960.9 1994 1960.8 1992.9 1959.9 1992.5 C +1959.6 1993.5 1958.3 1993.5 1958.2 1994.2 C +1958.1 1994.1 1958 1994 1958 1994 C +1957.2 1994.9 1958 1993.4 1956.8 1993 C +1955.6 1992.5 1956 1991 1956.3 1989.9 C +1956.5 1989.8 1956.6 1990 1956.8 1990.1 C +1957.1 1989 1956 1989.1 1955.8 1988.2 C +1955.1 1990.4 1956.2 1995 1954.8 1995.9 C +1954.1 1995.5 1954.5 1996.5 1954.4 1997.1 C +1955 1996.8 1954.8 1997.4 1955.6 1996.8 C +1956 1996 1956.3 1993.2 1958.7 1994.2 C +1958.9 1994.2 1959.7 1994.2 1959.2 1994.2 C +[0 1 1 0.23] vc +f +S +n +1958.2 1994 m +1958.4 1993.5 1959.7 1993.1 1959.9 1992 C +1959.7 1992.5 1959.3 1992 1959.4 1991.8 C +1959.4 1991.6 1959.4 1990.8 1959.4 1991.3 C +1959.2 1991.8 1958.8 1991.3 1958.9 1991.1 C +1958.9 1990.5 1958 1990.3 1957.5 1989.9 C +1956.8 1989.5 1956.9 1991 1956.3 1990.1 C +1956.7 1991 1955.4 1992.1 1956.5 1992.3 C +1956.8 1993.5 1958.3 1992.9 1957.2 1994 C +1957.8 1994.3 1958.1 1992.4 1958.2 1994 C +[0 0.5 0.5 0.2] vc +f +S +n +vmrs +1954.4 1982.7 m +1956.1 1982.7 1954.1 1982.5 1953.9 1982.9 C +1953.9 1983.7 1953.7 1984.7 1954.1 1985.3 C +1954.4 1984.2 1953.6 1983.6 1954.4 1982.7 C +[0 1 1 0.36] vc +f +0.4 w +2 J +2 M +S +n +1989.6 1982.9 m +1989.1 1982.7 1988.6 1982.3 1988 1982.4 C +1987.2 1982.8 1987.4 1984.8 1987.7 1985.1 C +1988.9 1985.6 1990.7 1984.4 1989.6 1982.9 C +[0 1 1 0.23] vc +f +S +n +1987 1980.3 m +1986.2 1980 1986 1979.1 1985.1 1979.8 C +1983.5 1981.4 1985.3 1981.4 1985.8 1982.2 C +1986.5 1981.7 1987.4 1981.5 1987 1980.3 C +f +S +n +1983.6 1977.2 m +1982.7 1977.5 1981.4 1976.3 1981.2 1977.4 C +1982.3 1978 1981.2 1979.9 1982.2 1979.1 C +1983.5 1979 1983.9 1978.5 1983.6 1977.2 C +f +S +n +1981.2 1976 m +1981.5 1974.9 1980.6 1974 1979.6 1974 C +1978.3 1974.3 1979.3 1975.4 1978.8 1975.5 C +1978.6 1976.4 1980.7 1977.4 1981.2 1976 C +f +S +n +1972.1 2082.3 m +1971.8 2081.8 1971.3 2080.9 1971.2 2080.1 C +1971.1 2072.9 1971.3 2064.6 1970.9 2058.3 C +1970.3 2058.5 1970.1 2057.7 1969.7 2058.5 C +1970.6 2058.5 1969.7 2059 1970.2 2059.2 C +1970.2 2065.4 1970.2 2072.4 1970.2 2077.7 C +1971.1 2078.9 1970.6 2078.9 1970.4 2079.9 C +1969.2 2080.2 1968.2 2080.4 1967.3 2079.6 C +1966.8 2077.8 1963.4 2076.3 1963.5 2075.1 C +1961.5 2075.5 1962 2071.5 1959.6 2072 C +1959.2 2070 1956.5 2069.3 1955.8 2067.6 C +1956 2068.4 1955.3 2069.7 1956.5 2069.8 C +1958.6 2068.9 1958.1 2073.5 1960.1 2072.4 C +1960.7 2075.9 1964.7 2074.6 1964.2 2078 C +1967.2 2078.6 1967.9 2081.6 1970.7 2080.6 C +1970.3 2081.1 1971.5 2081.2 1971.9 2082.3 C +1967.2 2084.3 1962.9 2087.1 1958.2 2089 C +1962.9 2087 1967.4 2084.4 1972.1 2082.3 C +[0 0.2 1 0] vc +f +S +n +1971.9 2080.1 m +1971.9 2075.1 1971.9 2070 1971.9 2065 C +1971.9 2070 1971.9 2075.1 1971.9 2080.1 C +[0 1 1 0.23] vc +f +S +n +2010.8 2050.6 m +2013.2 2049 2010.5 2050.1 2010.5 2051.3 C +2010.5 2057.7 2010.5 2064.1 2010.5 2070.5 C +2008.7 2072.4 2006 2073.3 2003.6 2074.4 C +2016.4 2073.7 2008 2058.4 2010.8 2050.6 C +[0.4 0.4 0 0] vc +f +S +n +2006.4 2066.9 m +2006.4 2061.9 2006.4 2056.8 2006.4 2051.8 C +2006.4 2056.8 2006.4 2061.9 2006.4 2066.9 C +[0 1 1 0.23] vc +f +S +n +1971.9 2060.7 m +1972.2 2060.3 1971.4 2068.2 1972.4 2061.9 C +1971.8 2061.6 1972.4 2060.9 1971.9 2060.7 C +f +S +n +vmrs +1986.5 2055.2 m +1987.5 2054.3 1986.3 2053.4 1986 2052.8 C +1983.8 2052.7 1983.6 2050.1 1981.7 2049.6 C +1981.2 2048.7 1980.8 2047 1980.3 2046.8 C +1978.5 2047 1978 2044.6 1976.7 2043.9 C +1974 2044.4 1972 2046.6 1969.2 2047 C +1969 2047.2 1968.8 2047.5 1968.5 2047.7 C +1970.6 2049.6 1973.1 2051.3 1974.3 2054.2 C +1975.7 2054.5 1977 2055.2 1976.4 2057.1 C +1976.7 2058 1975.5 2058.5 1976 2059.5 C +1979.2 2058 1983 2056.6 1986.5 2055.2 C +[0 0.5 0.5 0.2] vc +f +0.4 w +2 J +2 M +S +n +1970.2 2054.2 m +1971.5 2055.3 1972.5 2056.8 1972.1 2058.3 C +1972.8 2056.5 1971.6 2055.6 1970.2 2054.2 C +[0 1 1 0.23] vc +f +S +n +1992 2052.5 m +1992 2053.4 1992.2 2054.4 1991.8 2055.2 C +1992.2 2054.4 1992 2053.4 1992 2052.5 C +f +S +n +1957.2 2053 m +1958.1 2052.6 1959 2052.2 1959.9 2051.8 C +1959 2052.2 1958.1 2052.6 1957.2 2053 C +f +S +n +2006.4 2047.5 m +2006.8 2047.1 2006 2055 2006.9 2048.7 C +2006.4 2048.4 2007 2047.7 2006.4 2047.5 C +f +S +n +2004.8 2041 m +2006.1 2042.1 2007.1 2043.6 2006.7 2045.1 C +2007.3 2043.3 2006.2 2042.4 2004.8 2041 C +f +S +n +1976 2039.8 m +1975.6 2039.3 1975.2 2038.4 1975 2037.6 C +1974.9 2030.4 1975.2 2022.1 1974.8 2015.8 C +1974.2 2016 1974 2015.3 1973.6 2016 C +1974.4 2016 1973.5 2016.5 1974 2016.8 C +1974 2022.9 1974 2030 1974 2035.2 C +1974.9 2036.4 1974.4 2036.4 1974.3 2037.4 C +1973.1 2037.7 1972 2037.9 1971.2 2037.2 C +1970.6 2035.3 1967.3 2033.9 1967.3 2032.6 C +1965.3 2033 1965.9 2029.1 1963.5 2029.5 C +1963 2027.6 1960.4 2026.8 1959.6 2025.2 C +1959.8 2025.9 1959.2 2027.2 1960.4 2027.3 C +1962.5 2026.4 1961.9 2031 1964 2030 C +1964.6 2033.4 1968.5 2032.1 1968 2035.5 C +1971 2036.1 1971.8 2039.1 1974.5 2038.1 C +1974.2 2038.7 1975.3 2038.7 1975.7 2039.8 C +1971 2041.8 1966.7 2044.6 1962 2046.5 C +1966.8 2044.5 1971.3 2041.9 1976 2039.8 C +[0 0.2 1 0] vc +f +S +n +1975.7 2037.6 m +1975.7 2032.6 1975.7 2027.6 1975.7 2022.5 C +1975.7 2027.6 1975.7 2032.6 1975.7 2037.6 C +[0 1 1 0.23] vc +f +S +n +1992 2035.5 m +1992 2034.2 1992 2032.9 1992 2031.6 C +1992 2032.9 1992 2034.2 1992 2035.5 C +f +S +n +2015.3 2036 m +2015.4 2034.1 2013.3 2034 2012.9 2033.3 C +2011.5 2031 2009.3 2029.4 2007.4 2028 C +2006.9 2027.1 2006.6 2023.8 2005 2024.9 C +2004 2024.9 2002.9 2024.9 2001.9 2024.9 C +2001.4 2026.5 2001 2028.4 2003.8 2028.3 C +2006.6 2030.4 2008.9 2033.7 2011.2 2036.2 C +2011.8 2036.4 2012.9 2035.8 2012.9 2036.7 C +2013 2035.5 2015.3 2037.4 2015.3 2036 C +[0 0 0 0] vc +f +S +n +vmrs +2009.1 2030.4 m +2009.1 2029 2007.5 2029.4 2006.9 2028.3 C +2007.2 2027.1 2006.5 2025.5 2005.7 2024.7 C +2004.6 2025.1 2003.1 2024.9 2001.9 2024.9 C +2001.8 2026.2 2000.9 2027 2002.4 2028 C +2004.5 2027.3 2004.9 2029.4 2006.9 2029 C +2007 2030.2 2007.6 2030.7 2008.4 2031.4 C +2008.8 2031.5 2009.1 2031.1 2009.1 2030.4 C +[0 0 0 0.18] vc +f +0.4 w +2 J +2 M +S +n +2003.8 2029.5 m +2003 2029.4 2001.9 2029.1 2002.4 2030.4 C +2003.1 2031.3 2005.2 2030.3 2003.8 2029.5 C +[0 1 1 0.23] vc +f +S +n +1999.2 2025.2 m +1999.1 2025.6 1998 2025.7 1998.8 2026.6 C +2000.9 2028.5 1999.5 2023.4 1999.2 2025.2 C +f +S +n +2007.6 2024.2 m +2007.6 2022.9 2008.4 2024.2 2007.6 2022.8 C +2007.6 2017.5 2007.8 2009.1 2007.4 2003.8 C +2007.9 2003.7 2008.7 2002.8 2009.1 2002.1 C +2009.6 2000.8 2008.3 2000.8 2007.9 2000.2 C +2004.9 2000 2008.9 2001.3 2007.2 2002.1 C +2006.7 2007.7 2007 2015.1 2006.9 2021.1 C +2006.7 2022.1 2005.4 2022.8 2006.2 2023.5 C +2006.6 2023.1 2008 2025.9 2007.6 2024.2 C +f +S +n +1989.9 2023.5 m +1989.5 2022.6 1991.4 2023.2 1991.8 2023 C +1992.2 2023.2 1991.9 2023.7 1992.5 2023.5 C +1991.6 2023.2 1991.6 2022.2 1990.6 2021.8 C +1990.4 2022.8 1989 2022.8 1988.9 2023.5 C +1988.5 2023 1988.7 2022.6 1988.7 2023.5 C +1989.1 2023.5 1990.2 2023.5 1989.9 2023.5 C +f +[0 0.5 0.5 0.2] vc +S +n +2003.3 2023.5 m +2003.1 2023.3 2003.1 2023.2 2003.3 2023 C +2003.7 2023.1 2003.9 2022.9 2003.8 2022.5 C +2003.4 2022.2 2001.2 2022.3 2002.4 2023 C +2002.6 2022.9 2002.7 2023.1 2002.8 2023.2 C +2000.7 2023.7 2003.9 2023.4 2003.3 2023.5 C +[0 1 1 0.23] vc +f +S +n +1986.8 2019.4 m +1987.8 2019.8 1987.5 2018.6 1987.2 2018 C +1986.2 2017.8 1987.3 2020.5 1986.3 2019.2 C +1986.3 2017.7 1986.3 2020.6 1986.3 2021.3 C +1988.5 2023.1 1985.6 2020.3 1986.8 2019.4 C +f +S +n +1975.7 2018.2 m +1976.1 2017.8 1975.2 2025.7 1976.2 2019.4 C +1975.7 2019.2 1976.3 2018.4 1975.7 2018.2 C +f +S +n +1974 2011.7 m +1975.4 2012.8 1976.4 2014.3 1976 2015.8 C +1976.6 2014 1975.5 2013.1 1974 2011.7 C +f +S +n +1984.6 2006.7 m +1984.7 2004.8 1982.6 2004.8 1982.2 2004 C +1980.8 2001.7 1978.6 2000.1 1976.7 1998.8 C +1976.1 1997.8 1975.8 1994.5 1974.3 1995.6 C +1973.3 1995.6 1972.2 1995.6 1971.2 1995.6 C +1970.7 1997.2 1970.3 1999.1 1973.1 1999 C +1975.8 2001.2 1978.2 2004.4 1980.5 2006.9 C +1981.1 2007.1 1982.1 2006.5 1982.2 2007.4 C +1982.3 2006.2 1984.5 2008.1 1984.6 2006.7 C +[0 0 0 0] vc +f +S +n +vmrs +1978.4 2001.2 m +1978.4 1999.7 1976.8 2000.1 1976.2 1999 C +1976.5 1997.8 1975.8 1996.2 1975 1995.4 C +1973.9 1995.8 1972.4 1995.6 1971.2 1995.6 C +1971 1997 1970.2 1997.7 1971.6 1998.8 C +1973.8 1998 1974.2 2000.1 1976.2 1999.7 C +1976.3 2000.9 1976.9 2001.4 1977.6 2002.1 C +1978.1 2002.2 1978.4 2001.8 1978.4 2001.2 C +[0 0 0 0.18] vc +f +0.4 w +2 J +2 M +S +n +1973.1 2000.2 m +1972.3 2000.1 1971.2 1999.8 1971.6 2001.2 C +1972.4 2002 1974.5 2001 1973.1 2000.2 C +[0 1 1 0.23] vc +f +S +n +1960.8 1998.5 m +1961.6 1998.2 1962.6 2000.3 1963.2 2000.9 C +1962.3 2000.1 1962.2 1998.7 1960.8 1998.5 C +f +S +n +1968.5 1995.9 m +1968.4 1996.4 1967.3 1996.4 1968 1997.3 C +1970.1 1999.2 1968.8 1994.1 1968.5 1995.9 C +f +S +n +1976.9 1994.9 m +1976.9 1993.7 1977.6 1994.9 1976.9 1993.5 C +1976.9 1988.2 1977.1 1979.8 1976.7 1974.5 C +1977.2 1974.5 1978 1973.5 1978.4 1972.8 C +1978.8 1971.5 1977.6 1971.5 1977.2 1970.9 C +1974.2 1970.7 1978.2 1972 1976.4 1972.8 C +1976 1978.4 1976.3 1985.8 1976.2 1991.8 C +1976 1992.8 1974.6 1993.5 1975.5 1994.2 C +1975.9 1993.8 1977.3 1996.6 1976.9 1994.9 C +f +S +n +1972.6 1994.2 m +1972.4 1994 1972.4 1993.9 1972.6 1993.7 C +1973 1993.8 1973.1 1993.7 1973.1 1993.2 C +1972.7 1992.9 1970.5 1993.1 1971.6 1993.7 C +1971.9 1993.7 1972 1993.8 1972.1 1994 C +1970 1994.4 1973.1 1994.1 1972.6 1994.2 C +f +S +n +1948.1 2093.8 m +1947 2092.7 1945.9 2091.6 1944.8 2090.4 C +1945.9 2091.6 1947 2092.7 1948.1 2093.8 C +[0 0.4 1 0] vc +f +S +n +1953.4 2091.4 m +1954.8 2090.7 1956.3 2090 1957.7 2089.2 C +1956.3 2090 1954.8 2090.7 1953.4 2091.4 C +[0 0.2 1 0] vc +f +S +n +1954.1 2091.4 m +1956.6 2089.6 1957.2 2089.6 1954.1 2091.4 C +[0 0.4 1 0] vc +f +S +n +1962.3 2087.3 m +1963.7 2086.6 1965.2 2085.9 1966.6 2085.2 C +1965.2 2085.9 1963.7 2086.6 1962.3 2087.3 C +f +S +n +vmrs +1967.1 2084.9 m +1968.3 2084.4 1969.7 2083.8 1970.9 2083.2 C +1969.7 2083.8 1968.3 2084.4 1967.1 2084.9 C +[0 0.4 1 0] vc +f +0.4 w +2 J +2 M +S +n +1982.7 2080.6 m +1981.5 2079.5 1980.5 2078.4 1979.3 2077.2 C +1980.5 2078.4 1981.5 2079.5 1982.7 2080.6 C +f +S +n +1988 2078.2 m +1989.4 2077.5 1990.8 2076.8 1992.3 2076 C +1990.8 2076.8 1989.4 2077.5 1988 2078.2 C +[0 0.2 1 0] vc +f +S +n +1988.7 2078.2 m +1991.1 2076.4 1991.8 2076.4 1988.7 2078.2 C +[0 0.4 1 0] vc +f +S +n +1976.2 2063.8 m +1978.6 2062.2 1976 2063.3 1976 2064.5 C +1976.1 2067.8 1975.5 2071.4 1976.4 2074.4 C +1975.7 2071.1 1975.9 2067.2 1976.2 2063.8 C +f +S +n +1996.8 2074.1 m +1998.3 2073.4 1999.7 2072.7 2001.2 2072 C +1999.7 2072.7 1998.3 2073.4 1996.8 2074.1 C +f +S +n +2001.6 2071.7 m +2002.9 2071.2 2004.2 2070.6 2005.5 2070 C +2004.2 2070.6 2002.9 2071.2 2001.6 2071.7 C +f +S +n +1981.5 2060.7 m +1980.2 2061.2 1978.9 2061.5 1977.9 2062.6 C +1978.9 2061.5 1980.2 2061.2 1981.5 2060.7 C +f +S +n +1982 2060.4 m +1982.7 2060.1 1983.6 2059.8 1984.4 2059.5 C +1983.6 2059.8 1982.7 2060.1 1982 2060.4 C +f +S +n +1952 2051.3 m +1950.8 2050.2 1949.7 2049.1 1948.6 2048 C +1949.7 2049.1 1950.8 2050.2 1952 2051.3 C +f +S +n +vmrs +1977.4 2047.7 m +1975.8 2047.8 1974.8 2046.1 1974.5 2045.3 C +1974.9 2044.4 1976 2044.5 1976.7 2044.8 C +1977.9 2045 1977 2048.4 1979.3 2047.5 C +1979.9 2047.5 1980.8 2048.6 1979.8 2049.2 C +1978.2 2050.4 1980.8 2049.5 1980.3 2049.4 C +1981.4 2049.8 1980.3 2048.4 1980.3 2048 C +1979.8 2047.5 1979 2046.6 1978.4 2046.5 C +1977.3 2045.9 1977.2 2043.3 1975.2 2044.6 C +1974.7 2045.3 1973.6 2045 1973.3 2045.8 C +1975 2046.3 1975.8 2049.8 1978.1 2049.4 C +1978.4 2050.9 1978.7 2048.5 1977.9 2049.2 C +1977.7 2048.7 1977.2 2047.8 1977.4 2047.7 C +[0 0.5 0.5 0.2] vc +f +0.4 w +2 J +2 M +S +n +1957.2 2048.9 m +1958.7 2048.2 1960.1 2047.5 1961.6 2046.8 C +1960.1 2047.5 1958.7 2048.2 1957.2 2048.9 C +[0 0.2 1 0] vc +f +S +n +1958 2048.9 m +1960.4 2047.1 1961.1 2047.1 1958 2048.9 C +[0 0.4 1 0] vc +f +S +n +1966.1 2044.8 m +1967.6 2044.1 1969 2043.4 1970.4 2042.7 C +1969 2043.4 1967.6 2044.1 1966.1 2044.8 C +f +S +n +1970.9 2042.4 m +1972.2 2041.9 1973.5 2041.3 1974.8 2040.8 C +1973.5 2041.3 1972.2 2041.9 1970.9 2042.4 C +f +S +n +2012 2034.5 m +2010.4 2034.6 2009.3 2032.9 2009.1 2032.1 C +2009.4 2031 2010.3 2031.3 2011.2 2031.6 C +2012.5 2031.8 2011.6 2035.2 2013.9 2034.3 C +2014.4 2034.3 2015.4 2035.4 2014.4 2036 C +2012.7 2037.2 2015.3 2036.3 2014.8 2036.2 C +2015.9 2036.6 2014.8 2035.2 2014.8 2034.8 C +2014.4 2034.3 2013.6 2033.4 2012.9 2033.3 C +2011.5 2031 2009.3 2029.4 2007.4 2028 C +2007.5 2026.5 2007.3 2027.9 2007.2 2028.3 C +2007.9 2028.8 2008.7 2029.1 2009.3 2030 C +2009.6 2030.7 2009 2031.9 2008.4 2031.6 C +2006.7 2031 2007.7 2028 2005 2028.8 C +2004.8 2028.6 2004.3 2028.2 2003.8 2028.3 C +2006.6 2030.4 2008.9 2033.7 2011.2 2036.2 C +2011.8 2036.4 2012.9 2035.8 2012.9 2036.7 C +2012.7 2036.1 2011.8 2035 2012 2034.5 C +[0 0.5 0.5 0.2] vc +f +S +n +1981.2 2005.2 m +1979.7 2005.3 1978.6 2003.6 1978.4 2002.8 C +1978.7 2001.8 1979.6 2002.1 1980.5 2002.4 C +1981.8 2002.5 1980.9 2005.9 1983.2 2005 C +1983.7 2005.1 1984.7 2006.1 1983.6 2006.7 C +1982 2007.9 1984.6 2007 1984.1 2006.9 C +1985.2 2007.3 1984.1 2006 1984.1 2005.5 C +1983.6 2005 1982.9 2004.1 1982.2 2004 C +1980.8 2001.7 1978.6 2000.1 1976.7 1998.8 C +1976.7 1997.2 1976.6 1998.6 1976.4 1999 C +1977.2 1999.5 1978 1999.8 1978.6 2000.7 C +1978.8 2001.5 1978.3 2002.7 1977.6 2002.4 C +1976 2001.8 1977 1998.7 1974.3 1999.5 C +1974.1 1999.3 1973.6 1998.9 1973.1 1999 C +1975.8 2001.2 1978.2 2004.4 1980.5 2006.9 C +1981.1 2007.1 1982.1 2006.5 1982.2 2007.4 C +1982 2006.8 1981.1 2005.7 1981.2 2005.2 C +f +S +n +1966.8 1976.4 m +1969.4 1973 1974.4 1974.6 1976.2 1970.4 C +1972.7 1974 1968 1975.1 1964 1977.4 C +1960.9 1979.9 1957.1 1981.8 1953.9 1982.7 C +1958.4 1981.1 1962.6 1978.8 1966.8 1976.4 C +[0.18 0.18 0 0.78] vc +f +S +n +1948.4 2093.8 m +1949.8 2093.1 1951.2 2092.5 1952.7 2091.9 C +1951.2 2092.5 1949.8 2093.1 1948.4 2093.8 C +[0 0.2 1 0] vc +f +S +n +1948.1 2093.6 m +1947.3 2092.8 1946.5 2091.9 1945.7 2091.2 C +1946.5 2091.9 1947.3 2092.8 1948.1 2093.6 C +f +S +n +vmrs +1942.1 2087.8 m +1943.5 2088.4 1944.3 2089.5 1945.2 2090.7 C +1944.8 2089.3 1943.3 2088.3 1942.1 2087.8 C +[0 0.2 1 0] vc +f +0.4 w +2 J +2 M +S +n +1933.5 2078.4 m +1933.5 2078 1933.2 2079 1933.7 2079.4 C +1935 2080.4 1936.2 2081.3 1937.1 2082.8 C +1936.7 2080.7 1933.7 2080.7 1933.5 2078.4 C +f +S +n +1982.9 2080.6 m +1984.4 2079.9 1985.8 2079.3 1987.2 2078.7 C +1985.8 2079.3 1984.4 2079.9 1982.9 2080.6 C +f +S +n +1982.7 2080.4 m +1981.9 2079.6 1981.1 2078.7 1980.3 2078 C +1981.1 2078.7 1981.9 2079.6 1982.7 2080.4 C +f +S +n +1977.4 2075.1 m +1977.9 2075.3 1979.1 2076.4 1979.8 2077.5 C +1979 2076.8 1978.7 2075.1 1977.4 2075.1 C +f +S +n +1952.2 2051.3 m +1953.6 2050.7 1955.1 2050.1 1956.5 2049.4 C +1955.1 2050.1 1953.6 2050.7 1952.2 2051.3 C +f +S +n +1952 2051.1 m +1951.2 2050.3 1950.3 2049.5 1949.6 2048.7 C +1950.3 2049.5 1951.2 2050.3 1952 2051.1 C +f +S +n +1946 2045.3 m +1947.3 2045.9 1948.1 2047 1949.1 2048.2 C +1948.6 2046.8 1947.1 2045.8 1946 2045.3 C +f +S +n +1937.3 2036 m +1937.4 2035.5 1937 2036.5 1937.6 2036.9 C +1938.8 2037.9 1940.1 2038.8 1940.9 2040.3 C +1940.6 2038.2 1937.6 2038.2 1937.3 2036 C +f +S +n +1935.2 2073.2 m +1936.4 2069.9 1935.8 2061.8 1935.6 2056.4 C +1935.8 2055.9 1936.3 2055.7 1936.1 2055.2 C +1935.7 2054.7 1935 2055 1934.4 2054.9 C +1934.4 2061.5 1934.4 2068.7 1934.4 2074.6 C +1935.7 2075.1 1936 2073.7 1935.2 2073.2 C +[0 0.01 1 0] vc +f +S +n +vmrs +1939 2030.7 m +1940.3 2027.4 1939.7 2019.3 1939.5 2013.9 C +1939.7 2013.5 1940.1 2013.2 1940 2012.7 C +1939.5 2012.3 1938.8 2012.5 1938.3 2012.4 C +1938.3 2019 1938.3 2026.2 1938.3 2032.1 C +1939.5 2032.7 1939.8 2031.2 1939 2030.7 C +[0 0.01 1 0] vc +f +0.4 w +2 J +2 M +S +n +1975.2 2077.2 m +1975.3 2077.3 1975.4 2077.4 1975.5 2077.5 C +1974.7 2073.2 1974.9 2067.5 1975.2 2063.6 C +1975.4 2064 1974.6 2063.9 1974.8 2064.3 C +1974.9 2069.9 1974.3 2076.5 1975.2 2081.1 C +1974.9 2079.9 1974.9 2078.4 1975.2 2077.2 C +[0.92 0.92 0 0.67] vc +f +S +n +1930.8 2067.4 m +1931.5 2070.1 1929.6 2072.1 1930.6 2074.6 C +1931 2072.6 1930.8 2069.8 1930.8 2067.4 C +f +S +n +2010 2050.1 m +2009.8 2050.5 2009.5 2050.9 2009.3 2051.1 C +2009.5 2056.7 2008.9 2063.3 2009.8 2067.9 C +2009.5 2062.1 2009.3 2054.7 2010 2050.1 C +f +S +n +1930.1 2060.9 m +1929.3 2057.1 1930.7 2054.8 1929.9 2051.3 C +1930.2 2050.2 1931.1 2049.6 1931.8 2049.2 C +1931.4 2049.6 1930.4 2049.5 1930.1 2050.1 C +1928.4 2054.8 1933.4 2063.5 1925.3 2064.3 C +1927.2 2063.9 1928.5 2062.1 1930.1 2060.9 C +[0.07 0.06 0 0.58] vc +f +S +n +1929.6 2061.2 m +1929.6 2057.6 1929.6 2054.1 1929.6 2050.6 C +1930 2049.9 1930.5 2049.4 1931.1 2049.2 C +1930 2048.6 1930.5 2050.2 1929.4 2049.6 C +1928 2054.4 1932.8 2063 1925.3 2064 C +1926.9 2063.3 1928.3 2062.4 1929.6 2061.2 C +[0.4 0.4 0 0] vc +f +S +n +1930.8 2061.6 m +1930.5 2058 1931.6 2054 1930.8 2051.3 C +1930.3 2054.5 1930.9 2058.5 1930.4 2061.9 C +1930.5 2061.2 1931 2062.2 1930.8 2061.6 C +[0.92 0.92 0 0.67] vc +f +S +n +1941.2 2045.1 m +1939.7 2042.6 1937.3 2041.2 1935.4 2039.3 C +1934.2 2040 1933.7 2036.4 1934 2039.3 C +1934.9 2040.1 1936.1 2039.9 1936.8 2040.8 C +1935.3 2044.2 1942.3 2041.7 1939.5 2046 C +1937.1 2048.5 1940.5 2045.6 1941.2 2045.1 C +f +S +n +1910 2045.8 m +1910 2039.4 1910 2033 1910 2026.6 C +1910 2033 1910 2039.4 1910 2045.8 C +f +S +n +1978.8 2022.3 m +1979.1 2021.7 1979.4 2020.4 1978.6 2021.6 C +1978.6 2026.9 1978.6 2033 1978.6 2037.6 C +1979.2 2037 1979.1 2038.2 1979.1 2038.6 C +1978.7 2033.6 1978.9 2026.8 1978.8 2022.3 C +f +S +n +vmrs +2026.1 2041.2 m +2026.1 2034.8 2026.1 2028.3 2026.1 2021.8 C +2026.1 2028.5 2026.3 2035.4 2025.9 2042 C +2024.4 2042.9 2022.9 2044.1 2021.3 2044.8 C +2023.1 2044 2025.1 2042.8 2026.1 2041.2 C +[0.07 0.06 0 0.58] vc +f +0.4 w +2 J +2 M +S +n +2026.4 2021.8 m +2026.3 2028.5 2026.5 2035.4 2026.1 2042 C +2025.6 2042.8 2024.7 2042.7 2024.2 2043.4 C +2024.7 2042.7 2025.5 2042.7 2026.1 2042.2 C +2026.5 2035.5 2026.3 2027.9 2026.4 2021.8 C +[0.4 0.4 0 0] vc +f +S +n +2025.6 2038.4 m +2025.6 2033 2025.6 2027.6 2025.6 2022.3 C +2025.6 2027.6 2025.6 2033 2025.6 2038.4 C +[0.92 0.92 0 0.67] vc +f +S +n +1934 2023.5 m +1934 2024.7 1933.8 2026 1934.2 2027.1 C +1934 2025.5 1934.7 2024.6 1934 2023.5 C +f +S +n +1928.2 2023.5 m +1928 2024.6 1927.4 2023.1 1926.8 2023.2 C +1926.2 2021 1921.4 2019.3 1923.2 2018 C +1922.7 2016.5 1923.2 2019.3 1922.2 2018.2 C +1924.4 2020.4 1926.2 2023.3 1928.9 2024.9 C +1927.9 2024.2 1929.8 2023.5 1928.2 2023.5 C +[0.18 0.18 0 0.78] vc +f +S +n +1934 2019.2 m +1932 2019.6 1930.8 2022.6 1928.7 2021.8 C +1924.5 2016.5 1918.2 2011.8 1914 2006.7 C +1914 2005.7 1914 2004.6 1914 2003.6 C +1913.6 2004.3 1913.9 2005.8 1913.8 2006.9 C +1919 2012.4 1924.1 2016.5 1929.2 2022.3 C +1931 2021.7 1932.2 2019.8 1934 2019.2 C +f +S +n +1928.7 2024.9 m +1926.3 2022.7 1924.1 2020.4 1921.7 2018.2 C +1924.1 2020.4 1926.3 2022.7 1928.7 2024.9 C +[0.65 0.65 0 0.42] vc +f +S +n +1914.3 2006.7 m +1918.7 2011.8 1924.5 2016.4 1928.9 2021.6 C +1924.2 2016.1 1919 2012.1 1914.3 2006.7 C +[0.07 0.06 0 0.58] vc +f +S +n +1924.8 2020.8 m +1921.2 2016.9 1925.6 2022.5 1926 2021.1 C +1924.2 2021 1926.7 2019.6 1924.8 2020.8 C +[0.92 0.92 0 0.67] vc +f +S +n +1934 2018.4 m +1933.2 2014.7 1934.5 2012.3 1933.7 2008.8 C +1934 2007.8 1935 2007.2 1935.6 2006.7 C +1935.3 2007.1 1934.3 2007 1934 2007.6 C +1932.2 2012.3 1937.2 2021 1929.2 2021.8 C +1931.1 2021.4 1932.3 2019.6 1934 2018.4 C +[0.07 0.06 0 0.58] vc +f +S +n +vmrs +1933.5 2018.7 m +1933.5 2015.1 1933.5 2011.7 1933.5 2008.1 C +1933.8 2007.4 1934.3 2006.9 1934.9 2006.7 C +1933.8 2006.1 1934.3 2007.7 1933.2 2007.2 C +1931.9 2012 1936.7 2020.5 1929.2 2021.6 C +1930.7 2020.8 1932.2 2019.9 1933.5 2018.7 C +[0.4 0.4 0 0] vc +f +0.4 w +2 J +2 M +S +n +1934.7 2019.2 m +1934.3 2015.6 1935.4 2011.5 1934.7 2008.8 C +1934.1 2012 1934.7 2016 1934.2 2019.4 C +1934.4 2018.7 1934.8 2019.8 1934.7 2019.2 C +[0.92 0.92 0 0.67] vc +f +S +n +1917.6 2013.6 m +1917.8 2011.1 1916.8 2014.2 1917.2 2012.2 C +1916.3 2012.9 1914.8 2011.8 1914.3 2010.8 C +1914.2 2010.5 1914.4 2010.4 1914.5 2010.3 C +1913.9 2008.8 1913.9 2011.9 1914.3 2012 C +1916.3 2012 1917.6 2013.6 1916.7 2015.6 C +1913.7 2017.4 1919.6 2014.8 1917.6 2013.6 C +f +S +n +1887.2 2015.3 m +1887.2 2008.9 1887.2 2002.5 1887.2 1996.1 C +1887.2 2002.5 1887.2 2008.9 1887.2 2015.3 C +f +S +n +1916.7 2014.4 m +1917 2012.1 1913 2013 1913.8 2010.8 C +1912.1 2009.8 1910.9 2009.4 1910.7 2007.9 C +1910.4 2010.6 1913.4 2010.4 1914 2012.4 C +1914.9 2012.8 1916.6 2012.9 1916.4 2014.4 C +1916.9 2015.1 1914.5 2016.6 1916.2 2015.8 C +1916.4 2015.3 1916.7 2015 1916.7 2014.4 C +[0.65 0.65 0 0.42] vc +f +S +n +1914 2009.3 m +1912.8 2010.9 1909.6 2005.3 1911.9 2009.8 C +1912.3 2009.6 1913.6 2010.2 1914 2009.3 C +[0.92 0.92 0 0.67] vc +f +S +n +1951.2 1998.8 m +1949 1996.4 1951.5 1994 1950.3 1991.8 C +1949.1 1989.1 1954 1982.7 1948.8 1981.2 C +1949.2 1981.5 1951 1982.4 1950.8 1983.6 C +1951.9 1988.6 1947.1 1986.5 1948.1 1990.4 C +1948.5 1990.3 1948.7 1990.7 1948.6 1991.1 C +1949 1992.5 1947.3 1991.9 1948.1 1992.5 C +1947.1 1992.7 1945.7 1993.5 1945.2 1994.7 C +1944.5 1996.8 1947.7 2000.5 1943.8 2001.4 C +1943.4 2002 1943.7 2004 1942.4 2004.5 C +1945.2 2002.2 1948.9 2000.9 1951.2 1998.8 C +f +S +n +1994.9 1993 m +1995.1 1996.5 1994.5 2000.3 1995.4 2003.6 C +1994.5 2000.3 1995.1 1996.5 1994.9 1993 C +f +S +n +1913.8 2003.3 m +1913.8 1996.9 1913.8 1990.5 1913.8 1984.1 C +1913.8 1990.5 1913.8 1996.9 1913.8 2003.3 C +f +S +n +1941.9 1998 m +1940.5 1997.3 1940.7 1999.4 1940.7 2000 C +1942.8 2001.3 1942.6 1998.8 1941.9 1998 C +[0 0 0 0] vc +f +S +n +vmrs +1942.1 1999.2 m +1942.2 1998.9 1941.8 1998.8 1941.6 1998.5 C +1940.4 1998 1940.7 1999.7 1940.7 2000 C +1941.6 2000.3 1942.6 2000.4 1942.1 1999.2 C +[0.92 0.92 0 0.67] vc +f +0.4 w +2 J +2 M +S +n +1940 1997.1 m +1939.8 1996 1939.7 1995.9 1939.2 1995.2 C +1939.1 1995.3 1938.5 1997.9 1937.8 1996.4 C +1938 1997.3 1939.4 1998.6 1940 1997.1 C +f +S +n +1911.2 1995.9 m +1911.2 1991.6 1911.3 1987.2 1911.4 1982.9 C +1911.3 1987.2 1911.2 1991.6 1911.2 1995.9 C +f +S +n +1947.2 1979.1 m +1945.1 1978.8 1944.6 1975.7 1942.4 1975 C +1940.5 1972.6 1942.2 1973.7 1942.4 1975.7 C +1945.8 1975.5 1944.2 1979.8 1947.6 1979.6 C +1948.3 1982.3 1948.5 1980 1947.2 1979.1 C +f +S +n +1939.5 1973.3 m +1940.1 1972.6 1939.8 1974.2 1940.2 1973.1 C +1939.1 1972.8 1938.8 1968.5 1935.9 1969.7 C +1937.4 1969.2 1938.5 1970.6 1939 1971.4 C +1939.2 1972.7 1938.6 1973.9 1939.5 1973.3 C +f +S +n +1975.2 2073.2 m +1975.2 2070.2 1975.2 2067.2 1975.2 2064.3 C +1975.2 2067.2 1975.2 2070.2 1975.2 2073.2 C +[0.18 0.18 0 0.78] vc +f +S +n +1929.9 2065.7 m +1928.1 2065.6 1926 2068.8 1924.1 2066.9 C +1918.1 2060.9 1912.9 2055.7 1907.1 2049.9 C +1906.7 2047.1 1906.9 2043.9 1906.8 2041 C +1906.8 2043.9 1906.8 2046.8 1906.8 2049.6 C +1913.2 2055.5 1918.7 2061.9 1925.1 2067.6 C +1927.1 2067.9 1928.6 2064.4 1930.1 2066.2 C +1929.7 2070.3 1929.9 2074.7 1929.9 2078.9 C +1929.6 2074.4 1930.5 2070.1 1929.9 2065.7 C +[0.07 0.06 0 0.58] vc +f +S +n +1930.1 2061.6 m +1928.1 2062.1 1927 2065.1 1924.8 2064.3 C +1920.7 2058.9 1914.4 2054.3 1910.2 2049.2 C +1910.2 2048.1 1910.2 2047.1 1910.2 2046 C +1909.8 2046.8 1910 2048.3 1910 2049.4 C +1915.1 2054.9 1920.3 2059 1925.3 2064.8 C +1927.1 2064.2 1928.4 2062.3 1930.1 2061.6 C +[0.18 0.18 0 0.78] vc +f +S +n +1932 2049.9 m +1932.3 2050.3 1932 2050.4 1932.8 2050.4 C +1932 2050.4 1932.2 2049.2 1931.3 2049.6 C +1931.4 2050.5 1930.3 2050.4 1930.4 2051.3 C +1931.1 2051.1 1930.7 2049.4 1932 2049.9 C +f +S +n +1938.3 2046 m +1936.3 2046.8 1935.2 2047.2 1934.2 2048.9 C +1935.3 2047.7 1936.8 2046.2 1938.3 2046 C +[0.4 0.4 0 0] vc +f +S +n +vmrs +1938.3 2047 m +1937.9 2046.9 1936.6 2047.1 1936.1 2048 C +1936.5 2047.5 1937.3 2046.7 1938.3 2047 C +[0.18 0.18 0 0.78] vc +f +0.4 w +2 J +2 M +S +n +1910.2 2043.2 m +1910.1 2037.5 1910 2031.8 1910 2026.1 C +1910 2031.8 1910.1 2037.5 1910.2 2043.2 C +f +S +n +1933.5 2032.1 m +1933.7 2035.2 1932.8 2035.8 1933.7 2038.6 C +1933.3 2036.6 1934.6 2018 1933.5 2032.1 C +f +S +n +1907.3 2021.8 m +1906.6 2025.9 1909.4 2032.6 1903.2 2034 C +1902.8 2034.1 1902.4 2033.9 1902 2033.8 C +1897.9 2028.5 1891.6 2023.8 1887.4 2018.7 C +1887.4 2017.7 1887.4 2016.6 1887.4 2015.6 C +1887 2016.3 1887.2 2017.8 1887.2 2018.9 C +1892.3 2024.4 1897.5 2028.5 1902.5 2034.3 C +1904.3 2033.6 1905.7 2032 1907.3 2030.9 C +1907.3 2027.9 1907.3 2024.9 1907.3 2021.8 C +f +S +n +1933.7 2023.2 m +1932 2021.7 1931.1 2024.9 1929.4 2024.9 C +1931.2 2024.7 1932.4 2021.5 1933.7 2023.2 C +f +S +n +1989.2 2024.4 m +1987.4 2023.7 1985.8 2022.2 1985.1 2020.4 C +1984.6 2020.1 1986 2018.9 1985.1 2019.2 C +1985.6 2020.8 1984.1 2019.4 1984.6 2021.1 C +1986.3 2022.3 1988.1 2025.3 1989.2 2024.4 C +f +S +n +1904.4 2031.9 m +1903 2029.7 1905.3 2027.7 1904.2 2025.9 C +1904.5 2025 1903.7 2023 1904 2021.3 C +1904 2022.3 1903.2 2022 1902.5 2022 C +1901.3 2022.3 1902.2 2020.1 1901.6 2019.6 C +1902.5 2019.8 1902.6 2018.3 1903.5 2018.9 C +1903.7 2021.8 1905.6 2016.8 1905.6 2020.6 C +1905.9 2020 1906.3 2020.8 1906.1 2021.1 C +1905.8 2022.7 1906.7 2020.4 1906.4 2019.9 C +1906.4 2018.5 1908.2 2017.8 1906.8 2016.5 C +1906.9 2015.7 1907.7 2017.1 1907.1 2016.3 C +1908.5 2015.8 1910.3 2015.1 1911.6 2016 C +1912.2 2016.2 1911.9 2018 1911.6 2018 C +1914.5 2017.1 1910.4 2013.6 1913.3 2013.4 C +1912.4 2011.3 1910.5 2011.8 1909.5 2010 C +1910 2010.5 1909 2010.8 1908.8 2011.2 C +1907.5 2009.9 1906.1 2011.7 1904.9 2011.5 C +1904.7 2010.9 1904.3 2010.5 1904.4 2009.8 C +1905 2010.2 1904.6 2008.6 1905.4 2008.1 C +1906.6 2007.5 1907.7 2008.4 1908.5 2007.4 C +1908.9 2008.5 1909.7 2008.1 1909 2007.2 C +1908.1 2006.5 1905.9 2007.3 1905.4 2007.4 C +1903.9 2007.3 1905.2 2008.5 1904.2 2008.4 C +1904.6 2009.9 1902.8 2010.3 1902.3 2010.5 C +1901.5 2009.9 1900.4 2010 1899.4 2010 C +1898.6 2011.2 1898.2 2013.4 1896.5 2013.4 C +1896 2012.9 1894.4 2012.9 1893.6 2012.9 C +1893.1 2013.9 1892.9 2015.5 1891.5 2016 C +1890.3 2016.1 1889.2 2014 1888.6 2015.8 C +1890 2016 1891 2016.9 1892.9 2016.5 C +1894.1 2017.2 1892.8 2018.3 1893.2 2018.9 C +1892.6 2018.9 1891.1 2019.8 1890.5 2020.6 C +1891.1 2023.6 1893.2 2019.8 1893.9 2022.5 C +1894.1 2023.3 1892.7 2023.6 1893.9 2024 C +1894.2 2024.3 1897.4 2023.8 1896.5 2026.1 C +1896 2025.6 1897.4 2028.1 1897.5 2027.1 C +1898.4 2027.4 1899.3 2027 1899.6 2028.5 C +1899.5 2028.6 1899.4 2028.8 1899.2 2028.8 C +1899.3 2029.2 1899.6 2029.8 1900.1 2030.2 C +1900.4 2029.6 1901 2030 1901.8 2030.2 C +1903.1 2032.1 1900.4 2031.5 1902.8 2033.1 C +1903.3 2032.7 1904.5 2032 1904.4 2031.9 C +[0.21 0.21 0 0] vc +f +S +n +1909.2 2019.4 m +1908.8 2020.3 1910.2 2019.8 1909.2 2019.2 C +1908.3 2019.3 1907.6 2020.2 1907.6 2021.3 C +1908.5 2021 1907.6 2019 1909.2 2019.4 C +[0.18 0.18 0 0.78] vc +f +S +n +1915.5 2015.6 m +1913.5 2016.3 1912.4 2016.8 1911.4 2018.4 C +1912.5 2017.2 1914 2015.7 1915.5 2015.6 C +[0.4 0.4 0 0] vc +f +S +n +1915.5 2016.5 m +1915.1 2016.4 1913.8 2016.6 1913.3 2017.5 C +1913.7 2017 1914.5 2016.2 1915.5 2016.5 C +[0.18 0.18 0 0.78] vc +f +S +n +vmrs +1887.4 2012.7 m +1887.3 2007 1887.2 2001.3 1887.2 1995.6 C +1887.2 2001.3 1887.3 2007 1887.4 2012.7 C +[0.18 0.18 0 0.78] vc +f +0.4 w +2 J +2 M +S +n +1935.9 2007.4 m +1936.2 2007.8 1935.8 2007.9 1936.6 2007.9 C +1935.9 2007.9 1936.1 2006.7 1935.2 2007.2 C +1935.2 2008.1 1934.1 2007.9 1934.2 2008.8 C +1935 2008.7 1934.6 2006.9 1935.9 2007.4 C +f +S +n +1942.1 2003.6 m +1940.1 2004.3 1939.1 2004.8 1938 2006.4 C +1939.1 2005.2 1940.6 2003.7 1942.1 2003.6 C +[0.4 0.4 0 0] vc +f +S +n +1942.1 2004.5 m +1941.8 2004.4 1940.4 2004.6 1940 2005.5 C +1940.4 2005 1941.2 2004.2 1942.1 2004.5 C +[0.18 0.18 0 0.78] vc +f +S +n +1914 2000.7 m +1914 1995 1913.9 1989.3 1913.8 1983.6 C +1913.9 1989.3 1914 1995 1914 2000.7 C +f +S +n +1941.6 1998.3 m +1943.4 2001.9 1942.4 1996 1940.9 1998.3 C +1941.2 1998.3 1941.4 1998.3 1941.6 1998.3 C +f +S +n +1954.8 1989.9 m +1953.9 1989.6 1954.7 1991.6 1953.9 1991.1 C +1954.5 1993.1 1953.6 1998 1954.6 1993.2 C +1954 1992.2 1954.7 1990.7 1954.8 1989.9 C +f +S +n +1947.6 1992.5 m +1946.2 1993.5 1944.9 1993 1944.8 1994.7 C +1945.5 1994 1947 1992.2 1947.6 1992.5 C +f +S +n +1910.7 1982.2 m +1910.3 1981.8 1909.7 1982 1909.2 1982 C +1909.7 1982 1910.3 1981.9 1910.7 1982.2 C +1911 1987.1 1910 1992.6 1910.7 1997.3 C +1910.7 1992.3 1910.7 1987.2 1910.7 1982.2 C +[0.65 0.65 0 0.42] vc +f +S +n +1910.9 1992.8 m +1910.9 1991.3 1910.9 1989.7 1910.9 1988.2 C +1910.9 1989.7 1910.9 1991.3 1910.9 1992.8 C +[0.18 0.18 0 0.78] vc +f +S +n +vmrs +1953.6 1983.6 m +1954.1 1985.3 1953.2 1988.6 1954.8 1989.4 C +1954.1 1987.9 1954.4 1985.4 1953.6 1983.6 C +[0.18 0.18 0 0.78] vc +f +0.4 w +2 J +2 M +S +n +1910.7 1982 m +1911.6 1982.9 1911 1984.4 1911.2 1985.6 C +1911 1984.4 1911.6 1982.9 1910.7 1982 C +f +S +n +1947.2 1979.6 m +1947.5 1980.6 1948.3 1980.6 1947.4 1979.6 C +1946.2 1979.4 1945.7 1978.8 1947.2 1979.6 C +f +S +n +1930.4 2061.4 m +1930.4 2058 1930.4 2053.5 1930.4 2051.1 C +1930.7 2054.6 1929.8 2057.4 1930.1 2061.2 C +1929.5 2061.9 1929.7 2061.2 1930.4 2061.4 C +[0.65 0.65 0 0.42] vc +f +S +n +1939.5 2044.8 m +1940 2041.5 1935.2 2044.3 1936.4 2040.8 C +1934.9 2040.9 1934.1 2039.7 1933.5 2038.6 C +1933.3 2035.4 1933.2 2040 1934 2040.3 C +1936.2 2040.6 1936.3 2043.6 1938.5 2043.4 C +1939.7 2044.2 1939.4 2045.6 1938.3 2046.5 C +1939.1 2046.6 1939.6 2045.6 1939.5 2044.8 C +f +S +n +1910.4 2045.3 m +1910.4 2039.5 1910.4 2033.6 1910.4 2027.8 C +1910.4 2033.6 1910.4 2039.5 1910.4 2045.3 C +f +S +n +1906.8 2030.9 m +1907.6 2026.8 1905 2020.8 1909 2018.7 C +1906.5 2018.9 1906.8 2022.4 1906.8 2024.7 C +1906.4 2028.2 1907.9 2032 1903 2033.8 C +1902.2 2034 1903.8 2033.4 1904.2 2033.1 C +1905.1 2032.4 1905.9 2031.5 1906.8 2030.9 C +[0.07 0.06 0 0.58] vc +f +S +n +1907.1 2030.7 m +1907.1 2028.8 1907.1 2027 1907.1 2025.2 C +1907.1 2027 1907.1 2028.8 1907.1 2030.7 C +[0.65 0.65 0 0.42] vc +f +S +n +1932 2023.2 m +1932.2 2023.6 1931.7 2023.7 1931.6 2024 C +1932 2023.7 1932.3 2022.8 1933 2023 C +1933.9 2024.3 1933.3 2026.2 1933.5 2027.8 C +1933.5 2026.4 1934.9 2022.2 1932 2023.2 C +f +S +n +2026.1 2021.6 m +2026.1 2020.8 2026.1 2019.9 2026.1 2019.2 C +2026.1 2019.9 2026.1 2020.8 2026.1 2021.6 C +f +S +n +vmrs +1934.2 2018.9 m +1934.2 2015.5 1934.2 2011 1934.2 2008.6 C +1934.5 2012.1 1933.7 2014.9 1934 2018.7 C +1933.4 2019.5 1933.5 2018.7 1934.2 2018.9 C +[0.65 0.65 0 0.42] vc +f +0.4 w +2 J +2 M +S +n +1887.6 2014.8 m +1887.6 2009 1887.6 2003.1 1887.6 1997.3 C +1887.6 2003.1 1887.6 2009 1887.6 2014.8 C +f +S +n +1914.3 2002.8 m +1914.3 1997 1914.3 1991.1 1914.3 1985.3 C +1914.3 1991.1 1914.3 1997 1914.3 2002.8 C +f +S +n +1995.4 1992.3 m +1995.4 1991.5 1995.4 1990.7 1995.4 1989.9 C +1995.4 1990.7 1995.4 1991.5 1995.4 1992.3 C +f +S +n +1896 1988.4 m +1896.9 1988 1897.8 1987.7 1898.7 1987.2 C +1897.8 1987.7 1896.9 1988 1896 1988.4 C +f +S +n +1899.4 1986.8 m +1900.4 1986.3 1901.3 1985.8 1902.3 1985.3 C +1901.3 1985.8 1900.4 1986.3 1899.4 1986.8 C +f +S +n +1902.8 1985.1 m +1905.2 1984 1905.2 1984 1902.8 1985.1 C +f +S +n +1949.1 1983.4 m +1950.2 1984.4 1947.8 1984.6 1949.3 1985.1 C +1949.5 1984.4 1949.6 1984.1 1949.1 1983.4 C +[0.07 0.06 0 0.58] vc +f +S +n +1906.1 1983.4 m +1908.6 1982 1908.6 1982 1906.1 1983.4 C +[0.65 0.65 0 0.42] vc +f +S +n +1922.7 1976.4 m +1923.6 1976 1924.4 1975.7 1925.3 1975.2 C +1924.4 1975.7 1923.6 1976 1922.7 1976.4 C +f +S +n +vmrs +1926 1974.8 m +1927 1974.3 1928 1973.8 1928.9 1973.3 C +1928 1973.8 1927 1974.3 1926 1974.8 C +[0.65 0.65 0 0.42] vc +f +0.4 w +2 J +2 M +S +n +1929.4 1973.1 m +1931.9 1972 1931.9 1972 1929.4 1973.1 C +f +S +n +1932.8 1971.4 m +1935.3 1970 1935.3 1970 1932.8 1971.4 C +f +S +n +1949.6 2097.2 m +1951.1 2096.4 1952.6 2095.5 1954.1 2094.8 C +1952.6 2095.5 1951.1 2096.4 1949.6 2097.2 C +[0.07 0.06 0 0.58] vc +f +S +n +1955.1 2094.3 m +1956.7 2093.5 1958.3 2092.7 1959.9 2091.9 C +1958.3 2092.7 1956.7 2093.5 1955.1 2094.3 C +f +S +n +1960.4 2091.6 m +1961.3 2091.2 1962.1 2090.9 1963 2090.4 C +1962.1 2090.9 1961.3 2091.2 1960.4 2091.6 C +f +S +n +1963.5 2090.2 m +1964.4 2089.7 1965.2 2089.2 1966.1 2088.8 C +1965.2 2089.2 1964.4 2089.7 1963.5 2090.2 C +f +S +n +1966.6 2088.5 m +1969.5 2087.1 1972.4 2085.8 1975.2 2084.4 C +1972.4 2085.8 1969.5 2087.1 1966.6 2088.5 C +f +S +n +1965.2 2086.1 m +1965.9 2085.7 1966.8 2085.3 1967.6 2084.9 C +1966.8 2085.3 1965.9 2085.7 1965.2 2086.1 C +f +S +n +1968.3 2084.7 m +1969.2 2084.3 1970 2083.9 1970.9 2083.5 C +1970 2083.9 1969.2 2084.3 1968.3 2084.7 C +f +S +n +vmrs +1984.1 2084 m +1985.6 2083.2 1987.2 2082.3 1988.7 2081.6 C +1987.2 2082.3 1985.6 2083.2 1984.1 2084 C +[0.07 0.06 0 0.58] vc +f +0.4 w +2 J +2 M +S +n +1976 2078.7 m +1978.1 2080.1 1980 2082 1982 2083.7 C +1980 2081.9 1977.9 2080.3 1976 2078.2 C +1975.5 2079.9 1975.8 2081.9 1975.7 2083.7 C +1975.8 2082 1975.5 2080.2 1976 2078.7 C +f +S +n +1989.6 2081.1 m +1991.3 2080.3 1992.8 2079.5 1994.4 2078.7 C +1992.8 2079.5 1991.3 2080.3 1989.6 2081.1 C +f +S +n +1933.2 2074.6 m +1932.4 2076.2 1932.8 2077.5 1933 2078.7 C +1933 2077.6 1932.9 2074.8 1933.2 2074.6 C +f +S +n +1994.9 2078.4 m +1995.8 2078 1996.7 2077.7 1997.6 2077.2 C +1996.7 2077.7 1995.8 2078 1994.9 2078.4 C +f +S +n +1998 2077 m +1998.9 2076.5 1999.8 2076 2000.7 2075.6 C +1999.8 2076 1998.9 2076.5 1998 2077 C +f +S +n +2001.2 2075.3 m +2004 2073.9 2006.9 2072.6 2009.8 2071.2 C +2006.9 2072.6 2004 2073.9 2001.2 2075.3 C +f +S +n +1980.5 2060.7 m +1979.9 2060.7 1976.7 2062.8 1975.7 2064.5 C +1975.7 2067.5 1975.7 2070.5 1975.7 2073.4 C +1976.3 2068.7 1973.9 2061.6 1980.5 2060.7 C +f +S +n +1999.7 2072.9 m +2000.5 2072.5 2001.3 2072.1 2002.1 2071.7 C +2001.3 2072.1 2000.5 2072.5 1999.7 2072.9 C +f +S +n +2002.8 2071.5 m +2003.7 2071.1 2004.6 2070.7 2005.5 2070.3 C +2004.6 2070.7 2003.7 2071.1 2002.8 2071.5 C +f +S +n +vmrs +2015.1 2047.5 m +2014.4 2047.5 2011.2 2049.6 2010.3 2051.3 C +2010.3 2057.7 2010.3 2064.1 2010.3 2070.5 C +2010.3 2063.9 2010.1 2057.1 2010.5 2050.6 C +2012 2049.3 2013.5 2048.3 2015.1 2047.5 C +[0.07 0.06 0 0.58] vc +f +0.4 w +2 J +2 M +S +n +1910.4 2049.2 m +1914.8 2054.3 1920.7 2058.9 1925.1 2064 C +1920.4 2058.6 1915.1 2054.6 1910.4 2049.2 C +f +S +n +1988.2 2057.3 m +1989.1 2056.8 1989.9 2056.2 1990.8 2055.6 C +1989.9 2056.2 1989.1 2056.8 1988.2 2057.3 C +f +S +n +1991.6 2051.3 m +1991.6 2046.3 1991.6 2041.2 1991.6 2036.2 C +1991.6 2041.2 1991.6 2046.3 1991.6 2051.3 C +f +S +n +1935.6 2047.5 m +1932.9 2051.7 1939.7 2043.8 1935.6 2047.5 C +f +S +n +1938.8 2043.9 m +1938.1 2043.3 1938.2 2043.7 1937.3 2043.4 C +1938.7 2043 1938.2 2044.9 1939 2045.3 C +1938.2 2045.3 1938.7 2046.6 1937.8 2046.5 C +1939.1 2046.2 1939.1 2044.5 1938.8 2043.9 C +f +S +n +1972.4 2045.6 m +1973.4 2045 1974.5 2044.4 1975.5 2043.9 C +1974.5 2044.4 1973.4 2045 1972.4 2045.6 C +f +S +n +1969 2043.6 m +1969.8 2043.2 1970.6 2042.9 1971.4 2042.4 C +1970.6 2042.9 1969.8 2043.2 1969 2043.6 C +f +S +n +1972.1 2042.2 m +1973 2041.8 1973.9 2041.4 1974.8 2041 C +1973.9 2041.4 1973 2041.8 1972.1 2042.2 C +f +S +n +1906.6 2035 m +1905 2034.7 1904.8 2036.6 1903.5 2036.9 C +1904.9 2037 1905.8 2033.4 1907.1 2035.7 C +1907.1 2037.2 1907.1 2038.6 1907.1 2040 C +1906.9 2038.4 1907.5 2036.4 1906.6 2035 C +f +S +n +vmrs +1937.1 2032.1 m +1936.2 2033.7 1936.6 2035 1936.8 2036.2 C +1936.8 2035.1 1936.8 2032.4 1937.1 2032.1 C +[0.07 0.06 0 0.58] vc +f +0.4 w +2 J +2 M +S +n +1887.6 2018.7 m +1892 2023.8 1897.9 2028.4 1902.3 2033.6 C +1897.6 2028.1 1892.3 2024.1 1887.6 2018.7 C +f +S +n +1999.7 2031.4 m +1998.7 2030.3 1997.6 2029.2 1996.6 2028 C +1997.6 2029.2 1998.7 2030.3 1999.7 2031.4 C +f +S +n +1912.8 2017 m +1910.6 2021.1 1913.6 2015.3 1914.5 2016 C +1914 2016.3 1913.4 2016.7 1912.8 2017 C +f +S +n +1939.5 2005 m +1936.7 2009.2 1943.6 2001.3 1939.5 2005 C +f +S +n +1942.6 2001.4 m +1941.9 2000.8 1942 2001.2 1941.2 2000.9 C +1942.5 2000.6 1942.1 2002.4 1942.8 2002.8 C +1942 2002.8 1942.5 2004.1 1941.6 2004 C +1943 2003.7 1942.9 2002.1 1942.6 2001.4 C +f +S +n +2006.2 2000.7 m +2005.4 2001.5 2004 2002.8 2004 2002.8 C +2004.5 2002.4 2005.5 2001.4 2006.2 2000.7 C +f +S +n +1998.5 2001.6 m +1997.7 2002 1996.8 2002.4 1995.9 2002.6 C +1995.5 1999.3 1995.7 1995.7 1995.6 1992.3 C +1995.6 1995.7 1995.6 1999.2 1995.6 2002.6 C +1996.6 2002.4 1997.7 2002.2 1998.5 2001.6 C +[0.4 0.4 0 0] vc +f +S +n +1996.1 2002.8 m +1995.9 2002.8 1995.8 2002.8 1995.6 2002.8 C +1995.2 1999.5 1995.5 1995.9 1995.4 1992.5 C +1995.4 1995.9 1995.4 1999.4 1995.4 2002.8 C +1996.4 2003.1 1998.2 2001.6 1996.1 2002.8 C +[0.07 0.06 0 0.58] vc +f +S +n +1969 2002.1 m +1968 2001 1966.9 1999.9 1965.9 1998.8 C +1966.9 1999.9 1968 2001 1969 2002.1 C +f +S +n +vmrs +2000 2001.2 m +2002.1 2000 2004.1 1998.9 2006.2 1997.8 C +2004.1 1998.9 2002.1 2000 2000 2001.2 C +[0.07 0.06 0 0.58] vc +f +0.4 w +2 J +2 M +S +n +1895.8 1984.8 m +1898.3 1983.6 1900.8 1982.3 1903.2 1981 C +1900.8 1982.3 1898.3 1983.6 1895.8 1984.8 C +f +S +n +1905.2 1980.3 m +1906.4 1979.9 1907.6 1979.5 1908.8 1979.1 C +1907.6 1979.5 1906.4 1979.9 1905.2 1980.3 C +f +S +n +1964.7 1977.4 m +1963.8 1977.5 1962.5 1980.2 1960.8 1980 C +1962.5 1980.2 1963.3 1978 1964.7 1977.4 C +f +S +n +1952 1979.6 m +1955.2 1979.2 1955.2 1979.2 1952 1979.6 C +f +S +n +1937.8 1966.4 m +1941.2 1969.5 1946.1 1976.4 1951.5 1979.3 C +1946.1 1976.7 1942.8 1970.4 1937.8 1966.4 C +f +S +n +1911.9 1978.6 m +1914.3 1977.4 1916.7 1976.2 1919.1 1975 C +1916.7 1976.2 1914.3 1977.4 1911.9 1978.6 C +f +S +n +1975.5 1971.4 m +1974.6 1972.2 1973.3 1973.6 1973.3 1973.6 C +1973.7 1973.1 1974.8 1972.1 1975.5 1971.4 C +f +S +n +1922.4 1972.8 m +1924.9 1971.6 1927.4 1970.3 1929.9 1969 C +1927.4 1970.3 1924.9 1971.6 1922.4 1972.8 C +f +S +n +1969.2 1971.9 m +1971.1 1970.9 1972.9 1969.8 1974.8 1968.8 C +1972.9 1969.8 1971.1 1970.9 1969.2 1971.9 C +f +S +n +vmrs +1931.8 1968.3 m +1933 1967.9 1934.2 1967.5 1935.4 1967.1 C +1934.2 1967.5 1933 1967.9 1931.8 1968.3 C +[0.07 0.06 0 0.58] vc +f +0.4 w +2 J +2 M +S +n +1940.7 2072.4 m +1941.5 2072.4 1942.3 2072.3 1943.1 2072.2 C +1942.3 2072.3 1941.5 2072.4 1940.7 2072.4 C +[0 0 0 0.18] vc +f +S +n +1948.6 2069.3 m +1947 2069.5 1945.7 2068.9 1944.8 2069.8 C +1945.9 2068.5 1948.4 2070.2 1948.6 2069.3 C +f +S +n +1954.6 2066.4 m +1954.7 2067.9 1955.6 2067.3 1955.6 2068.8 C +1955.4 2067.8 1956 2066.6 1954.6 2066.4 C +f +S +n +1929.2 2061.2 m +1927.8 2062.1 1926.3 2064.1 1924.8 2063.3 C +1926.3 2064.6 1928 2062 1929.2 2061.2 C +f +S +n +1924.4 2067.4 m +1918.5 2061.6 1912.7 2055.9 1906.8 2050.1 C +1912.7 2055.9 1918.5 2061.6 1924.4 2067.4 C +[0.4 0.4 0 0] vc +f +S +n +1924.6 2062.8 m +1923.9 2062.1 1923.2 2061.2 1922.4 2060.4 C +1923.2 2061.2 1923.9 2062.1 1924.6 2062.8 C +[0 0 0 0.18] vc +f +S +n +1919.3 2057.3 m +1917.5 2055.6 1915.7 2053.8 1913.8 2052 C +1915.7 2053.8 1917.5 2055.6 1919.3 2057.3 C +f +S +n +1929.2 2055.2 m +1929.2 2054.2 1929.2 2053.2 1929.2 2052.3 C +1929.2 2053.2 1929.2 2054.2 1929.2 2055.2 C +f +S +n +1926.3 2049.6 m +1925.4 2049 1925.4 2050.5 1924.4 2050.4 C +1925.3 2051.3 1924.5 2051.9 1925.6 2052.5 C +1926.9 2052.6 1926 2050.6 1926.3 2049.6 C +f +S +n +vmrs +1911.2 2046.8 m +1910.1 2048.9 1911.9 2050.1 1913.1 2051.3 C +1912.1 2049.9 1910.6 2048.8 1911.2 2046.8 C +[0 0 0 0.18] vc +f +0.4 w +2 J +2 M +S +n +1934 2048.7 m +1932.6 2048.7 1930.1 2047.7 1929.6 2049.4 C +1930.9 2048.6 1933.3 2049 1934 2048.7 C +f +S +n +1980 2048.4 m +1979.5 2046.8 1976.3 2047.9 1977.2 2045.6 C +1976.8 2045.1 1976.1 2044.7 1975.2 2044.8 C +1973.7 2046 1976.3 2046.4 1976.7 2047.5 C +1977.8 2047.2 1978.2 2050 1979.6 2049.2 C +1980 2049 1979.6 2048.6 1980 2048.4 C +f +S +n +1938.3 2045.6 m +1938.2 2044.4 1936.8 2043.8 1935.9 2043.4 C +1936.4 2044.4 1939.1 2044.3 1937.6 2045.8 C +1937 2046.1 1935.9 2046.1 1935.9 2046.8 C +1936.7 2046.3 1937.8 2046.2 1938.3 2045.6 C +f +S +n +1932.5 2040 m +1932.8 2038.1 1932 2038.9 1932.3 2040.3 C +1933.1 2040.3 1932.7 2041.7 1933.7 2041.5 C +1933.1 2041 1932.9 2040.5 1932.5 2040 C +f +S +n +2014.6 2035.2 m +2014.1 2033.6 2010.9 2034.7 2011.7 2032.4 C +2011.3 2031.9 2009.4 2030.7 2009.3 2032.1 C +2009.5 2033.7 2012.9 2033.8 2012.4 2035.7 C +2013 2036.4 2014.2 2036.5 2014.6 2035.2 C +f +S +n +1906.4 2030.7 m +1905 2031.6 1903.5 2033.6 1902 2032.8 C +1903.4 2034 1905.6 2031.4 1906.4 2030.7 C +f +S +n +1901.8 2037.2 m +1899.5 2034.8 1897.2 2032.5 1894.8 2030.2 C +1897.2 2032.5 1899.5 2034.8 1901.8 2037.2 C +[0.4 0.4 0 0] vc +f +S +n +1901.8 2032.4 m +1901.1 2031.6 1900.4 2030.7 1899.6 2030 C +1900.4 2030.7 1901.1 2031.6 1901.8 2032.4 C +[0 0 0 0.18] vc +f +S +n +1944.5 2030 m +1945.3 2029.9 1946.1 2029.8 1946.9 2029.7 C +1946.1 2029.8 1945.3 2029.9 1944.5 2030 C +f +S +n +vmrs +1997.8 2027.8 m +1997.7 2027.9 1997.6 2028.1 1997.3 2028 C +1997.4 2029.1 1998.5 2029.5 1999.2 2030 C +2000.1 2029.5 1998.9 2028 1997.8 2027.8 C +[0 0 0 0.18] vc +f +0.4 w +2 J +2 M +S +n +1906.4 2029.2 m +1906.4 2026.6 1906.4 2024 1906.4 2021.3 C +1906.4 2024 1906.4 2026.6 1906.4 2029.2 C +f +S +n +2006.2 2025.9 m +2006 2025.9 2005.8 2025.8 2005.7 2025.6 C +2005.7 2025.5 2005.7 2025.3 2005.7 2025.2 C +2004.6 2025.8 2002.7 2024.7 2001.9 2026.1 C +2001.9 2027.9 2007.8 2029.2 2006.2 2025.9 C +[0 0 0 0] vc +f +S +n +1952.4 2026.8 m +1950.9 2027 1949.6 2026.4 1948.6 2027.3 C +1949.7 2026.1 1952.2 2027.7 1952.4 2026.8 C +[0 0 0 0.18] vc +f +S +n +1896.5 2026.8 m +1894.7 2025.1 1892.9 2023.3 1891 2021.6 C +1892.9 2023.3 1894.7 2025.1 1896.5 2026.8 C +f +S +n +1958.4 2024 m +1958.5 2025.5 1959.4 2024.8 1959.4 2026.4 C +1959.3 2025.3 1959.8 2024.1 1958.4 2024 C +f +S +n +1903.5 2019.2 m +1902.6 2018.6 1902.6 2020 1901.6 2019.9 C +1902.5 2020.8 1901.7 2021.4 1902.8 2022 C +1904.1 2022.2 1903.2 2020.1 1903.5 2019.2 C +f +S +n +1933 2018.7 m +1931.7 2019.6 1930.1 2021.6 1928.7 2020.8 C +1930.1 2022.1 1931.8 2019.5 1933 2018.7 C +f +S +n +1888.4 2016.3 m +1887.3 2018.4 1889.1 2019.6 1890.3 2020.8 C +1889.3 2019.5 1887.8 2018.3 1888.4 2016.3 C +f +S +n +1928.4 2020.4 m +1927.7 2019.6 1927 2018.7 1926.3 2018 C +1927 2018.7 1927.7 2019.6 1928.4 2020.4 C +f +S +n +vmrs +1911.2 2018.2 m +1909.8 2018.3 1907.3 2017.2 1906.8 2018.9 C +1908.1 2018.1 1910.5 2018.6 1911.2 2018.2 C +[0 0 0 0.18] vc +f +0.4 w +2 J +2 M +S +n +1915.5 2015.1 m +1915.4 2013.9 1914 2013.3 1913.1 2012.9 C +1913.6 2013.9 1916.3 2013.8 1914.8 2015.3 C +1914.2 2015.6 1913.1 2015.6 1913.1 2016.3 C +1913.9 2015.9 1915 2015.7 1915.5 2015.1 C +f +S +n +1923.2 2014.8 m +1921.3 2013.1 1919.5 2011.3 1917.6 2009.6 C +1919.5 2011.3 1921.3 2013.1 1923.2 2014.8 C +f +S +n +1933 2012.7 m +1933 2011.7 1933 2010.8 1933 2009.8 C +1933 2010.8 1933 2011.7 1933 2012.7 C +f +S +n +1909.7 2008.1 m +1908.9 2009.2 1910.1 2009.9 1910.4 2011 C +1911.1 2010.7 1908.9 2009.7 1909.7 2008.1 C +f +S +n +1930.1 2007.2 m +1929.2 2006.6 1929.2 2008 1928.2 2007.9 C +1929.1 2008.8 1928.4 2009.4 1929.4 2010 C +1930.7 2010.2 1929.9 2008.1 1930.1 2007.2 C +f +S +n +1915 2004.3 m +1914 2006.4 1915.7 2007.6 1916.9 2008.8 C +1915.9 2007.5 1914.4 2006.3 1915 2004.3 C +f +S +n +1937.8 2006.2 m +1936.4 2006.3 1934 2005.2 1933.5 2006.9 C +1934.7 2006.1 1937.1 2006.6 1937.8 2006.2 C +f +S +n +1983.9 2006 m +1983.3 2004.3 1980.2 2005.4 1981 2003.1 C +1980.6 2002.7 1978.7 2001.5 1978.6 2002.8 C +1978.8 2004.4 1982.1 2004.5 1981.7 2006.4 C +1982.3 2007.2 1983.5 2007.2 1983.9 2006 C +f +S +n +1942.1 2003.1 m +1942 2001.9 1940.6 2001.3 1939.7 2000.9 C +1940.2 2001.9 1943 2001.8 1941.4 2003.3 C +1940.9 2003.6 1939.7 2003.6 1939.7 2004.3 C +1940.5 2003.9 1941.6 2003.7 1942.1 2003.1 C +f +S +n +vmrs +1967.1 1998.5 m +1967 1998.6 1966.8 1998.8 1966.6 1998.8 C +1966.7 1999.8 1967.8 2000.2 1968.5 2000.7 C +1969.4 2000.2 1968.2 1998.8 1967.1 1998.5 C +[0 0 0 0.18] vc +f +0.4 w +2 J +2 M +S +n +1936.4 1997.6 m +1936.7 1995.6 1935.8 1996.4 1936.1 1997.8 C +1936.9 1997.9 1936.5 1999.2 1937.6 1999 C +1937 1998.5 1936.8 1998 1936.4 1997.6 C +f +S +n +1975.5 1996.6 m +1975.2 1996.7 1975.1 1996.5 1975 1996.4 C +1975 1996.2 1975 1996.1 1975 1995.9 C +1973.9 1996.5 1972 1995.5 1971.2 1996.8 C +1971.2 1998.6 1977 1999.9 1975.5 1996.6 C +[0 0 0 0] vc +f +S +n +1949.3 2097.4 m +1950.3 2096.9 1951.2 2096.4 1952.2 2096 C +1951.2 2096.4 1950.3 2096.9 1949.3 2097.4 C +[0.4 0.4 0 0] vc +f +S +n +1960.8 2091.6 m +1961.7 2091.2 1962.6 2090.9 1963.5 2090.4 C +1962.6 2090.9 1961.7 2091.2 1960.8 2091.6 C +f +S +n +1964.4 2090 m +1965.7 2089.2 1967 2088.5 1968.3 2087.8 C +1967 2088.5 1965.7 2089.2 1964.4 2090 C +f +S +n +1976 2083.7 m +1976.3 2082.3 1975.2 2079.1 1976.9 2079.4 C +1978.8 2080.7 1980.3 2082.9 1982.2 2084.2 C +1980.6 2083.1 1978.2 2080.2 1976 2078.9 C +1975.6 2081.2 1977 2084.9 1973.8 2085.4 C +1972.2 2086.1 1970.7 2087 1969 2087.6 C +1971.4 2086.5 1974.1 2085.6 1976 2083.7 C +f +S +n +1983.9 2084.2 m +1984.8 2083.7 1985.8 2083.2 1986.8 2082.8 C +1985.8 2083.2 1984.8 2083.7 1983.9 2084.2 C +f +S +n +1995.4 2078.4 m +1996.3 2078 1997.1 2077.7 1998 2077.2 C +1997.1 2077.7 1996.3 2078 1995.4 2078.4 C +f +S +n +1999 2076.8 m +2000.3 2076 2001.6 2075.3 2002.8 2074.6 C +2001.6 2075.3 2000.3 2076 1999 2076.8 C +f +S +n +vmrs +1929.6 2065.7 m +1930.1 2065.6 1929.8 2068.6 1929.9 2070 C +1929.8 2068.6 1930.1 2067 1929.6 2065.7 C +[0.4 0.4 0 0] vc +f +0.4 w +2 J +2 M +S +n +1906.6 2049.4 m +1906.6 2046.7 1906.6 2043.9 1906.6 2041.2 C +1906.6 2043.9 1906.6 2046.7 1906.6 2049.4 C +f +S +n +2016 2047.5 m +2014.8 2048 2013.5 2048.3 2012.4 2049.4 C +2013.5 2048.3 2014.8 2048 2016 2047.5 C +f +S +n +2016.5 2047.2 m +2017.3 2046.9 2018.1 2046.6 2018.9 2046.3 C +2018.1 2046.6 2017.3 2046.9 2016.5 2047.2 C +f +S +n +1912.4 2028.5 m +1911.8 2032.4 1912.4 2037.2 1911.9 2041.2 C +1911.5 2037.2 1911.7 2032.9 1911.6 2028.8 C +1911.6 2033.5 1911.6 2038.9 1911.6 2042.9 C +1912.5 2042.2 1911.6 2043.9 1912.6 2043.6 C +1912.9 2039.3 1913.1 2033.3 1912.4 2028.5 C +[0.21 0.21 0 0] vc +f +S +n +1906.8 2040.8 m +1906.8 2039 1906.8 2037.2 1906.8 2035.5 C +1906.8 2037.2 1906.8 2039 1906.8 2040.8 C +[0.4 0.4 0 0] vc +f +S +n +1905.9 2035.2 m +1904.9 2036.4 1903.7 2037.2 1902.3 2037.4 C +1903.7 2037.2 1904.9 2036.4 1905.9 2035.2 C +f +S +n +1906.1 2031.2 m +1907 2031.1 1906.4 2028 1906.6 2030.7 C +1905.5 2032.1 1904 2032.8 1902.5 2033.6 C +1903.9 2033.2 1905 2032.1 1906.1 2031.2 C +f +S +n +1908.3 2018.7 m +1905.2 2018.6 1907.1 2023.2 1906.6 2025.4 C +1906.8 2023 1905.9 2019.5 1908.3 2018.7 C +f +S +n +1889.6 1998 m +1889 2001.9 1889.6 2006.7 1889.1 2010.8 C +1888.7 2006.7 1888.9 2002.4 1888.8 1998.3 C +1888.8 2003 1888.8 2008.4 1888.8 2012.4 C +1889.7 2011.7 1888.8 2013.4 1889.8 2013.2 C +1890.1 2008.8 1890.3 2002.8 1889.6 1998 C +[0.21 0.21 0 0] vc +f +S +n +vmrs +1999 2001.4 m +2001 2000.3 2003 1999.2 2005 1998 C +2003 1999.2 2001 2000.3 1999 2001.4 C +[0.4 0.4 0 0] vc +f +0.4 w +2 J +2 M +S +n +1916.2 1986 m +1915.7 1989.9 1916.3 1994.7 1915.7 1998.8 C +1915.3 1994.7 1915.5 1990.4 1915.5 1986.3 C +1915.5 1991 1915.5 1996.4 1915.5 2000.4 C +1916.3 1999.7 1915.5 2001.4 1916.4 2001.2 C +1916.7 1996.8 1917 1990.8 1916.2 1986 C +[0.21 0.21 0 0] vc +f +S +n +1886.9 1989.6 m +1887.8 1989.2 1888.7 1988.9 1889.6 1988.4 C +1888.7 1988.9 1887.8 1989.2 1886.9 1989.6 C +[0.4 0.4 0 0] vc +f +S +n +1892.4 1986.8 m +1895.1 1985.1 1897.9 1983.6 1900.6 1982 C +1897.9 1983.6 1895.1 1985.1 1892.4 1986.8 C +f +S +n +1907.3 1979.3 m +1908.5 1978.9 1909.7 1978.5 1910.9 1978.1 C +1909.7 1978.5 1908.5 1978.9 1907.3 1979.3 C +f +S +n +1938.5 1966.6 m +1942.6 1970.1 1945.9 1976.4 1951.7 1979.1 C +1946.2 1976.1 1943.1 1970.9 1938.5 1966.6 C +f +S +n +1955.1 1978.6 m +1955.9 1978.2 1956.7 1977.8 1957.5 1977.4 C +1956.7 1977.8 1955.9 1978.2 1955.1 1978.6 C +f +S +n +1913.6 1977.6 m +1914.5 1977.2 1915.3 1976.9 1916.2 1976.4 C +1915.3 1976.9 1914.5 1977.2 1913.6 1977.6 C +f +S +n +1919.1 1974.8 m +1921.8 1973.1 1924.5 1971.6 1927.2 1970 C +1924.5 1971.6 1921.8 1973.1 1919.1 1974.8 C +f +S +n +1963.5 1974.5 m +1964.5 1974 1965.6 1973.4 1966.6 1972.8 C +1965.6 1973.4 1964.5 1974 1963.5 1974.5 C +f +S +n +vmrs +1967.8 1972.4 m +1970 1971.2 1972.1 1970 1974.3 1968.8 C +1972.1 1970 1970 1971.2 1967.8 1972.4 C +[0.4 0.4 0 0] vc +f +0.4 w +2 J +2 M +S +n +1934 1967.3 m +1935.2 1966.9 1936.4 1966.5 1937.6 1966.1 C +1936.4 1966.5 1935.2 1966.9 1934 1967.3 C +f +S +n +1928.9 2061.2 m +1928.9 2059.2 1928.9 2057.3 1928.9 2055.4 C +1928.9 2057.3 1928.9 2059.2 1928.9 2061.2 C +[0.21 0.21 0 0] vc +f +S +n +1917.2 2047 m +1917.8 2046.5 1919.6 2046.8 1920 2047.2 C +1920 2046.5 1920.9 2046.8 1921 2046.3 C +1921.9 2047.3 1921.3 2044.1 1921.5 2044.1 C +1919.7 2044.8 1915.7 2043.5 1916.2 2046 C +1916.2 2048.3 1917 2045.9 1917.2 2047 C +[0 0 0 0] vc +f +S +n +1922 2044.1 m +1923.5 2043.2 1927 2045.4 1927.5 2042.9 C +1927.1 2042.6 1927.3 2040.9 1927.2 2041.5 C +1924.9 2042.3 1920.9 2040.6 1922 2044.1 C +f +S +n +1934.9 2043.9 m +1935.2 2043.4 1934.4 2042.7 1934 2042.2 C +1933.2 2041.8 1932.4 2042.8 1932.8 2043.2 C +1932.9 2044 1934.3 2043.3 1934.9 2043.9 C +f +S +n +1906.1 2030.7 m +1906.1 2028.8 1906.1 2027 1906.1 2025.2 C +1906.1 2027 1906.1 2028.8 1906.1 2030.7 C +[0.21 0.21 0 0] vc +f +S +n +1932.8 2018.7 m +1932.8 2016.8 1932.8 2014.8 1932.8 2012.9 C +1932.8 2014.8 1932.8 2016.8 1932.8 2018.7 C +f +S +n +1894.4 2016.5 m +1895 2016 1896.8 2016.3 1897.2 2016.8 C +1897.2 2016 1898.1 2016.3 1898.2 2015.8 C +1899.1 2016.8 1898.5 2013.6 1898.7 2013.6 C +1896.9 2014.4 1892.9 2013 1893.4 2015.6 C +1893.4 2017.8 1894.2 2015.4 1894.4 2016.5 C +[0 0 0 0] vc +f +S +n +1899.2 2013.6 m +1900.7 2012.7 1904.2 2014.9 1904.7 2012.4 C +1904.3 2012.1 1904.5 2010.5 1904.4 2011 C +1902.1 2011.8 1898.1 2010.1 1899.2 2013.6 C +f +S +n +vmrs +1912.1 2013.4 m +1912.4 2012.9 1911.6 2012.3 1911.2 2011.7 C +1910.4 2011.4 1909.6 2012.3 1910 2012.7 C +1910.1 2013.5 1911.5 2012.9 1912.1 2013.4 C +[0 0 0 0] vc +f +0.4 w +2 J +2 M +S +n +1921 2004.5 m +1921.6 2004 1923.4 2004.3 1923.9 2004.8 C +1923.8 2004 1924.8 2004.3 1924.8 2003.8 C +1925.7 2004.8 1925.1 2001.6 1925.3 2001.6 C +1923.6 2002.4 1919.6 2001 1920 2003.6 C +1920 2005.8 1920.8 2003.4 1921 2004.5 C +f +S +n +1925.8 2001.6 m +1927.3 2000.7 1930.8 2002.9 1931.3 2000.4 C +1930.9 2000.1 1931.1 1998.5 1931.1 1999 C +1928.7 1999.8 1924.8 1998.1 1925.8 2001.6 C +f +S +n +1938.8 2001.4 m +1939 2000.9 1938.2 2000.3 1937.8 1999.7 C +1937.1 1999.4 1936.2 2000.3 1936.6 2000.7 C +1936.7 2001.5 1938.1 2000.9 1938.8 2001.4 C +f +S +n +1908.6691 2008.1348 m +1897.82 2010.0477 L +1894.1735 1989.3671 L +1905.0226 1987.4542 L +1908.6691 2008.1348 L +n +q +_bfh +%%IncludeResource: font Symbol +_efh +{ +f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont +1895.041763 1994.291153 m +0 0 32 0 0 (l) ts +} +true +[0 0 0 1]sts +Q +1979.2185 1991.7809 m +1960.6353 1998.5452 L +1953.4532 1978.8124 L +1972.0363 1972.0481 L +1979.2185 1991.7809 L +n +q +_bfh +%%IncludeResource: font Symbol +_efh +{ +f0 [18.793335 -6.84082 6.84021 18.793335 0 0] makesetfont +1955.163254 1983.510773 m +0 0 32 0 0 (\256) ts +} +true +[0 0 0 1]sts +Q +1952.1544 2066.5423 m +1938.0739 2069.025 L +1934.4274 2048.3444 L +1948.5079 2045.8617 L +1952.1544 2066.5423 L +n +q +_bfh +%%IncludeResource: font Symbol +_efh +{ +f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont +1935.29567 2053.268433 m +0 0 32 0 0 (") ts +} +true +[0 0 0 1]sts +Q +1931.7231 2043.621 m +1919.3084 2048.14 L +1910.6898 2024.4607 L +1923.1046 2019.9417 L +1931.7231 2043.621 L +n +q +_bfh +%%IncludeResource: font Symbol +_efh +{ +f0 [22.552002 -8.208984 8.208252 22.552002 0 0] makesetfont +1912.741867 2030.098648 m +0 0 32 0 0 (=) ts +} +true +[0 0 0 1]sts +Q +1944 2024.5 m +1944 2014 L +0.8504 w +0 J +3.863693 M +[0 0 0 1] vc +false setoverprint +S +n +1944.25 2019.1673 m +1952.5 2015.9173 L +S +n +1931.0787 2124.423 m +1855.5505 2043.4285 L +1871.0419 2013.0337 L +1946.5701 2094.0282 L +1931.0787 2124.423 L +n +q +_bfh +%%IncludeResource: font ZapfHumanist601BT-Bold +_efh +{ +f1 [22.155762 23.759277 -14.753906 28.947754 0 0] makesetfont +1867.35347 2020.27063 m +0 0 32 0 0 (Isabelle) ts +} +true +[0 0 0 1]sts +Q +1933.5503 1996.9547 m +1922.7012 1998.8677 L +1919.0547 1978.1871 L +1929.9038 1976.2741 L +1933.5503 1996.9547 L +n +q +_bfh +%%IncludeResource: font Symbol +_efh +{ +f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont +1919.922913 1983.111069 m +0 0 32 0 0 (b) ts +} +true +[0 0 0 1]sts +Q +2006.3221 2025.7184 m +1993.8573 2027.9162 L +1990.2108 2007.2356 L +2002.6756 2005.0378 L +2006.3221 2025.7184 L +n +q +_bfh +%%IncludeResource: font Symbol +_efh +{ +f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont +1991.07901 2012.159653 m +0 0 32 0 0 (a) ts +} +true +[0 0 0 1]sts +Q +vmrs +2030.0624 2094.056 m +1956.3187 2120.904 L +1956.321 2095.3175 L +2030.0647 2068.4695 L +2030.0624 2094.056 L +n +q +_bfh +%%IncludeResource: font ZapfHumanist601BT-Bold +_efh +{ +f1 [22.898804 -8.336792 -0.002197 24.368408 0 0] makesetfont +1956.320496 2101.409561 m +0 0 32 0 0 (Isar) ts +} +true +[0 0 0 1]sts +Q +vmr +vmr +end +%%Trailer +%%DocumentNeededResources: font Symbol +%%+ font ZapfHumanist601BT-Bold +%%DocumentFonts: Symbol +%%+ ZapfHumanist601BT-Bold +%%DocumentNeededFonts: Symbol +%%+ ZapfHumanist601BT-Bold diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/isabelle_isar.pdf Binary file doc-src/Functions/isabelle_isar.pdf has changed diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/mathpartir.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/mathpartir.sty Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,421 @@ +% Mathpartir --- Math Paragraph for Typesetting Inference Rules +% +% Copyright (C) 2001, 2002, 2003, 2004, 2005 Didier Rémy +% +% Author : Didier Remy +% Version : 1.2.0 +% Bug Reports : to author +% Web Site : http://pauillac.inria.fr/~remy/latex/ +% +% Mathpartir is free software; you can redistribute it and/or modify +% it under the terms of the GNU General Public License as published by +% the Free Software Foundation; either version 2, or (at your option) +% any later version. +% +% Mathpartir is distributed in the hope that it will be useful, +% but WITHOUT ANY WARRANTY; without even the implied warranty of +% MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the +% GNU General Public License for more details +% (http://pauillac.inria.fr/~remy/license/GPL). +% +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +% File mathpartir.sty (LaTeX macros) +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +\NeedsTeXFormat{LaTeX2e} +\ProvidesPackage{mathpartir} + [2005/12/20 version 1.2.0 Math Paragraph for Typesetting Inference Rules] + +%% + +%% Identification +%% Preliminary declarations + +\RequirePackage {keyval} + +%% Options +%% More declarations + +%% PART I: Typesetting maths in paragraphe mode + +\newdimen \mpr@tmpdim + +% To ensure hevea \hva compatibility, \hva should expands to nothing +% in mathpar or in inferrule +\let \mpr@hva \empty + +%% normal paragraph parametters, should rather be taken dynamically +\def \mpr@savepar {% + \edef \MathparNormalpar + {\noexpand \lineskiplimit \the\lineskiplimit + \noexpand \lineskip \the\lineskip}% + } + +\def \mpr@rulelineskip {\lineskiplimit=0.3em\lineskip=0.2em plus 0.1em} +\def \mpr@lesslineskip {\lineskiplimit=0.6em\lineskip=0.5em plus 0.2em} +\def \mpr@lineskip {\lineskiplimit=1.2em\lineskip=1.2em plus 0.2em} +\let \MathparLineskip \mpr@lineskip +\def \mpr@paroptions {\MathparLineskip} +\let \mpr@prebindings \relax + +\newskip \mpr@andskip \mpr@andskip 2em plus 0.5fil minus 0.5em + +\def \mpr@goodbreakand + {\hskip -\mpr@andskip \penalty -1000\hskip \mpr@andskip} +\def \mpr@and {\hskip \mpr@andskip} +\def \mpr@andcr {\penalty 50\mpr@and} +\def \mpr@cr {\penalty -10000\mpr@and} +\def \mpr@eqno #1{\mpr@andcr #1\hskip 0em plus -1fil \penalty 10} + +\def \mpr@bindings {% + \let \and \mpr@andcr + \let \par \mpr@andcr + \let \\\mpr@cr + \let \eqno \mpr@eqno + \let \hva \mpr@hva + } +\let \MathparBindings \mpr@bindings + +% \@ifundefined {ignorespacesafterend} +% {\def \ignorespacesafterend {\aftergroup \ignorespaces} + +\newenvironment{mathpar}[1][] + {$$\mpr@savepar \parskip 0em \hsize \linewidth \centering + \vbox \bgroup \mpr@prebindings \mpr@paroptions #1\ifmmode $\else + \noindent $\displaystyle\fi + \MathparBindings} + {\unskip \ifmmode $\fi\egroup $$\ignorespacesafterend} + +% \def \math@mathpar #1{\setbox0 \hbox {$\displaystyle #1$}\ifnum +% \wd0 < \hsize $$\box0$$\else \bmathpar #1\emathpar \fi} + +%%% HOV BOXES + +\def \mathvbox@ #1{\hbox \bgroup \mpr@normallineskip + \vbox \bgroup \tabskip 0em \let \\ \cr + \halign \bgroup \hfil $##$\hfil\cr #1\crcr \egroup \egroup + \egroup} + +\def \mathhvbox@ #1{\setbox0 \hbox {\let \\\qquad $#1$}\ifnum \wd0 < \hsize + \box0\else \mathvbox {#1}\fi} + + +%% Part II -- operations on lists + +\newtoks \mpr@lista +\newtoks \mpr@listb + +\long \def\mpr@cons #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter +{#2}\edef #2{\the \mpr@lista \the \mpr@listb}} + +\long \def\mpr@snoc #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter +{#2}\edef #2{\the \mpr@listb\the\mpr@lista}} + +\long \def \mpr@concat#1=#2\mpr@to#3{\mpr@lista \expandafter {#2}\mpr@listb +\expandafter {#3}\edef #1{\the \mpr@listb\the\mpr@lista}} + +\def \mpr@head #1\mpr@to #2{\expandafter \mpr@head@ #1\mpr@head@ #1#2} +\long \def \mpr@head@ #1#2\mpr@head@ #3#4{\def #4{#1}\def#3{#2}} + +\def \mpr@flatten #1\mpr@to #2{\expandafter \mpr@flatten@ #1\mpr@flatten@ #1#2} +\long \def \mpr@flatten@ \\#1\\#2\mpr@flatten@ #3#4{\def #4{#1}\def #3{\\#2}} + +\def \mpr@makelist #1\mpr@to #2{\def \mpr@all {#1}% + \mpr@lista {\\}\mpr@listb \expandafter {\mpr@all}\edef \mpr@all {\the + \mpr@lista \the \mpr@listb \the \mpr@lista}\let #2\empty + \def \mpr@stripof ##1##2\mpr@stripend{\def \mpr@stripped{##2}}\loop + \mpr@flatten \mpr@all \mpr@to \mpr@one + \expandafter \mpr@snoc \mpr@one \mpr@to #2\expandafter \mpr@stripof + \mpr@all \mpr@stripend + \ifx \mpr@stripped \empty \let \mpr@isempty 0\else \let \mpr@isempty 1\fi + \ifx 1\mpr@isempty + \repeat +} + +\def \mpr@rev #1\mpr@to #2{\let \mpr@tmp \empty + \def \\##1{\mpr@cons ##1\mpr@to \mpr@tmp}#1\let #2\mpr@tmp} + +%% Part III -- Type inference rules + +\newif \if@premisse +\newbox \mpr@hlist +\newbox \mpr@vlist +\newif \ifmpr@center \mpr@centertrue +\def \mpr@htovlist {% + \setbox \mpr@hlist + \hbox {\strut + \ifmpr@center \hskip -0.5\wd\mpr@hlist\fi + \unhbox \mpr@hlist}% + \setbox \mpr@vlist + \vbox {\if@premisse \box \mpr@hlist \unvbox \mpr@vlist + \else \unvbox \mpr@vlist \box \mpr@hlist + \fi}% +} +% OLD version +% \def \mpr@htovlist {% +% \setbox \mpr@hlist +% \hbox {\strut \hskip -0.5\wd\mpr@hlist \unhbox \mpr@hlist}% +% \setbox \mpr@vlist +% \vbox {\if@premisse \box \mpr@hlist \unvbox \mpr@vlist +% \else \unvbox \mpr@vlist \box \mpr@hlist +% \fi}% +% } + +\def \mpr@item #1{$\displaystyle #1$} +\def \mpr@sep{2em} +\def \mpr@blank { } +\def \mpr@hovbox #1#2{\hbox + \bgroup + \ifx #1T\@premissetrue + \else \ifx #1B\@premissefalse + \else + \PackageError{mathpartir} + {Premisse orientation should either be T or B} + {Fatal error in Package}% + \fi \fi + \def \@test {#2}\ifx \@test \mpr@blank\else + \setbox \mpr@hlist \hbox {}% + \setbox \mpr@vlist \vbox {}% + \if@premisse \let \snoc \mpr@cons \else \let \snoc \mpr@snoc \fi + \let \@hvlist \empty \let \@rev \empty + \mpr@tmpdim 0em + \expandafter \mpr@makelist #2\mpr@to \mpr@flat + \if@premisse \mpr@rev \mpr@flat \mpr@to \@rev \else \let \@rev \mpr@flat \fi + \def \\##1{% + \def \@test {##1}\ifx \@test \empty + \mpr@htovlist + \mpr@tmpdim 0em %%% last bug fix not extensively checked + \else + \setbox0 \hbox{\mpr@item {##1}}\relax + \advance \mpr@tmpdim by \wd0 + %\mpr@tmpdim 1.02\mpr@tmpdim + \ifnum \mpr@tmpdim < \hsize + \ifnum \wd\mpr@hlist > 0 + \if@premisse + \setbox \mpr@hlist + \hbox {\unhbox0 \hskip \mpr@sep \unhbox \mpr@hlist}% + \else + \setbox \mpr@hlist + \hbox {\unhbox \mpr@hlist \hskip \mpr@sep \unhbox0}% + \fi + \else + \setbox \mpr@hlist \hbox {\unhbox0}% + \fi + \else + \ifnum \wd \mpr@hlist > 0 + \mpr@htovlist + \mpr@tmpdim \wd0 + \fi + \setbox \mpr@hlist \hbox {\unhbox0}% + \fi + \advance \mpr@tmpdim by \mpr@sep + \fi + }% + \@rev + \mpr@htovlist + \ifmpr@center \hskip \wd\mpr@vlist\fi \box \mpr@vlist + \fi + \egroup +} + +%%% INFERENCE RULES + +\@ifundefined{@@over}{% + \let\@@over\over % fallback if amsmath is not loaded + \let\@@overwithdelims\overwithdelims + \let\@@atop\atop \let\@@atopwithdelims\atopwithdelims + \let\@@above\above \let\@@abovewithdelims\abovewithdelims + }{} + +%% The default + +\def \mpr@@fraction #1#2{\hbox {\advance \hsize by -0.5em + $\displaystyle {#1\mpr@over #2}$}} +\let \mpr@fraction \mpr@@fraction + +%% A generic solution to arrow + +\def \mpr@make@fraction #1#2#3#4#5{\hbox {% + \def \mpr@tail{#1}% + \def \mpr@body{#2}% + \def \mpr@head{#3}% + \setbox1=\hbox{$#4$}\setbox2=\hbox{$#5$}% + \setbox3=\hbox{$\mkern -3mu\mpr@body\mkern -3mu$}% + \setbox3=\hbox{$\mkern -3mu \mpr@body\mkern -3mu$}% + \dimen0=\dp1\advance\dimen0 by \ht3\relax\dp1\dimen0\relax + \dimen0=\ht2\advance\dimen0 by \dp3\relax\ht2\dimen0\relax + \setbox0=\hbox {$\box1 \@@atop \box2$}% + \dimen0=\wd0\box0 + \box0 \hskip -\dimen0\relax + \hbox to \dimen0 {$% + \mathrel{\mpr@tail}\joinrel + \xleaders\hbox{\copy3}\hfil\joinrel\mathrel{\mpr@head}% + $}}} + +%% Old stuff should be removed in next version +\def \mpr@@reduce #1#2{\hbox + {$\lower 0.01pt \mpr@@fraction {#1}{#2}\mkern -15mu\rightarrow$}} +\def \mpr@@rewrite #1#2#3{\hbox + {$\lower 0.01pt \mpr@@fraction {#2}{#3}\mkern -8mu#1$}} +\def \mpr@infercenter #1{\vcenter {\mpr@hovbox{T}{#1}}} + +\def \mpr@empty {} +\def \mpr@inferrule + {\bgroup + \ifnum \linewidth<\hsize \hsize \linewidth\fi + \mpr@rulelineskip + \let \and \qquad + \let \hva \mpr@hva + \let \@rulename \mpr@empty + \let \@rule@options \mpr@empty + \let \mpr@over \@@over + \mpr@inferrule@} +\newcommand {\mpr@inferrule@}[3][] + {\everymath={\displaystyle}% + \def \@test {#2}\ifx \empty \@test + \setbox0 \hbox {$\vcenter {\mpr@hovbox{B}{#3}}$}% + \else + \def \@test {#3}\ifx \empty \@test + \setbox0 \hbox {$\vcenter {\mpr@hovbox{T}{#2}}$}% + \else + \setbox0 \mpr@fraction {\mpr@hovbox{T}{#2}}{\mpr@hovbox{B}{#3}}% + \fi \fi + \def \@test {#1}\ifx \@test\empty \box0 + \else \vbox +%%% Suggestion de Francois pour les etiquettes longues +%%% {\hbox to \wd0 {\RefTirName {#1}\hfil}\box0}\fi + {\hbox {\RefTirName {#1}}\box0}\fi + \egroup} + +\def \mpr@vdotfil #1{\vbox to #1{\leaders \hbox{$\cdot$} \vfil}} + +% They are two forms +% \inferrule [label]{[premisses}{conclusions} +% or +% \inferrule* [options]{[premisses}{conclusions} +% +% Premisses and conclusions are lists of elements separated by \\ +% Each \\ produces a break, attempting horizontal breaks if possible, +% and vertical breaks if needed. +% +% An empty element obtained by \\\\ produces a vertical break in all cases. +% +% The former rule is aligned on the fraction bar. +% The optional label appears on top of the rule +% The second form to be used in a derivation tree is aligned on the last +% line of its conclusion +% +% The second form can be parameterized, using the key=val interface. The +% folloiwng keys are recognized: +% +% width set the width of the rule to val +% narrower set the width of the rule to val\hsize +% before execute val at the beginning/left +% lab put a label [Val] on top of the rule +% lskip add negative skip on the right +% left put a left label [Val] +% Left put a left label [Val], ignoring its width +% right put a right label [Val] +% Right put a right label [Val], ignoring its width +% leftskip skip negative space on the left-hand side +% rightskip skip negative space on the right-hand side +% vdots lift the rule by val and fill vertical space with dots +% after execute val at the end/right +% +% Note that most options must come in this order to avoid strange +% typesetting (in particular leftskip must preceed left and Left and +% rightskip must follow Right or right; vdots must come last +% or be only followed by rightskip. +% + +%% Keys that make sence in all kinds of rules +\def \mprset #1{\setkeys{mprset}{#1}} +\define@key {mprset}{flushleft}[]{\mpr@centerfalse} +\define@key {mprset}{center}[]{\mpr@centertrue} +\define@key {mprset}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite} +\define@key {mprset}{myfraction}[]{\let \mpr@fraction #1} +\define@key {mprset}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}} + +\newbox \mpr@right +\define@key {mpr}{flushleft}[]{\mpr@centerfalse} +\define@key {mpr}{center}[]{\mpr@centertrue} +\define@key {mpr}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite} +\define@key {mpr}{myfraction}[]{\let \mpr@fraction #1} +\define@key {mpr}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}} +\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax + \advance \hsize by -\wd0\box0} +\define@key {mpr}{width}{\hsize #1} +\define@key {mpr}{sep}{\def\mpr@sep{#1}} +\define@key {mpr}{before}{#1} +\define@key {mpr}{lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}} +\define@key {mpr}{Lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}} +\define@key {mpr}{narrower}{\hsize #1\hsize} +\define@key {mpr}{leftskip}{\hskip -#1} +\define@key {mpr}{reduce}[]{\let \mpr@fraction \mpr@@reduce} +\define@key {mpr}{rightskip} + {\setbox \mpr@right \hbox {\unhbox \mpr@right \hskip -#1}} +\define@key {mpr}{LEFT}{\setbox0 \hbox {$#1$}\relax + \advance \hsize by -\wd0\box0} +\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax + \advance \hsize by -\wd0\box0} +\define@key {mpr}{Left}{\llap{$\TirName {#1}\;$}} +\define@key {mpr}{right} + {\setbox0 \hbox {$\;\TirName {#1}$}\relax \advance \hsize by -\wd0 + \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}} +\define@key {mpr}{RIGHT} + {\setbox0 \hbox {$#1$}\relax \advance \hsize by -\wd0 + \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}} +\define@key {mpr}{Right} + {\setbox \mpr@right \hbox {\unhbox \mpr@right \rlap {$\;\TirName {#1}$}}} +\define@key {mpr}{vdots}{\def \mpr@vdots {\@@atop \mpr@vdotfil{#1}}} +\define@key {mpr}{after}{\edef \mpr@after {\mpr@after #1}} + +\newdimen \rule@dimen +\newcommand \mpr@inferstar@ [3][]{\setbox0 + \hbox {\let \mpr@rulename \mpr@empty \let \mpr@vdots \relax + \setbox \mpr@right \hbox{}% + $\setkeys{mpr}{#1}% + \ifx \mpr@rulename \mpr@empty \mpr@inferrule {#2}{#3}\else + \mpr@inferrule [{\mpr@rulename}]{#2}{#3}\fi + \box \mpr@right \mpr@vdots$} + \setbox1 \hbox {\strut} + \rule@dimen \dp0 \advance \rule@dimen by -\dp1 + \raise \rule@dimen \box0} + +\def \mpr@infer {\@ifnextchar *{\mpr@inferstar}{\mpr@inferrule}} +\newcommand \mpr@err@skipargs[3][]{} +\def \mpr@inferstar*{\ifmmode + \let \@do \mpr@inferstar@ + \else + \let \@do \mpr@err@skipargs + \PackageError {mathpartir} + {\string\inferrule* can only be used in math mode}{}% + \fi \@do} + + +%%% Exports + +% Envirnonment mathpar + +\let \inferrule \mpr@infer + +% make a short name \infer is not already defined +\@ifundefined {infer}{\let \infer \mpr@infer}{} + +\def \TirNameStyle #1{\small \textsc{#1}} +\def \tir@name #1{\hbox {\small \TirNameStyle{#1}}} +\let \TirName \tir@name +\let \DefTirName \TirName +\let \RefTirName \TirName + +%%% Other Exports + +% \let \listcons \mpr@cons +% \let \listsnoc \mpr@snoc +% \let \listhead \mpr@head +% \let \listmake \mpr@makelist + + + + +\endinput diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Functions/style.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Functions/style.sty Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,46 @@ +%% toc +\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} +\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} + +%% references +\newcommand{\secref}[1]{\S\ref{#1}} +\newcommand{\chref}[1]{chapter~\ref{#1}} +\newcommand{\figref}[1]{figure~\ref{#1}} + +%% math +\newcommand{\text}[1]{\mbox{#1}} +\newcommand{\isasymvartheta}{\isamath{\theta}} +\newcommand{\isactrlvec}[1]{\emph{$\overline{#1}$}} + +\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} + +\pagestyle{headings} +\sloppy +\binperiod +\underscoreon + +\renewcommand{\isadigit}[1]{\isamath{#1}} + +\newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup} + +\isafoldtag{FIXME} +\isakeeptag{mlref} +\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small} +\renewcommand{\endisatagmlref}{\endgroup} + +\newcommand{\isasymGUESS}{\isakeyword{guess}} +\newcommand{\isasymOBTAIN}{\isakeyword{obtain}} +\newcommand{\isasymTHEORY}{\isakeyword{theory}} +\newcommand{\isasymUSES}{\isakeyword{uses}} +\newcommand{\isasymEND}{\isakeyword{end}} +\newcommand{\isasymCONSTS}{\isakeyword{consts}} +\newcommand{\isasymDEFS}{\isakeyword{defs}} +\newcommand{\isasymTHEOREM}{\isakeyword{theorem}} +\newcommand{\isasymDEFINITION}{\isakeyword{definition}} + +\isabellestyle{it} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "implementation" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Intro/intro.tex --- a/doc-src/Intro/intro.tex Wed Mar 04 11:05:02 2009 +0100 +++ b/doc-src/Intro/intro.tex Wed Mar 04 11:05:29 2009 +0100 @@ -7,7 +7,7 @@ %prth *(\(.*\)); \1; %{\\out \(.*\)} {\\out val it = "\1" : thm} -\title{\includegraphics[scale=0.5]{isabelle} \\[4ex] Introduction to Isabelle} +\title{\includegraphics[scale=0.5]{isabelle} \\[4ex] Old Introduction to Isabelle} \author{{\em Lawrence C. Paulson}\\ Computer Laboratory \\ University of Cambridge \\ \texttt{lcp@cl.cam.ac.uk}\\[3ex] diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Classes/IsaMakefile --- a/doc-src/IsarAdvanced/Classes/IsaMakefile Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,33 +0,0 @@ - -## targets - -default: Thy -images: -test: Thy - -all: images test - - -## global settings - -SRC = $(ISABELLE_HOME)/src -OUT = $(ISABELLE_OUTPUT) -LOG = $(OUT)/log - -USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -d false -C false -D document - - -## Thy - -THY = $(LOG)/HOL-Thy.gz - -Thy: $(THY) - -$(THY): Thy/ROOT.ML Thy/Setup.thy Thy/Classes.thy - @$(USEDIR) HOL Thy - - -## clean - -clean: - @rm -f $(THY) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Classes/Makefile --- a/doc-src/IsarAdvanced/Classes/Makefile Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,38 +0,0 @@ -# -# $Id$ -# - -## targets - -default: dvi - - -## dependencies - -include ../Makefile.in - -NAME = classes - -FILES = $(NAME).tex classes.tex Thy/document/Classes.tex \ - style.sty ../../iman.sty ../../extra.sty ../../isar.sty \ - ../../isabelle.sty ../../isabellesym.sty ../../pdfsetup.sty \ - ../../manual.bib ../../proof.sty - -dvi: $(NAME).dvi - -$(NAME).dvi: $(FILES) isabelle_isar.eps - $(LATEX) $(NAME) - $(BIBTEX) $(NAME) - $(LATEX) $(NAME) - $(LATEX) $(NAME) - -pdf: $(NAME).pdf - -$(NAME).pdf: $(FILES) isabelle_isar.pdf - $(PDFLATEX) $(NAME) - $(BIBTEX) $(NAME) - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) - $(FIXBOOKMARKS) $(NAME).out - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Classes/Thy/Classes.thy --- a/doc-src/IsarAdvanced/Classes/Thy/Classes.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,635 +0,0 @@ -theory Classes -imports Main Setup -begin - -chapter {* Haskell-style classes with Isabelle/Isar *} - -section {* Introduction *} - -text {* - Type classes were introduces by Wadler and Blott \cite{wadler89how} - into the Haskell language, to allow for a reasonable implementation - of overloading\footnote{throughout this tutorial, we are referring - to classical Haskell 1.0 type classes, not considering - later additions in expressiveness}. - As a canonical example, a polymorphic equality function - @{text "eq \ \ \ \ \ bool"} which is overloaded on different - types for @{text "\"}, which is achieved by splitting introduction - of the @{text eq} function from its overloaded definitions by means - of @{text class} and @{text instance} declarations: - - \begin{quote} - - \noindent@{text "class eq where"}\footnote{syntax here is a kind of isabellized Haskell} \\ - \hspace*{2ex}@{text "eq \ \ \ \ \ bool"} - - \medskip\noindent@{text "instance nat \ eq where"} \\ - \hspace*{2ex}@{text "eq 0 0 = True"} \\ - \hspace*{2ex}@{text "eq 0 _ = False"} \\ - \hspace*{2ex}@{text "eq _ 0 = False"} \\ - \hspace*{2ex}@{text "eq (Suc n) (Suc m) = eq n m"} - - \medskip\noindent@{text "instance (\\eq, \\eq) pair \ eq where"} \\ - \hspace*{2ex}@{text "eq (x1, y1) (x2, y2) = eq x1 x2 \ eq y1 y2"} - - \medskip\noindent@{text "class ord extends eq where"} \\ - \hspace*{2ex}@{text "less_eq \ \ \ \ \ bool"} \\ - \hspace*{2ex}@{text "less \ \ \ \ \ bool"} - - \end{quote} - - \noindent Type variables are annotated with (finitely many) classes; - these annotations are assertions that a particular polymorphic type - provides definitions for overloaded functions. - - Indeed, type classes not only allow for simple overloading - but form a generic calculus, an instance of order-sorted - algebra \cite{Nipkow-Prehofer:1993,nipkow-sorts93,Wenzel:1997:TPHOL}. - - From a software engeneering point of view, type classes - roughly correspond to interfaces in object-oriented languages like Java; - so, it is naturally desirable that type classes do not only - provide functions (class parameters) but also state specifications - implementations must obey. For example, the @{text "class eq"} - above could be given the following specification, demanding that - @{text "class eq"} is an equivalence relation obeying reflexivity, - symmetry and transitivity: - - \begin{quote} - - \noindent@{text "class eq where"} \\ - \hspace*{2ex}@{text "eq \ \ \ \ \ bool"} \\ - @{text "satisfying"} \\ - \hspace*{2ex}@{text "refl: eq x x"} \\ - \hspace*{2ex}@{text "sym: eq x y \ eq x y"} \\ - \hspace*{2ex}@{text "trans: eq x y \ eq y z \ eq x z"} - - \end{quote} - - \noindent From a theoretic point of view, type classes are lightweight - modules; Haskell type classes may be emulated by - SML functors \cite{classes_modules}. - Isabelle/Isar offers a discipline of type classes which brings - all those aspects together: - - \begin{enumerate} - \item specifying abstract parameters together with - corresponding specifications, - \item instantiating those abstract parameters by a particular - type - \item in connection with a ``less ad-hoc'' approach to overloading, - \item with a direct link to the Isabelle module system - (aka locales \cite{kammueller-locales}). - \end{enumerate} - - \noindent Isar type classes also directly support code generation - in a Haskell like fashion. - - This tutorial demonstrates common elements of structured specifications - and abstract reasoning with type classes by the algebraic hierarchy of - semigroups, monoids and groups. Our background theory is that of - Isabelle/HOL \cite{isa-tutorial}, for which some - familiarity is assumed. - - Here we merely present the look-and-feel for end users. - Internally, those are mapped to more primitive Isabelle concepts. - See \cite{Haftmann-Wenzel:2006:classes} for more detail. -*} - -section {* A simple algebra example \label{sec:example} *} - -subsection {* Class definition *} - -text {* - Depending on an arbitrary type @{text "\"}, class @{text - "semigroup"} introduces a binary operator @{text "(\)"} that is - assumed to be associative: -*} - -class %quote semigroup = - fixes mult :: "\ \ \ \ \" (infixl "\" 70) - assumes assoc: "(x \ y) \ z = x \ (y \ z)" - -text {* - \noindent This @{command class} specification consists of two - parts: the \qn{operational} part names the class parameter - (@{element "fixes"}), the \qn{logical} part specifies properties on them - (@{element "assumes"}). The local @{element "fixes"} and - @{element "assumes"} are lifted to the theory toplevel, - yielding the global - parameter @{term [source] "mult \ \\semigroup \ \ \ \"} and the - global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\x y - z \ \\semigroup. (x \ y) \ z = x \ (y \ z)"}. -*} - - -subsection {* Class instantiation \label{sec:class_inst} *} - -text {* - The concrete type @{typ int} is made a @{class semigroup} - instance by providing a suitable definition for the class parameter - @{text "(\)"} and a proof for the specification of @{fact assoc}. - This is accomplished by the @{command instantiation} target: -*} - -instantiation %quote int :: semigroup -begin - -definition %quote - mult_int_def: "i \ j = i + (j\int)" - -instance %quote proof - fix i j k :: int have "(i + j) + k = i + (j + k)" by simp - then show "(i \ j) \ k = i \ (j \ k)" - unfolding mult_int_def . -qed - -end %quote - -text {* - \noindent @{command instantiation} allows to define class parameters - at a particular instance using common specification tools (here, - @{command definition}). The concluding @{command instance} - opens a proof that the given parameters actually conform - to the class specification. Note that the first proof step - is the @{method default} method, - which for such instance proofs maps to the @{method intro_classes} method. - This boils down an instance judgement to the relevant primitive - proof goals and should conveniently always be the first method applied - in an instantiation proof. - - From now on, the type-checker will consider @{typ int} - as a @{class semigroup} automatically, i.e.\ any general results - are immediately available on concrete instances. - - \medskip Another instance of @{class semigroup} are the natural numbers: -*} - -instantiation %quote nat :: semigroup -begin - -primrec %quote mult_nat where - "(0\nat) \ n = n" - | "Suc m \ n = Suc (m \ n)" - -instance %quote proof - fix m n q :: nat - show "m \ n \ q = m \ (n \ q)" - by (induct m) auto -qed - -end %quote - -text {* - \noindent Note the occurence of the name @{text mult_nat} - in the primrec declaration; by default, the local name of - a class operation @{text f} to instantiate on type constructor - @{text \} are mangled as @{text f_\}. In case of uncertainty, - these names may be inspected using the @{command "print_context"} command - or the corresponding ProofGeneral button. -*} - -subsection {* Lifting and parametric types *} - -text {* - Overloaded definitions giving on class instantiation - may include recursion over the syntactic structure of types. - As a canonical example, we model product semigroups - using our simple algebra: -*} - -instantiation %quote * :: (semigroup, semigroup) semigroup -begin - -definition %quote - mult_prod_def: "p\<^isub>1 \ p\<^isub>2 = (fst p\<^isub>1 \ fst p\<^isub>2, snd p\<^isub>1 \ snd p\<^isub>2)" - -instance %quote proof - fix p\<^isub>1 p\<^isub>2 p\<^isub>3 :: "\\semigroup \ \\semigroup" - show "p\<^isub>1 \ p\<^isub>2 \ p\<^isub>3 = p\<^isub>1 \ (p\<^isub>2 \ p\<^isub>3)" - unfolding mult_prod_def by (simp add: assoc) -qed - -end %quote - -text {* - \noindent Associativity from product semigroups is - established using - the definition of @{text "(\)"} on products and the hypothetical - associativity of the type components; these hypotheses - are facts due to the @{class semigroup} constraints imposed - on the type components by the @{command instance} proposition. - Indeed, this pattern often occurs with parametric types - and type classes. -*} - - -subsection {* Subclassing *} - -text {* - We define a subclass @{text monoidl} (a semigroup with a left-hand neutral) - by extending @{class semigroup} - with one additional parameter @{text neutral} together - with its property: -*} - -class %quote monoidl = semigroup + - fixes neutral :: "\" ("\") - assumes neutl: "\ \ x = x" - -text {* - \noindent Again, we prove some instances, by - providing suitable parameter definitions and proofs for the - additional specifications. Observe that instantiations - for types with the same arity may be simultaneous: -*} - -instantiation %quote nat and int :: monoidl -begin - -definition %quote - neutral_nat_def: "\ = (0\nat)" - -definition %quote - neutral_int_def: "\ = (0\int)" - -instance %quote proof - fix n :: nat - show "\ \ n = n" - unfolding neutral_nat_def by simp -next - fix k :: int - show "\ \ k = k" - unfolding neutral_int_def mult_int_def by simp -qed - -end %quote - -instantiation %quote * :: (monoidl, monoidl) monoidl -begin - -definition %quote - neutral_prod_def: "\ = (\, \)" - -instance %quote proof - fix p :: "\\monoidl \ \\monoidl" - show "\ \ p = p" - unfolding neutral_prod_def mult_prod_def by (simp add: neutl) -qed - -end %quote - -text {* - \noindent Fully-fledged monoids are modelled by another subclass - which does not add new parameters but tightens the specification: -*} - -class %quote monoid = monoidl + - assumes neutr: "x \ \ = x" - -instantiation %quote nat and int :: monoid -begin - -instance %quote proof - fix n :: nat - show "n \ \ = n" - unfolding neutral_nat_def by (induct n) simp_all -next - fix k :: int - show "k \ \ = k" - unfolding neutral_int_def mult_int_def by simp -qed - -end %quote - -instantiation %quote * :: (monoid, monoid) monoid -begin - -instance %quote proof - fix p :: "\\monoid \ \\monoid" - show "p \ \ = p" - unfolding neutral_prod_def mult_prod_def by (simp add: neutr) -qed - -end %quote - -text {* - \noindent To finish our small algebra example, we add a @{text group} class - with a corresponding instance: -*} - -class %quote group = monoidl + - fixes inverse :: "\ \ \" ("(_\
)" [1000] 999) - assumes invl: "x\
\ x = \" - -instantiation %quote int :: group -begin - -definition %quote - inverse_int_def: "i\
= - (i\int)" - -instance %quote proof - fix i :: int - have "-i + i = 0" by simp - then show "i\
\ i = \" - unfolding mult_int_def neutral_int_def inverse_int_def . -qed - -end %quote - - -section {* Type classes as locales *} - -subsection {* A look behind the scene *} - -text {* - The example above gives an impression how Isar type classes work - in practice. As stated in the introduction, classes also provide - a link to Isar's locale system. Indeed, the logical core of a class - is nothing else than a locale: -*} - -class %quote idem = - fixes f :: "\ \ \" - assumes idem: "f (f x) = f x" - -text {* - \noindent essentially introduces the locale -*} setup %invisible {* Sign.add_path "foo" *} - -locale %quote idem = - fixes f :: "\ \ \" - assumes idem: "f (f x) = f x" - -text {* \noindent together with corresponding constant(s): *} - -consts %quote f :: "\ \ \" - -text {* - \noindent The connection to the type system is done by means - of a primitive axclass -*} setup %invisible {* Sign.add_path "foo" *} - -axclass %quote idem < type - idem: "f (f x) = f x" setup %invisible {* Sign.parent_path *} - -text {* \noindent together with a corresponding interpretation: *} - -interpretation %quote idem_class: - idem "f \ (\\idem) \ \" -proof qed (rule idem) - -text {* - \noindent This gives you at hand the full power of the Isabelle module system; - conclusions in locale @{text idem} are implicitly propagated - to class @{text idem}. -*} setup %invisible {* Sign.parent_path *} - -subsection {* Abstract reasoning *} - -text {* - Isabelle locales enable reasoning at a general level, while results - are implicitly transferred to all instances. For example, we can - now establish the @{text "left_cancel"} lemma for groups, which - states that the function @{text "(x \)"} is injective: -*} - -lemma %quote (in group) left_cancel: "x \ y = x \ z \ y = z" -proof - assume "x \ y = x \ z" - then have "x\
\ (x \ y) = x\
\ (x \ z)" by simp - then have "(x\
\ x) \ y = (x\
\ x) \ z" using assoc by simp - then show "y = z" using neutl and invl by simp -next - assume "y = z" - then show "x \ y = x \ z" by simp -qed - -text {* - \noindent Here the \qt{@{keyword "in"} @{class group}} target specification - indicates that the result is recorded within that context for later - use. This local theorem is also lifted to the global one @{fact - "group.left_cancel:"} @{prop [source] "\x y z \ \\group. x \ y = x \ - z \ y = z"}. Since type @{text "int"} has been made an instance of - @{text "group"} before, we may refer to that fact as well: @{prop - [source] "\x y z \ int. x \ y = x \ z \ y = z"}. -*} - - -subsection {* Derived definitions *} - -text {* - Isabelle locales support a concept of local definitions - in locales: -*} - -primrec %quote (in monoid) pow_nat :: "nat \ \ \ \" where - "pow_nat 0 x = \" - | "pow_nat (Suc n) x = x \ pow_nat n x" - -text {* - \noindent If the locale @{text group} is also a class, this local - definition is propagated onto a global definition of - @{term [source] "pow_nat \ nat \ \\monoid \ \\monoid"} - with corresponding theorems - - @{thm pow_nat.simps [no_vars]}. - - \noindent As you can see from this example, for local - definitions you may use any specification tool - which works together with locales (e.g. \cite{krauss2006}). -*} - - -subsection {* A functor analogy *} - -text {* - We introduced Isar classes by analogy to type classes - functional programming; if we reconsider this in the - context of what has been said about type classes and locales, - we can drive this analogy further by stating that type - classes essentially correspond to functors which have - a canonical interpretation as type classes. - Anyway, there is also the possibility of other interpretations. - For example, also @{text list}s form a monoid with - @{text append} and @{term "[]"} as operations, but it - seems inappropriate to apply to lists - the same operations as for genuinely algebraic types. - In such a case, we simply can do a particular interpretation - of monoids for lists: -*} - -interpretation %quote list_monoid!: monoid append "[]" - proof qed auto - -text {* - \noindent This enables us to apply facts on monoids - to lists, e.g. @{thm list_monoid.neutl [no_vars]}. - - When using this interpretation pattern, it may also - be appropriate to map derived definitions accordingly: -*} - -primrec %quote replicate :: "nat \ \ list \ \ list" where - "replicate 0 _ = []" - | "replicate (Suc n) xs = xs @ replicate n xs" - -interpretation %quote list_monoid!: monoid append "[]" where - "monoid.pow_nat append [] = replicate" -proof - - interpret monoid append "[]" .. - show "monoid.pow_nat append [] = replicate" - proof - fix n - show "monoid.pow_nat append [] n = replicate n" - by (induct n) auto - qed -qed intro_locales - - -subsection {* Additional subclass relations *} - -text {* - Any @{text "group"} is also a @{text "monoid"}; this - can be made explicit by claiming an additional - subclass relation, - together with a proof of the logical difference: -*} - -subclass %quote (in group) monoid -proof - fix x - from invl have "x\
\ x = \" by simp - with assoc [symmetric] neutl invl have "x\
\ (x \ \) = x\
\ x" by simp - with left_cancel show "x \ \ = x" by simp -qed - -text {* - \noindent The logical proof is carried out on the locale level. - Afterwards it is propagated - to the type system, making @{text group} an instance of - @{text monoid} by adding an additional edge - to the graph of subclass relations - (cf.\ \figref{fig:subclass}). - - \begin{figure}[htbp] - \begin{center} - \small - \unitlength 0.6mm - \begin{picture}(40,60)(0,0) - \put(20,60){\makebox(0,0){@{text semigroup}}} - \put(20,40){\makebox(0,0){@{text monoidl}}} - \put(00,20){\makebox(0,0){@{text monoid}}} - \put(40,00){\makebox(0,0){@{text group}}} - \put(20,55){\vector(0,-1){10}} - \put(15,35){\vector(-1,-1){10}} - \put(25,35){\vector(1,-3){10}} - \end{picture} - \hspace{8em} - \begin{picture}(40,60)(0,0) - \put(20,60){\makebox(0,0){@{text semigroup}}} - \put(20,40){\makebox(0,0){@{text monoidl}}} - \put(00,20){\makebox(0,0){@{text monoid}}} - \put(40,00){\makebox(0,0){@{text group}}} - \put(20,55){\vector(0,-1){10}} - \put(15,35){\vector(-1,-1){10}} - \put(05,15){\vector(3,-1){30}} - \end{picture} - \caption{Subclass relationship of monoids and groups: - before and after establishing the relationship - @{text "group \ monoid"}; transitive edges left out.} - \label{fig:subclass} - \end{center} - \end{figure} -7 - For illustration, a derived definition - in @{text group} which uses @{text pow_nat}: -*} - -definition %quote (in group) pow_int :: "int \ \ \ \" where - "pow_int k x = (if k >= 0 - then pow_nat (nat k) x - else (pow_nat (nat (- k)) x)\
)" - -text {* - \noindent yields the global definition of - @{term [source] "pow_int \ int \ \\group \ \\group"} - with the corresponding theorem @{thm pow_int_def [no_vars]}. -*} - -subsection {* A note on syntax *} - -text {* - As a commodity, class context syntax allows to refer - to local class operations and their global counterparts - uniformly; type inference resolves ambiguities. For example: -*} - -context %quote semigroup -begin - -term %quote "x \ y" -- {* example 1 *} -term %quote "(x\nat) \ y" -- {* example 2 *} - -end %quote - -term %quote "x \ y" -- {* example 3 *} - -text {* - \noindent Here in example 1, the term refers to the local class operation - @{text "mult [\]"}, whereas in example 2 the type constraint - enforces the global class operation @{text "mult [nat]"}. - In the global context in example 3, the reference is - to the polymorphic global class operation @{text "mult [?\ \ semigroup]"}. -*} - -section {* Further issues *} - -subsection {* Type classes and code generation *} - -text {* - Turning back to the first motivation for type classes, - namely overloading, it is obvious that overloading - stemming from @{command class} statements and - @{command instantiation} - targets naturally maps to Haskell type classes. - The code generator framework \cite{isabelle-codegen} - takes this into account. Concerning target languages - lacking type classes (e.g.~SML), type classes - are implemented by explicit dictionary construction. - As example, let's go back to the power function: -*} - -definition %quote example :: int where - "example = pow_int 10 (-2)" - -text {* - \noindent This maps to Haskell as: -*} - -text %quote {*@{code_stmts example (Haskell)}*} - -text {* - \noindent The whole code in SML with explicit dictionary passing: -*} - -text %quote {*@{code_stmts example (SML)}*} - -subsection {* Inspecting the type class universe *} - -text {* - To facilitate orientation in complex subclass structures, - two diagnostics commands are provided: - - \begin{description} - - \item[@{command "print_classes"}] print a list of all classes - together with associated operations etc. - - \item[@{command "class_deps"}] visualizes the subclass relation - between all classes as a Hasse diagram. - - \end{description} -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Classes/Thy/ROOT.ML --- a/doc-src/IsarAdvanced/Classes/Thy/ROOT.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,6 +0,0 @@ - -(* $Id$ *) - -no_document use_thy "Setup"; - -use_thy "Classes"; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Classes/Thy/Setup.thy --- a/doc-src/IsarAdvanced/Classes/Thy/Setup.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,34 +0,0 @@ -theory Setup -imports Main Code_Integer -uses - "../../../antiquote_setup" - "../../../more_antiquote" -begin - -ML {* Code_Target.code_width := 74 *} - -syntax - "_alpha" :: "type" ("\") - "_alpha_ofsort" :: "sort \ type" ("\()\_" [0] 1000) - "_beta" :: "type" ("\") - "_beta_ofsort" :: "sort \ type" ("\()\_" [0] 1000) - -parse_ast_translation {* - let - fun alpha_ast_tr [] = Syntax.Variable "'a" - | alpha_ast_tr asts = raise Syntax.AST ("alpha_ast_tr", asts); - fun alpha_ofsort_ast_tr [ast] = - Syntax.Appl [Syntax.Constant "_ofsort", Syntax.Variable "'a", ast] - | alpha_ofsort_ast_tr asts = raise Syntax.AST ("alpha_ast_tr", asts); - fun beta_ast_tr [] = Syntax.Variable "'b" - | beta_ast_tr asts = raise Syntax.AST ("beta_ast_tr", asts); - fun beta_ofsort_ast_tr [ast] = - Syntax.Appl [Syntax.Constant "_ofsort", Syntax.Variable "'b", ast] - | beta_ofsort_ast_tr asts = raise Syntax.AST ("beta_ast_tr", asts); - in [ - ("_alpha", alpha_ast_tr), ("_alpha_ofsort", alpha_ofsort_ast_tr), - ("_beta", beta_ast_tr), ("_beta_ofsort", beta_ofsort_ast_tr) - ] end -*} - -end \ No newline at end of file diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Classes/Thy/document/Classes.tex --- a/doc-src/IsarAdvanced/Classes/Thy/document/Classes.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1346 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Classes}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Classes\isanewline -\isakeyword{imports}\ Main\ Setup\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Haskell-style classes with Isabelle/Isar% -} -\isamarkuptrue% -% -\isamarkupsection{Introduction% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Type classes were introduces by Wadler and Blott \cite{wadler89how} - into the Haskell language, to allow for a reasonable implementation - of overloading\footnote{throughout this tutorial, we are referring - to classical Haskell 1.0 type classes, not considering - later additions in expressiveness}. - As a canonical example, a polymorphic equality function - \isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} which is overloaded on different - types for \isa{{\isasymalpha}}, which is achieved by splitting introduction - of the \isa{eq} function from its overloaded definitions by means - of \isa{class} and \isa{instance} declarations: - - \begin{quote} - - \noindent\isa{class\ eq\ where}\footnote{syntax here is a kind of isabellized Haskell} \\ - \hspace*{2ex}\isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} - - \medskip\noindent\isa{instance\ nat\ {\isasymColon}\ eq\ where} \\ - \hspace*{2ex}\isa{eq\ {\isadigit{0}}\ {\isadigit{0}}\ {\isacharequal}\ True} \\ - \hspace*{2ex}\isa{eq\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ False} \\ - \hspace*{2ex}\isa{eq\ {\isacharunderscore}\ {\isadigit{0}}\ {\isacharequal}\ False} \\ - \hspace*{2ex}\isa{eq\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharparenleft}Suc\ m{\isacharparenright}\ {\isacharequal}\ eq\ n\ m} - - \medskip\noindent\isa{instance\ {\isacharparenleft}{\isasymalpha}{\isasymColon}eq{\isacharcomma}\ {\isasymbeta}{\isasymColon}eq{\isacharparenright}\ pair\ {\isasymColon}\ eq\ where} \\ - \hspace*{2ex}\isa{eq\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ eq\ x{\isadigit{1}}\ x{\isadigit{2}}\ {\isasymand}\ eq\ y{\isadigit{1}}\ y{\isadigit{2}}} - - \medskip\noindent\isa{class\ ord\ extends\ eq\ where} \\ - \hspace*{2ex}\isa{less{\isacharunderscore}eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \\ - \hspace*{2ex}\isa{less\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} - - \end{quote} - - \noindent Type variables are annotated with (finitely many) classes; - these annotations are assertions that a particular polymorphic type - provides definitions for overloaded functions. - - Indeed, type classes not only allow for simple overloading - but form a generic calculus, an instance of order-sorted - algebra \cite{Nipkow-Prehofer:1993,nipkow-sorts93,Wenzel:1997:TPHOL}. - - From a software engeneering point of view, type classes - roughly correspond to interfaces in object-oriented languages like Java; - so, it is naturally desirable that type classes do not only - provide functions (class parameters) but also state specifications - implementations must obey. For example, the \isa{class\ eq} - above could be given the following specification, demanding that - \isa{class\ eq} is an equivalence relation obeying reflexivity, - symmetry and transitivity: - - \begin{quote} - - \noindent\isa{class\ eq\ where} \\ - \hspace*{2ex}\isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \\ - \isa{satisfying} \\ - \hspace*{2ex}\isa{refl{\isacharcolon}\ eq\ x\ x} \\ - \hspace*{2ex}\isa{sym{\isacharcolon}\ eq\ x\ y\ {\isasymlongleftrightarrow}\ eq\ x\ y} \\ - \hspace*{2ex}\isa{trans{\isacharcolon}\ eq\ x\ y\ {\isasymand}\ eq\ y\ z\ {\isasymlongrightarrow}\ eq\ x\ z} - - \end{quote} - - \noindent From a theoretic point of view, type classes are lightweight - modules; Haskell type classes may be emulated by - SML functors \cite{classes_modules}. - Isabelle/Isar offers a discipline of type classes which brings - all those aspects together: - - \begin{enumerate} - \item specifying abstract parameters together with - corresponding specifications, - \item instantiating those abstract parameters by a particular - type - \item in connection with a ``less ad-hoc'' approach to overloading, - \item with a direct link to the Isabelle module system - (aka locales \cite{kammueller-locales}). - \end{enumerate} - - \noindent Isar type classes also directly support code generation - in a Haskell like fashion. - - This tutorial demonstrates common elements of structured specifications - and abstract reasoning with type classes by the algebraic hierarchy of - semigroups, monoids and groups. Our background theory is that of - Isabelle/HOL \cite{isa-tutorial}, for which some - familiarity is assumed. - - Here we merely present the look-and-feel for end users. - Internally, those are mapped to more primitive Isabelle concepts. - See \cite{Haftmann-Wenzel:2006:classes} for more detail.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{A simple algebra example \label{sec:example}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Class definition% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Depending on an arbitrary type \isa{{\isasymalpha}}, class \isa{semigroup} introduces a binary operator \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} that is - assumed to be associative:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{class}\isamarkupfalse% -\ semigroup\ {\isacharequal}\isanewline -\ \ \isakeyword{fixes}\ mult\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymotimes}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} specification consists of two - parts: the \qn{operational} part names the class parameter - (\hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}}), the \qn{logical} part specifies properties on them - (\hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}}). The local \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} and - \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} are lifted to the theory toplevel, - yielding the global - parameter \isa{{\isachardoublequote}mult\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequote}} and the - global theorem \hyperlink{fact.semigroup.assoc:}{\mbox{\isa{semigroup{\isachardot}assoc{\isacharcolon}}}}~\isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup{\isachardot}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequote}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Class instantiation \label{sec:class_inst}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The concrete type \isa{int} is made a \isa{semigroup} - instance by providing a suitable definition for the class parameter - \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} and a proof for the specification of \hyperlink{fact.assoc}{\mbox{\isa{assoc}}}. - This is accomplished by the \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} target:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{instantiation}\isamarkupfalse% -\ int\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ mult{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}i\ {\isasymotimes}\ j\ {\isacharequal}\ i\ {\isacharplus}\ {\isacharparenleft}j{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ i\ j\ k\ {\isacharcolon}{\isacharcolon}\ int\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isacharplus}\ j{\isacharparenright}\ {\isacharplus}\ k\ {\isacharequal}\ i\ {\isacharplus}\ {\isacharparenleft}j\ {\isacharplus}\ k{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isasymotimes}\ j{\isacharparenright}\ {\isasymotimes}\ k\ {\isacharequal}\ i\ {\isasymotimes}\ {\isacharparenleft}j\ {\isasymotimes}\ k{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} allows to define class parameters - at a particular instance using common specification tools (here, - \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}). The concluding \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} - opens a proof that the given parameters actually conform - to the class specification. Note that the first proof step - is the \hyperlink{method.default}{\mbox{\isa{default}}} method, - which for such instance proofs maps to the \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isacharunderscore}classes}}} method. - This boils down an instance judgement to the relevant primitive - proof goals and should conveniently always be the first method applied - in an instantiation proof. - - From now on, the type-checker will consider \isa{int} - as a \isa{semigroup} automatically, i.e.\ any general results - are immediately available on concrete instances. - - \medskip Another instance of \isa{semigroup} are the natural numbers:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{instantiation}\isamarkupfalse% -\ nat\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{primrec}\isamarkupfalse% -\ mult{\isacharunderscore}nat\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}Suc\ m\ {\isasymotimes}\ n\ {\isacharequal}\ Suc\ {\isacharparenleft}m\ {\isasymotimes}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ m\ n\ q\ {\isacharcolon}{\isacharcolon}\ nat\ \isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymotimes}\ q\ {\isacharequal}\ m\ {\isasymotimes}\ {\isacharparenleft}n\ {\isasymotimes}\ q{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ m{\isacharparenright}\ auto\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Note the occurence of the name \isa{mult{\isacharunderscore}nat} - in the primrec declaration; by default, the local name of - a class operation \isa{f} to instantiate on type constructor - \isa{{\isasymkappa}} are mangled as \isa{f{\isacharunderscore}{\isasymkappa}}. In case of uncertainty, - these names may be inspected using the \hyperlink{command.print-context}{\mbox{\isa{\isacommand{print{\isacharunderscore}context}}}} command - or the corresponding ProofGeneral button.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Lifting and parametric types% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Overloaded definitions giving on class instantiation - may include recursion over the syntactic structure of types. - As a canonical example, we model product semigroups - using our simple algebra:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{instantiation}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}semigroup{\isacharcomma}\ semigroup{\isacharparenright}\ semigroup\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ mult{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{2}}\ {\isacharequal}\ {\isacharparenleft}fst\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ fst\ p\isactrlisub {\isadigit{2}}{\isacharcomma}\ snd\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ snd\ p\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ p\isactrlisub {\isadigit{1}}\ p\isactrlisub {\isadigit{2}}\ p\isactrlisub {\isadigit{3}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}semigroup\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}semigroup{\isachardoublequoteclose}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{2}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{3}}\ {\isacharequal}\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ {\isacharparenleft}p\isactrlisub {\isadigit{2}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{3}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ assoc{\isacharparenright}\isanewline -\isacommand{qed}\isamarkupfalse% -\ \ \ \ \ \ \isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Associativity from product semigroups is - established using - the definition of \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} on products and the hypothetical - associativity of the type components; these hypotheses - are facts due to the \isa{semigroup} constraints imposed - on the type components by the \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} proposition. - Indeed, this pattern often occurs with parametric types - and type classes.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Subclassing% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -We define a subclass \isa{monoidl} (a semigroup with a left-hand neutral) - by extending \isa{semigroup} - with one additional parameter \isa{neutral} together - with its property:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{class}\isamarkupfalse% -\ monoidl\ {\isacharequal}\ semigroup\ {\isacharplus}\isanewline -\ \ \isakeyword{fixes}\ neutral\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isachardoublequoteclose}\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ neutl{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Again, we prove some instances, by - providing suitable parameter definitions and proofs for the - additional specifications. Observe that instantiations - for types with the same arity may be simultaneous:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{instantiation}\isamarkupfalse% -\ nat\ \isakeyword{and}\ int\ {\isacharcolon}{\isacharcolon}\ monoidl\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ neutral{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ neutral{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ k\ {\isacharequal}\ k{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoidl{\isacharcomma}\ monoidl{\isacharparenright}\ monoidl\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ neutral{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isasymone}{\isacharcomma}\ {\isasymone}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ p\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}monoidl\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}monoidl{\isachardoublequoteclose}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ neutral{\isacharunderscore}prod{\isacharunderscore}def\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ neutl{\isacharparenright}\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Fully-fledged monoids are modelled by another subclass - which does not add new parameters but tightens the specification:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{class}\isamarkupfalse% -\ monoid\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline -\ \ \isakeyword{assumes}\ neutr{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ nat\ \isakeyword{and}\ int\ {\isacharcolon}{\isacharcolon}\ monoid\ \isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}n\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ neutral{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ n{\isacharparenright}\ simp{\isacharunderscore}all\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}k\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ k{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoid{\isacharcomma}\ monoid{\isacharparenright}\ monoid\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\ \isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ p\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}monoid\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}monoid{\isachardoublequoteclose}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}p\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ neutral{\isacharunderscore}prod{\isacharunderscore}def\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ neutr{\isacharparenright}\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent To finish our small algebra example, we add a \isa{group} class - with a corresponding instance:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{class}\isamarkupfalse% -\ group\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline -\ \ \isakeyword{fixes}\ inverse\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}{\isachardoublequoteopen}{\isacharparenleft}{\isacharunderscore}{\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ invl{\isacharcolon}\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ int\ {\isacharcolon}{\isacharcolon}\ group\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ inverse{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}i{\isasymdiv}\ {\isacharequal}\ {\isacharminus}\ {\isacharparenleft}i{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ i\ {\isacharcolon}{\isacharcolon}\ int\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharminus}i\ {\isacharplus}\ i\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}i{\isasymdiv}\ {\isasymotimes}\ i\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{unfolding}\isamarkupfalse% -\ mult{\isacharunderscore}int{\isacharunderscore}def\ neutral{\isacharunderscore}int{\isacharunderscore}def\ inverse{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isamarkupsection{Type classes as locales% -} -\isamarkuptrue% -% -\isamarkupsubsection{A look behind the scene% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The example above gives an impression how Isar type classes work - in practice. As stated in the introduction, classes also provide - a link to Isar's locale system. Indeed, the logical core of a class - is nothing else than a locale:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{class}\isamarkupfalse% -\ idem\ {\isacharequal}\isanewline -\ \ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{assumes}\ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent essentially introduces the locale% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadeliminvisible -\ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{setup}\isamarkupfalse% -\ {\isacharverbatimopen}\ Sign{\isachardot}add{\isacharunderscore}path\ {\isachardoublequote}foo{\isachardoublequote}\ {\isacharverbatimclose}% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -\isanewline -% -\isadelimquote -\isanewline -% -\endisadelimquote -% -\isatagquote -\isacommand{locale}\isamarkupfalse% -\ idem\ {\isacharequal}\isanewline -\ \ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{assumes}\ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent together with corresponding constant(s):% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{consts}\isamarkupfalse% -\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The connection to the type system is done by means - of a primitive axclass% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadeliminvisible -\ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{setup}\isamarkupfalse% -\ {\isacharverbatimopen}\ Sign{\isachardot}add{\isacharunderscore}path\ {\isachardoublequote}foo{\isachardoublequote}\ {\isacharverbatimclose}% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -\isanewline -% -\isadelimquote -\isanewline -% -\endisadelimquote -% -\isatagquote -\isacommand{axclass}\isamarkupfalse% -\ idem\ {\isacharless}\ type\isanewline -\ \ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isadeliminvisible -\ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{setup}\isamarkupfalse% -\ {\isacharverbatimopen}\ Sign{\isachardot}parent{\isacharunderscore}path\ {\isacharverbatimclose}% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\begin{isamarkuptext}% -\noindent together with a corresponding interpretation:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{interpretation}\isamarkupfalse% -\ idem{\isacharunderscore}class{\isacharcolon}\isanewline -\ \ idem\ {\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isasymalpha}{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline -\isacommand{proof}\isamarkupfalse% -\ \isacommand{qed}\isamarkupfalse% -\ {\isacharparenleft}rule\ idem{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This gives you at hand the full power of the Isabelle module system; - conclusions in locale \isa{idem} are implicitly propagated - to class \isa{idem}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadeliminvisible -\ % -\endisadeliminvisible -% -\isataginvisible -\isacommand{setup}\isamarkupfalse% -\ {\isacharverbatimopen}\ Sign{\isachardot}parent{\isacharunderscore}path\ {\isacharverbatimclose}% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isamarkupsubsection{Abstract reasoning% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle locales enable reasoning at a general level, while results - are implicitly transferred to all instances. For example, we can - now establish the \isa{left{\isacharunderscore}cancel} lemma for groups, which - states that the function \isa{{\isacharparenleft}x\ {\isasymotimes}{\isacharparenright}} is injective:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ left{\isacharunderscore}cancel{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline -\isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z{\isachardoublequoteclose}\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isacharequal}\ x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymdiv}\ {\isasymotimes}\ x{\isacharparenright}\ {\isasymotimes}\ y\ {\isacharequal}\ {\isacharparenleft}x{\isasymdiv}\ {\isasymotimes}\ x{\isacharparenright}\ {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% -\ assoc\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% -\ neutl\ \isakeyword{and}\ invl\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Here the \qt{\hyperlink{keyword.in}{\mbox{\isa{\isakeyword{in}}}} \isa{group}} target specification - indicates that the result is recorded within that context for later - use. This local theorem is also lifted to the global one \hyperlink{fact.group.left-cancel:}{\mbox{\isa{group{\isachardot}left{\isacharunderscore}cancel{\isacharcolon}}}} \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}group{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}. Since type \isa{int} has been made an instance of - \isa{group} before, we may refer to that fact as well: \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ int{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Derived definitions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle locales support a concept of local definitions - in locales:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{primrec}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ monoid{\isacharparenright}\ pow{\isacharunderscore}nat\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent If the locale \isa{group} is also a class, this local - definition is propagated onto a global definition of - \isa{{\isachardoublequote}pow{\isacharunderscore}nat\ {\isasymColon}\ nat\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid{\isachardoublequote}} - with corresponding theorems - - \isa{pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}\isasep\isanewline% -pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x}. - - \noindent As you can see from this example, for local - definitions you may use any specification tool - which works together with locales (e.g. \cite{krauss2006}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{A functor analogy% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -We introduced Isar classes by analogy to type classes - functional programming; if we reconsider this in the - context of what has been said about type classes and locales, - we can drive this analogy further by stating that type - classes essentially correspond to functors which have - a canonical interpretation as type classes. - Anyway, there is also the possibility of other interpretations. - For example, also \isa{list}s form a monoid with - \isa{append} and \isa{{\isacharbrackleft}{\isacharbrackright}} as operations, but it - seems inappropriate to apply to lists - the same operations as for genuinely algebraic types. - In such a case, we simply can do a particular interpretation - of monoids for lists:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{interpretation}\isamarkupfalse% -\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ \isacommand{qed}\isamarkupfalse% -\ auto% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This enables us to apply facts on monoids - to lists, e.g. \isa{{\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ x\ {\isacharequal}\ x}. - - When using this interpretation pattern, it may also - be appropriate to map derived definitions accordingly:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{primrec}\isamarkupfalse% -\ replicate\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isasymalpha}\ list\ {\isasymRightarrow}\ {\isasymalpha}\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}replicate\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}replicate\ {\isacharparenleft}Suc\ n{\isacharparenright}\ xs\ {\isacharequal}\ xs\ {\isacharat}\ replicate\ n\ xs{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{interpretation}\isamarkupfalse% -\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline -\isacommand{proof}\isamarkupfalse% -\ {\isacharminus}\isanewline -\ \ \isacommand{interpret}\isamarkupfalse% -\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ n\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ n\ {\isacharequal}\ replicate\ n{\isachardoublequoteclose}\isanewline -\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ n{\isacharparenright}\ auto\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -\ intro{\isacharunderscore}locales% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isamarkupsubsection{Additional subclass relations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Any \isa{group} is also a \isa{monoid}; this - can be made explicit by claiming an additional - subclass relation, - together with a proof of the logical difference:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{subclass}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ monoid\isanewline -\isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ x\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ invl\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{with}\isamarkupfalse% -\ assoc\ {\isacharbrackleft}symmetric{\isacharbrackright}\ neutl\ invl\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ {\isasymone}{\isacharparenright}\ {\isacharequal}\ x{\isasymdiv}\ {\isasymotimes}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{with}\isamarkupfalse% -\ left{\isacharunderscore}cancel\ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The logical proof is carried out on the locale level. - Afterwards it is propagated - to the type system, making \isa{group} an instance of - \isa{monoid} by adding an additional edge - to the graph of subclass relations - (cf.\ \figref{fig:subclass}). - - \begin{figure}[htbp] - \begin{center} - \small - \unitlength 0.6mm - \begin{picture}(40,60)(0,0) - \put(20,60){\makebox(0,0){\isa{semigroup}}} - \put(20,40){\makebox(0,0){\isa{monoidl}}} - \put(00,20){\makebox(0,0){\isa{monoid}}} - \put(40,00){\makebox(0,0){\isa{group}}} - \put(20,55){\vector(0,-1){10}} - \put(15,35){\vector(-1,-1){10}} - \put(25,35){\vector(1,-3){10}} - \end{picture} - \hspace{8em} - \begin{picture}(40,60)(0,0) - \put(20,60){\makebox(0,0){\isa{semigroup}}} - \put(20,40){\makebox(0,0){\isa{monoidl}}} - \put(00,20){\makebox(0,0){\isa{monoid}}} - \put(40,00){\makebox(0,0){\isa{group}}} - \put(20,55){\vector(0,-1){10}} - \put(15,35){\vector(-1,-1){10}} - \put(05,15){\vector(3,-1){30}} - \end{picture} - \caption{Subclass relationship of monoids and groups: - before and after establishing the relationship - \isa{group\ {\isasymsubseteq}\ monoid}; transitive edges left out.} - \label{fig:subclass} - \end{center} - \end{figure} -7 - For illustration, a derived definition - in \isa{group} which uses \isa{pow{\isacharunderscore}nat}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{definition}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ pow{\isacharunderscore}int\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}int\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ k\ {\isachargreater}{\isacharequal}\ {\isadigit{0}}\isanewline -\ \ \ \ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\isanewline -\ \ \ \ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent yields the global definition of - \isa{{\isachardoublequote}pow{\isacharunderscore}int\ {\isasymColon}\ int\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group{\isachardoublequote}} - with the corresponding theorem \isa{pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{0}}\ {\isasymle}\ k\ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{A note on syntax% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -As a commodity, class context syntax allows to refer - to local class operations and their global counterparts - uniformly; type inference resolves ambiguities. For example:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{context}\isamarkupfalse% -\ semigroup\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{term}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % -\isamarkupcmt{example 1% -} -\isanewline -\isacommand{term}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymColon}nat{\isacharparenright}\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % -\isamarkupcmt{example 2% -} -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{term}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % -\isamarkupcmt{example 3% -} -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Here in example 1, the term refers to the local class operation - \isa{mult\ {\isacharbrackleft}{\isasymalpha}{\isacharbrackright}}, whereas in example 2 the type constraint - enforces the global class operation \isa{mult\ {\isacharbrackleft}nat{\isacharbrackright}}. - In the global context in example 3, the reference is - to the polymorphic global class operation \isa{mult\ {\isacharbrackleft}{\isacharquery}{\isasymalpha}\ {\isasymColon}\ semigroup{\isacharbrackright}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Further issues% -} -\isamarkuptrue% -% -\isamarkupsubsection{Type classes and code generation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Turning back to the first motivation for type classes, - namely overloading, it is obvious that overloading - stemming from \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} statements and - \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} - targets naturally maps to Haskell type classes. - The code generator framework \cite{isabelle-codegen} - takes this into account. Concerning target languages - lacking type classes (e.g.~SML), type classes - are implemented by explicit dictionary construction. - As example, let's go back to the power function:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{definition}\isamarkupfalse% -\ example\ {\isacharcolon}{\isacharcolon}\ int\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}example\ {\isacharequal}\ pow{\isacharunderscore}int\ {\isadigit{1}}{\isadigit{0}}\ {\isacharparenleft}{\isacharminus}{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This maps to Haskell as:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}module Example where {\char123}\\ -\hspace*{0pt}\\ -\hspace*{0pt}\\ -\hspace*{0pt}data Nat = Suc Nat | Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}nat{\char95}aux ::~Integer -> Nat -> Nat;\\ -\hspace*{0pt}nat{\char95}aux i n = (if i <= 0 then n else nat{\char95}aux (i - 1) (Suc n));\\ -\hspace*{0pt}\\ -\hspace*{0pt}nat ::~Integer -> Nat;\\ -\hspace*{0pt}nat i = nat{\char95}aux i Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}class Semigroup a where {\char123}\\ -\hspace*{0pt} ~mult ::~a -> a -> a;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}class (Semigroup a) => Monoidl a where {\char123}\\ -\hspace*{0pt} ~neutral ::~a;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}class (Monoidl a) => Monoid a where {\char123}\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}class (Monoid a) => Group a where {\char123}\\ -\hspace*{0pt} ~inverse ::~a -> a;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}inverse{\char95}int ::~Integer -> Integer;\\ -\hspace*{0pt}inverse{\char95}int i = negate i;\\ -\hspace*{0pt}\\ -\hspace*{0pt}neutral{\char95}int ::~Integer;\\ -\hspace*{0pt}neutral{\char95}int = 0;\\ -\hspace*{0pt}\\ -\hspace*{0pt}mult{\char95}int ::~Integer -> Integer -> Integer;\\ -\hspace*{0pt}mult{\char95}int i j = i + j;\\ -\hspace*{0pt}\\ -\hspace*{0pt}instance Semigroup Integer where {\char123}\\ -\hspace*{0pt} ~mult = mult{\char95}int;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}instance Monoidl Integer where {\char123}\\ -\hspace*{0pt} ~neutral = neutral{\char95}int;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}instance Monoid Integer where {\char123}\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}instance Group Integer where {\char123}\\ -\hspace*{0pt} ~inverse = inverse{\char95}int;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}pow{\char95}nat ::~forall a.~(Monoid a) => Nat -> a -> a;\\ -\hspace*{0pt}pow{\char95}nat Zero{\char95}nat x = neutral;\\ -\hspace*{0pt}pow{\char95}nat (Suc n) x = mult x (pow{\char95}nat n x);\\ -\hspace*{0pt}\\ -\hspace*{0pt}pow{\char95}int ::~forall a.~(Group a) => Integer -> a -> a;\\ -\hspace*{0pt}pow{\char95}int k x =\\ -\hspace*{0pt} ~(if 0 <= k then pow{\char95}nat (nat k) x\\ -\hspace*{0pt} ~~~else inverse (pow{\char95}nat (nat (negate k)) x));\\ -\hspace*{0pt}\\ -\hspace*{0pt}example ::~Integer;\\ -\hspace*{0pt}example = pow{\char95}int 10 (-2);\\ -\hspace*{0pt}\\ -\hspace*{0pt}{\char125}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The whole code in SML with explicit dictionary passing:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype nat = Suc of nat | Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun nat{\char95}aux i n =\\ -\hspace*{0pt} ~(if IntInf.<= (i,~(0 :~IntInf.int)) then n\\ -\hspace*{0pt} ~~~else nat{\char95}aux (IntInf.- (i,~(1 :~IntInf.int))) (Suc n));\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun nat i = nat{\char95}aux i Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a semigroup = {\char123}mult :~'a -> 'a -> 'a{\char125};\\ -\hspace*{0pt}fun mult (A{\char95}:'a semigroup) = {\char35}mult A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a monoidl =\\ -\hspace*{0pt} ~{\char123}Classes{\char95}{\char95}semigroup{\char95}monoidl :~'a semigroup,~neutral :~'a{\char125};\\ -\hspace*{0pt}fun semigroup{\char95}monoidl (A{\char95}:'a monoidl) = {\char35}Classes{\char95}{\char95}semigroup{\char95}monoidl A{\char95};\\ -\hspace*{0pt}fun neutral (A{\char95}:'a monoidl) = {\char35}neutral A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a monoid = {\char123}Classes{\char95}{\char95}monoidl{\char95}monoid :~'a monoidl{\char125};\\ -\hspace*{0pt}fun monoidl{\char95}monoid (A{\char95}:'a monoid) = {\char35}Classes{\char95}{\char95}monoidl{\char95}monoid A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a group = {\char123}Classes{\char95}{\char95}monoid{\char95}group :~'a monoid,~inverse :~'a -> 'a{\char125};\\ -\hspace*{0pt}fun monoid{\char95}group (A{\char95}:'a group) = {\char35}Classes{\char95}{\char95}monoid{\char95}group A{\char95};\\ -\hspace*{0pt}fun inverse (A{\char95}:'a group) = {\char35}inverse A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun inverse{\char95}int i = IntInf.{\char126}~i;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val neutral{\char95}int :~IntInf.int = (0 :~IntInf.int)\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun mult{\char95}int i j = IntInf.+ (i,~j);\\ -\hspace*{0pt}\\ -\hspace*{0pt}val semigroup{\char95}int = {\char123}mult = mult{\char95}int{\char125}~:~IntInf.int semigroup;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val monoidl{\char95}int =\\ -\hspace*{0pt} ~{\char123}Classes{\char95}{\char95}semigroup{\char95}monoidl = semigroup{\char95}int,~neutral = neutral{\char95}int{\char125}~:\\ -\hspace*{0pt} ~IntInf.int monoidl;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val monoid{\char95}int = {\char123}Classes{\char95}{\char95}monoidl{\char95}monoid = monoidl{\char95}int{\char125}~:\\ -\hspace*{0pt} ~IntInf.int monoid;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val group{\char95}int =\\ -\hspace*{0pt} ~{\char123}Classes{\char95}{\char95}monoid{\char95}group = monoid{\char95}int,~inverse = inverse{\char95}int{\char125}~:\\ -\hspace*{0pt} ~IntInf.int group;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun pow{\char95}nat A{\char95}~Zero{\char95}nat x = neutral (monoidl{\char95}monoid A{\char95})\\ -\hspace*{0pt} ~| pow{\char95}nat A{\char95}~(Suc n) x =\\ -\hspace*{0pt} ~~~mult ((semigroup{\char95}monoidl o monoidl{\char95}monoid) A{\char95}) x (pow{\char95}nat A{\char95}~n x);\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun pow{\char95}int A{\char95}~k x =\\ -\hspace*{0pt} ~(if IntInf.<= ((0 :~IntInf.int),~k)\\ -\hspace*{0pt} ~~~then pow{\char95}nat (monoid{\char95}group A{\char95}) (nat k) x\\ -\hspace*{0pt} ~~~else inverse A{\char95}~(pow{\char95}nat (monoid{\char95}group A{\char95}) (nat (IntInf.{\char126}~k)) x));\\ -\hspace*{0pt}\\ -\hspace*{0pt}val example :~IntInf.int =\\ -\hspace*{0pt} ~pow{\char95}int group{\char95}int (10 :~IntInf.int) ({\char126}2 :~IntInf.int)\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isamarkupsubsection{Inspecting the type class universe% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -To facilitate orientation in complex subclass structures, - two diagnostics commands are provided: - - \begin{description} - - \item[\hyperlink{command.print-classes}{\mbox{\isa{\isacommand{print{\isacharunderscore}classes}}}}] print a list of all classes - together with associated operations etc. - - \item[\hyperlink{command.class-deps}{\mbox{\isa{\isacommand{class{\isacharunderscore}deps}}}}] visualizes the subclass relation - between all classes as a Hasse diagram. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Classes/classes.tex --- a/doc-src/IsarAdvanced/Classes/classes.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,50 +0,0 @@ - -\documentclass[12pt,a4paper,fleqn]{report} -\usepackage{latexsym,graphicx} -\usepackage[refpage]{nomencl} -\usepackage{../../iman,../../extra,../../isar,../../proof} -\usepackage{../../isabelle,../../isabellesym} -\usepackage{style} -\usepackage{../../pdfsetup} - - -\hyphenation{Isabelle} -\hyphenation{Isar} -\isadroptag{theory} - -\title{\includegraphics[scale=0.5]{isabelle_isar} - \\[4ex] Haskell-style type classes with Isabelle/Isar} -\author{\emph{Florian Haftmann}} - -\begin{document} - -\maketitle - -\begin{abstract} - This tutorial introduces the look-and-feel of Isar type classes - to the end-user; Isar type classes are a convenient mechanism - for organizing specifications, overcoming some drawbacks - of raw axiomatic type classes. Essentially, they combine - an operational aspect (in the manner of Haskell) with - a logical aspect, both managed uniformly. -\end{abstract} - -\thispagestyle{empty}\clearpage - -\pagenumbering{roman} -\clearfirst - -\input{Thy/document/Classes.tex} - -\begingroup -\bibliographystyle{plain} \small\raggedright\frenchspacing -\bibliography{../../manual} -\endgroup - -\end{document} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: t -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Classes/style.sty --- a/doc-src/IsarAdvanced/Classes/style.sty Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,43 +0,0 @@ - -%% toc -\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} -\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} - -%% references -\newcommand{\secref}[1]{\S\ref{#1}} -\newcommand{\figref}[1]{figure~\ref{#1}} - -%% logical markup -\newcommand{\strong}[1]{{\bfseries {#1}}} -\newcommand{\qn}[1]{\emph{#1}} - -%% typographic conventions -\newcommand{\qt}[1]{``{#1}''} - -%% verbatim text -\newcommand{\isatypewriter}{\fontsize{9pt}{0pt}\tt\renewcommand{\baselinestretch}{1}\setlength{\baselineskip}{9pt}} - -%% quoted segments -\makeatletter -\isakeeptag{quote} -\newenvironment{quotesegment}{\begin{quote}\isa@parindent\parindent\parindent0pt\isa@parskip\parskip\parskip0pt}{\end{quote}} -\renewcommand{\isatagquote}{\begin{quotesegment}} -\renewcommand{\endisatagquote}{\end{quotesegment}} -\makeatother - -%% presentation -\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} - -\pagestyle{headings} -\binperiod -\underscoreoff - -\renewcommand{\isadigit}[1]{\isamath{#1}} - -\isabellestyle{it} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "implementation" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/IsaMakefile --- a/doc-src/IsarAdvanced/Codegen/IsaMakefile Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,33 +0,0 @@ - -## targets - -default: Thy -images: -test: Thy - -all: images test - - -## global settings - -SRC = $(ISABELLE_HOME)/src -OUT = $(ISABELLE_OUTPUT) -LOG = $(OUT)/log - -USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -d false -C false -D document - - -## Thy - -THY = $(LOG)/HOL-Thy.gz - -Thy: $(THY) - -$(THY): Thy/ROOT.ML Thy/*.thy ../../antiquote_setup.ML - @$(USEDIR) HOL Thy - - -## clean - -clean: - @rm -f $(THY) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Makefile --- a/doc-src/IsarAdvanced/Codegen/Makefile Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,38 +0,0 @@ -# -# $Id$ -# - -## targets - -default: dvi - - -## dependencies - -include ../Makefile.in - -NAME = codegen - -FILES = $(NAME).tex Thy/document/*.tex \ - style.sty ../../iman.sty ../../extra.sty ../../isar.sty \ - ../../isabelle.sty ../../isabellesym.sty ../../pdfsetup.sty \ - ../../manual.bib ../../proof.sty - -dvi: $(NAME).dvi - -$(NAME).dvi: $(FILES) isabelle_isar.eps codegen_process.ps - $(LATEX) $(NAME) - $(BIBTEX) $(NAME) - $(LATEX) $(NAME) - $(LATEX) $(NAME) - -pdf: $(NAME).pdf - -$(NAME).pdf: $(FILES) isabelle_isar.pdf codegen_process.pdf - $(PDFLATEX) $(NAME) - $(BIBTEX) $(NAME) - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) - $(FIXBOOKMARKS) $(NAME).out - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/Adaption.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Adaption.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,361 +0,0 @@ -theory Adaption -imports Setup -begin - -setup %invisible {* Code_Target.extend_target ("\", ("SML", K I)) *} - -section {* Adaption to target languages \label{sec:adaption} *} - -subsection {* Adapting code generation *} - -text {* - The aspects of code generation introduced so far have two aspects - in common: - - \begin{itemize} - \item They act uniformly, without reference to a specific - target language. - \item They are \emph{safe} in the sense that as long as you trust - the code generator meta theory and implementation, you cannot - produce programs that yield results which are not derivable - in the logic. - \end{itemize} - - \noindent In this section we will introduce means to \emph{adapt} the serialiser - to a specific target language, i.e.~to print program fragments - in a way which accommodates \qt{already existing} ingredients of - a target language environment, for three reasons: - - \begin{itemize} - \item improving readability and aesthetics of generated code - \item gaining efficiency - \item interface with language parts which have no direct counterpart - in @{text "HOL"} (say, imperative data structures) - \end{itemize} - - \noindent Generally, you should avoid using those features yourself - \emph{at any cost}: - - \begin{itemize} - \item The safe configuration methods act uniformly on every target language, - whereas for adaption you have to treat each target language separate. - \item Application is extremely tedious since there is no abstraction - which would allow for a static check, making it easy to produce garbage. - \item More or less subtle errors can be introduced unconsciously. - \end{itemize} - - \noindent However, even if you ought refrain from setting up adaption - yourself, already the @{text "HOL"} comes with some reasonable default - adaptions (say, using target language list syntax). There also some - common adaption cases which you can setup by importing particular - library theories. In order to understand these, we provide some clues here; - these however are not supposed to replace a careful study of the sources. -*} - -subsection {* The adaption principle *} - -text {* - The following figure illustrates what \qt{adaption} is conceptually - supposed to be: - - \begin{figure}[here] - \begin{tikzpicture}[scale = 0.5] - \tikzstyle water=[color = blue, thick] - \tikzstyle ice=[color = black, very thick, cap = round, join = round, fill = white] - \tikzstyle process=[color = green, semithick, ->] - \tikzstyle adaption=[color = red, semithick, ->] - \tikzstyle target=[color = black] - \foreach \x in {0, ..., 24} - \draw[style=water] (\x, 0.25) sin + (0.25, 0.25) cos + (0.25, -0.25) sin - + (0.25, -0.25) cos + (0.25, 0.25); - \draw[style=ice] (1, 0) -- - (3, 6) node[above, fill=white] {logic} -- (5, 0) -- cycle; - \draw[style=ice] (9, 0) -- - (11, 6) node[above, fill=white] {intermediate language} -- (13, 0) -- cycle; - \draw[style=ice] (15, -6) -- - (19, 6) node[above, fill=white] {target language} -- (23, -6) -- cycle; - \draw[style=process] - (3.5, 3) .. controls (7, 5) .. node[fill=white] {translation} (10.5, 3); - \draw[style=process] - (11.5, 3) .. controls (15, 5) .. node[fill=white] (serialisation) {serialisation} (18.5, 3); - \node (adaption) at (11, -2) [style=adaption] {adaption}; - \node at (19, 3) [rotate=90] {generated}; - \node at (19.5, -5) {language}; - \node at (19.5, -3) {library}; - \node (includes) at (19.5, -1) {includes}; - \node (reserved) at (16.5, -3) [rotate=72] {reserved}; % proper 71.57 - \draw[style=process] - (includes) -- (serialisation); - \draw[style=process] - (reserved) -- (serialisation); - \draw[style=adaption] - (adaption) -- (serialisation); - \draw[style=adaption] - (adaption) -- (includes); - \draw[style=adaption] - (adaption) -- (reserved); - \end{tikzpicture} - \caption{The adaption principle} - \label{fig:adaption} - \end{figure} - - \noindent In the tame view, code generation acts as broker between - @{text logic}, @{text "intermediate language"} and - @{text "target language"} by means of @{text translation} and - @{text serialisation}; for the latter, the serialiser has to observe - the structure of the @{text language} itself plus some @{text reserved} - keywords which have to be avoided for generated code. - However, if you consider @{text adaption} mechanisms, the code generated - by the serializer is just the tip of the iceberg: - - \begin{itemize} - \item @{text serialisation} can be \emph{parametrised} such that - logical entities are mapped to target-specific ones - (e.g. target-specific list syntax, - see also \secref{sec:adaption_mechanisms}) - \item Such parametrisations can involve references to a - target-specific standard @{text library} (e.g. using - the @{text Haskell} @{verbatim Maybe} type instead - of the @{text HOL} @{type "option"} type); - if such are used, the corresponding identifiers - (in our example, @{verbatim Maybe}, @{verbatim Nothing} - and @{verbatim Just}) also have to be considered @{text reserved}. - \item Even more, the user can enrich the library of the - target-language by providing code snippets - (\qt{@{text "includes"}}) which are prepended to - any generated code (see \secref{sec:include}); this typically - also involves further @{text reserved} identifiers. - \end{itemize} - - \noindent As figure \ref{fig:adaption} illustrates, all these adaption mechanisms - have to act consistently; it is at the discretion of the user - to take care for this. -*} - -subsection {* Common adaption patterns *} - -text {* - The @{theory HOL} @{theory Main} theory already provides a code - generator setup - which should be suitable for most applications. Common extensions - and modifications are available by certain theories of the @{text HOL} - library; beside being useful in applications, they may serve - as a tutorial for customising the code generator setup (see below - \secref{sec:adaption_mechanisms}). - - \begin{description} - - \item[@{theory "Code_Integer"}] represents @{text HOL} integers by big - integer literals in target languages. - \item[@{theory "Code_Char"}] represents @{text HOL} characters by - character literals in target languages. - \item[@{theory "Code_Char_chr"}] like @{text "Code_Char"}, - but also offers treatment of character codes; includes - @{theory "Code_Char"}. - \item[@{theory "Efficient_Nat"}] \label{eff_nat} implements natural numbers by integers, - which in general will result in higher efficiency; pattern - matching with @{term "0\nat"} / @{const "Suc"} - is eliminated; includes @{theory "Code_Integer"} - and @{theory "Code_Index"}. - \item[@{theory "Code_Index"}] provides an additional datatype - @{typ index} which is mapped to target-language built-in integers. - Useful for code setups which involve e.g. indexing of - target-language arrays. - \item[@{theory "Code_Message"}] provides an additional datatype - @{typ message_string} which is isomorphic to strings; - @{typ message_string}s are mapped to target-language strings. - Useful for code setups which involve e.g. printing (error) messages. - - \end{description} - - \begin{warn} - When importing any of these theories, they should form the last - items in an import list. Since these theories adapt the - code generator setup in a non-conservative fashion, - strange effects may occur otherwise. - \end{warn} -*} - - -subsection {* Parametrising serialisation \label{sec:adaption_mechanisms} *} - -text {* - Consider the following function and its corresponding - SML code: -*} - -primrec %quote in_interval :: "nat \ nat \ nat \ bool" where - "in_interval (k, l) n \ k \ n \ n \ l" -(*<*) -code_type %invisible bool - (SML) -code_const %invisible True and False and "op \" and Not - (SML and and and) -(*>*) -text %quote {*@{code_stmts in_interval (SML)}*} - -text {* - \noindent Though this is correct code, it is a little bit unsatisfactory: - boolean values and operators are materialised as distinguished - entities with have nothing to do with the SML-built-in notion - of \qt{bool}. This results in less readable code; - additionally, eager evaluation may cause programs to - loop or break which would perfectly terminate when - the existing SML @{verbatim "bool"} would be used. To map - the HOL @{typ bool} on SML @{verbatim "bool"}, we may use - \qn{custom serialisations}: -*} - -code_type %quotett bool - (SML "bool") -code_const %quotett True and False and "op \" - (SML "true" and "false" and "_ andalso _") - -text {* - \noindent The @{command code_type} command takes a type constructor - as arguments together with a list of custom serialisations. - Each custom serialisation starts with a target language - identifier followed by an expression, which during - code serialisation is inserted whenever the type constructor - would occur. For constants, @{command code_const} implements - the corresponding mechanism. Each ``@{verbatim "_"}'' in - a serialisation expression is treated as a placeholder - for the type constructor's (the constant's) arguments. -*} - -text %quote {*@{code_stmts in_interval (SML)}*} - -text {* - \noindent This still is not perfect: the parentheses - around the \qt{andalso} expression are superfluous. - Though the serialiser - by no means attempts to imitate the rich Isabelle syntax - framework, it provides some common idioms, notably - associative infixes with precedences which may be used here: -*} - -code_const %quotett "op \" - (SML infixl 1 "andalso") - -text %quote {*@{code_stmts in_interval (SML)}*} - -text {* - \noindent The attentive reader may ask how we assert that no generated - code will accidentally overwrite. For this reason the serialiser has - an internal table of identifiers which have to be avoided to be used - for new declarations. Initially, this table typically contains the - keywords of the target language. It can be extended manually, thus avoiding - accidental overwrites, using the @{command "code_reserved"} command: -*} - -code_reserved %quote "\" bool true false andalso - -text {* - \noindent Next, we try to map HOL pairs to SML pairs, using the - infix ``@{verbatim "*"}'' type constructor and parentheses: -*} -(*<*) -code_type %invisible * - (SML) -code_const %invisible Pair - (SML) -(*>*) -code_type %quotett * - (SML infix 2 "*") -code_const %quotett Pair - (SML "!((_),/ (_))") - -text {* - \noindent The initial bang ``@{verbatim "!"}'' tells the serialiser - never to put - parentheses around the whole expression (they are already present), - while the parentheses around argument place holders - tell not to put parentheses around the arguments. - The slash ``@{verbatim "/"}'' (followed by arbitrary white space) - inserts a space which may be used as a break if necessary - during pretty printing. - - These examples give a glimpse what mechanisms - custom serialisations provide; however their usage - requires careful thinking in order not to introduce - inconsistencies -- or, in other words: - custom serialisations are completely axiomatic. - - A further noteworthy details is that any special - character in a custom serialisation may be quoted - using ``@{verbatim "'"}''; thus, in - ``@{verbatim "fn '_ => _"}'' the first - ``@{verbatim "_"}'' is a proper underscore while the - second ``@{verbatim "_"}'' is a placeholder. -*} - - -subsection {* @{text Haskell} serialisation *} - -text {* - For convenience, the default - @{text HOL} setup for @{text Haskell} maps the @{class eq} class to - its counterpart in @{text Haskell}, giving custom serialisations - for the class @{class eq} (by command @{command code_class}) and its operation - @{const HOL.eq} -*} - -code_class %quotett eq - (Haskell "Eq") - -code_const %quotett "op =" - (Haskell infixl 4 "==") - -text {* - \noindent A problem now occurs whenever a type which - is an instance of @{class eq} in @{text HOL} is mapped - on a @{text Haskell}-built-in type which is also an instance - of @{text Haskell} @{text Eq}: -*} - -typedecl %quote bar - -instantiation %quote bar :: eq -begin - -definition %quote "eq_class.eq (x\bar) y \ x = y" - -instance %quote by default (simp add: eq_bar_def) - -end %quote - -code_type %quotett bar - (Haskell "Integer") - -text {* - \noindent The code generator would produce - an additional instance, which of course is rejected by the @{text Haskell} - compiler. - To suppress this additional instance, use - @{text "code_instance"}: -*} - -code_instance %quotett bar :: eq - (Haskell -) - - -subsection {* Enhancing the target language context \label{sec:include} *} - -text {* - In rare cases it is necessary to \emph{enrich} the context of a - target language; this is accomplished using the @{command "code_include"} - command: -*} - -code_include %quotett Haskell "Errno" -{*errno i = error ("Error number: " ++ show i)*} - -code_reserved %quotett Haskell Errno - -text {* - \noindent Such named @{text include}s are then prepended to every generated code. - Inspect such code in order to find out how @{command "code_include"} behaves - with respect to a particular target language. -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/Codegen.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Codegen.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2 +0,0 @@ - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/Further.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Further.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,113 +0,0 @@ -theory Further -imports Setup -begin - -section {* Further issues \label{sec:further} *} - -subsection {* Further reading *} - -text {* - Do dive deeper into the issue of code generation, you should visit - the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} which - contains exhaustive syntax diagrams. -*} - -subsection {* Modules *} - -text {* - When invoking the @{command export_code} command it is possible to leave - out the @{keyword "module_name"} part; then code is distributed over - different modules, where the module name space roughly is induced - by the @{text Isabelle} theory name space. - - Then sometimes the awkward situation occurs that dependencies between - definitions introduce cyclic dependencies between modules, which in the - @{text Haskell} world leaves you to the mercy of the @{text Haskell} implementation - you are using, while for @{text SML}/@{text OCaml} code generation is not possible. - - A solution is to declare module names explicitly. - Let use assume the three cyclically dependent - modules are named \emph{A}, \emph{B} and \emph{C}. - Then, by stating -*} - -code_modulename %quote SML - A ABC - B ABC - C ABC - -text {* - we explicitly map all those modules on \emph{ABC}, - resulting in an ad-hoc merge of this three modules - at serialisation time. -*} - -subsection {* Evaluation oracle *} - -text {* - Code generation may also be used to \emph{evaluate} expressions - (using @{text SML} as target language of course). - For instance, the @{command value} allows to reduce an expression to a - normal form with respect to the underlying code equations: -*} - -value %quote "42 / (12 :: rat)" - -text {* - \noindent will display @{term "7 / (2 :: rat)"}. - - The @{method eval} method tries to reduce a goal by code generation to @{term True} - and solves it in that case, but fails otherwise: -*} - -lemma %quote "42 / (12 :: rat) = 7 / 2" - by %quote eval - -text {* - \noindent The soundness of the @{method eval} method depends crucially - on the correctness of the code generator; this is one of the reasons - why you should not use adaption (see \secref{sec:adaption}) frivolously. -*} - -subsection {* Code antiquotation *} - -text {* - In scenarios involving techniques like reflection it is quite common - that code generated from a theory forms the basis for implementing - a proof procedure in @{text SML}. To facilitate interfacing of generated code - with system code, the code generator provides a @{text code} antiquotation: -*} - -datatype %quote form = T | F | And form form | Or form form - -ML %quote {* - fun eval_form @{code T} = true - | eval_form @{code F} = false - | eval_form (@{code And} (p, q)) = - eval_form p andalso eval_form q - | eval_form (@{code Or} (p, q)) = - eval_form p orelse eval_form q; -*} - -text {* - \noindent @{text code} takes as argument the name of a constant; after the - whole @{text SML} is read, the necessary code is generated transparently - and the corresponding constant names are inserted. This technique also - allows to use pattern matching on constructors stemming from compiled - @{text datatypes}. - - For a less simplistic example, theory @{theory Ferrack} is - a good reference. -*} - -subsection {* Imperative data structures *} - -text {* - If you consider imperative data structures as inevitable for a specific - application, you should consider - \emph{Imperative Functional Programming with Isabelle/HOL} - (\cite{bulwahn-et-al:2008:imperative}); - the framework described there is available in theory @{theory Imperative_HOL}. -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/Introduction.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Introduction.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,206 +0,0 @@ -theory Introduction -imports Setup -begin - -chapter {* Code generation from @{text "Isabelle/HOL"} theories *} - -section {* Introduction and Overview *} - -text {* - This tutorial introduces a generic code generator for the - @{text Isabelle} system. - Generic in the sense that the - \qn{target language} for which code shall ultimately be - generated is not fixed but may be an arbitrary state-of-the-art - functional programming language (currently, the implementation - supports @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml} and @{text Haskell} - \cite{haskell-revised-report}). - - Conceptually the code generator framework is part - of Isabelle's @{theory Pure} meta logic framework; the logic - @{theory HOL} which is an extension of @{theory Pure} - already comes with a reasonable framework setup and thus provides - a good working horse for raising code-generation-driven - applications. So, we assume some familiarity and experience - with the ingredients of the @{theory HOL} distribution theories. - (see also \cite{isa-tutorial}). - - The code generator aims to be usable with no further ado - in most cases while allowing for detailed customisation. - This manifests in the structure of this tutorial: after a short - conceptual introduction with an example (\secref{sec:intro}), - we discuss the generic customisation facilities (\secref{sec:program}). - A further section (\secref{sec:adaption}) is dedicated to the matter of - \qn{adaption} to specific target language environments. After some - further issues (\secref{sec:further}) we conclude with an overview - of some ML programming interfaces (\secref{sec:ml}). - - \begin{warn} - Ultimately, the code generator which this tutorial deals with - is supposed to replace the existing code generator - by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. - So, for the moment, there are two distinct code generators - in Isabelle. In case of ambiguity, we will refer to the framework - described here as @{text "generic code generator"}, to the - other as @{text "SML code generator"}. - Also note that while the framework itself is - object-logic independent, only @{theory HOL} provides a reasonable - framework setup. - \end{warn} - -*} - -subsection {* Code generation via shallow embedding \label{sec:intro} *} - -text {* - The key concept for understanding @{text Isabelle}'s code generation is - \emph{shallow embedding}, i.e.~logical entities like constants, types and - classes are identified with corresponding concepts in the target language. - - Inside @{theory HOL}, the @{command datatype} and - @{command definition}/@{command primrec}/@{command fun} declarations form - the core of a functional programming language. The default code generator setup - allows to turn those into functional programs immediately. - This means that \qt{naive} code generation can proceed without further ado. - For example, here a simple \qt{implementation} of amortised queues: -*} - -datatype %quote 'a queue = AQueue "'a list" "'a list" - -definition %quote empty :: "'a queue" where - "empty = AQueue [] []" - -primrec %quote enqueue :: "'a \ 'a queue \ 'a queue" where - "enqueue x (AQueue xs ys) = AQueue (x # xs) ys" - -fun %quote dequeue :: "'a queue \ 'a option \ 'a queue" where - "dequeue (AQueue [] []) = (None, AQueue [] [])" - | "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)" - | "dequeue (AQueue xs []) = - (case rev xs of y # ys \ (Some y, AQueue [] ys))" - -text {* \noindent Then we can generate code e.g.~for @{text SML} as follows: *} - -export_code %quote empty dequeue enqueue in SML - module_name Example file "examples/example.ML" - -text {* \noindent resulting in the following code: *} - -text %quote {*@{code_stmts empty enqueue dequeue (SML)}*} - -text {* - \noindent The @{command export_code} command takes a space-separated list of - constants for which code shall be generated; anything else needed for those - is added implicitly. Then follows a target language identifier - (@{text SML}, @{text OCaml} or @{text Haskell}) and a freely chosen module name. - A file name denotes the destination to store the generated code. Note that - the semantics of the destination depends on the target language: for - @{text SML} and @{text OCaml} it denotes a \emph{file}, for @{text Haskell} - it denotes a \emph{directory} where a file named as the module name - (with extension @{text ".hs"}) is written: -*} - -export_code %quote empty dequeue enqueue in Haskell - module_name Example file "examples/" - -text {* - \noindent This is how the corresponding code in @{text Haskell} looks like: -*} - -text %quote {*@{code_stmts empty enqueue dequeue (Haskell)}*} - -text {* - \noindent This demonstrates the basic usage of the @{command export_code} command; - for more details see \secref{sec:further}. -*} - -subsection {* Code generator architecture \label{sec:concept} *} - -text {* - What you have seen so far should be already enough in a lot of cases. If you - are content with this, you can quit reading here. Anyway, in order to customise - and adapt the code generator, it is inevitable to gain some understanding - how it works. - - \begin{figure}[h] - \begin{tikzpicture}[x = 4.2cm, y = 1cm] - \tikzstyle entity=[rounded corners, draw, thick, color = black, fill = white]; - \tikzstyle process=[ellipse, draw, thick, color = green, fill = white]; - \tikzstyle process_arrow=[->, semithick, color = green]; - \node (HOL) at (0, 4) [style=entity] {@{text "Isabelle/HOL"} theory}; - \node (eqn) at (2, 2) [style=entity] {code equations}; - \node (iml) at (2, 0) [style=entity] {intermediate language}; - \node (seri) at (1, 0) [style=process] {serialisation}; - \node (SML) at (0, 3) [style=entity] {@{text SML}}; - \node (OCaml) at (0, 2) [style=entity] {@{text OCaml}}; - \node (further) at (0, 1) [style=entity] {@{text "\"}}; - \node (Haskell) at (0, 0) [style=entity] {@{text Haskell}}; - \draw [style=process_arrow] (HOL) .. controls (2, 4) .. - node [style=process, near start] {selection} - node [style=process, near end] {preprocessing} - (eqn); - \draw [style=process_arrow] (eqn) -- node (transl) [style=process] {translation} (iml); - \draw [style=process_arrow] (iml) -- (seri); - \draw [style=process_arrow] (seri) -- (SML); - \draw [style=process_arrow] (seri) -- (OCaml); - \draw [style=process_arrow, dashed] (seri) -- (further); - \draw [style=process_arrow] (seri) -- (Haskell); - \end{tikzpicture} - \caption{Code generator architecture} - \label{fig:arch} - \end{figure} - - The code generator employs a notion of executability - for three foundational executable ingredients known - from functional programming: - \emph{code equations}, \emph{datatypes}, and - \emph{type classes}. A code equation as a first approximation - is a theorem of the form @{text "f t\<^isub>1 t\<^isub>2 \ t\<^isub>n \ t"} - (an equation headed by a constant @{text f} with arguments - @{text "t\<^isub>1 t\<^isub>2 \ t\<^isub>n"} and right hand side @{text t}). - Code generation aims to turn code equations - into a functional program. This is achieved by three major - components which operate sequentially, i.e. the result of one is - the input - of the next in the chain, see diagram \ref{fig:arch}: - - \begin{itemize} - - \item Out of the vast collection of theorems proven in a - \qn{theory}, a reasonable subset modelling - code equations is \qn{selected}. - - \item On those selected theorems, certain - transformations are carried out - (\qn{preprocessing}). Their purpose is to turn theorems - representing non- or badly executable - specifications into equivalent but executable counterparts. - The result is a structured collection of \qn{code theorems}. - - \item Before the selected code equations are continued with, - they can be \qn{preprocessed}, i.e. subjected to theorem - transformations. This \qn{preprocessor} is an interface which - allows to apply - the full expressiveness of ML-based theorem transformations - to code generation; motivating examples are shown below, see - \secref{sec:preproc}. - The result of the preprocessing step is a structured collection - of code equations. - - \item These code equations are \qn{translated} to a program - in an abstract intermediate language. Think of it as a kind - of \qt{Mini-Haskell} with four \qn{statements}: @{text data} - (for datatypes), @{text fun} (stemming from code equations), - also @{text class} and @{text inst} (for type classes). - - \item Finally, the abstract program is \qn{serialised} into concrete - source code of a target language. - - \end{itemize} - - \noindent From these steps, only the two last are carried out outside the logic; by - keeping this layer as thin as possible, the amount of code to trust is - kept to a minimum. -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/ML.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/ML.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,177 +0,0 @@ -theory "ML" -imports Setup -begin - -section {* ML system interfaces \label{sec:ml} *} - -text {* - Since the code generator framework not only aims to provide - a nice Isar interface but also to form a base for - code-generation-based applications, here a short - description of the most important ML interfaces. -*} - -subsection {* Executable theory content: @{text Code} *} - -text {* - This Pure module implements the core notions of - executable content of a theory. -*} - -subsubsection {* Managing executable content *} - -text %mlref {* - \begin{mldecls} - @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\ - @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\ - @{index_ML Code.add_eqnl: "string * (thm * bool) list lazy -> theory -> theory"} \\ - @{index_ML Code.map_pre: "(simpset -> simpset) -> theory -> theory"} \\ - @{index_ML Code.map_post: "(simpset -> simpset) -> theory -> theory"} \\ - @{index_ML Code.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option) - -> theory -> theory"} \\ - @{index_ML Code.del_functrans: "string -> theory -> theory"} \\ - @{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\ - @{index_ML Code.get_datatype: "theory -> string - -> (string * sort) list * (string * typ list) list"} \\ - @{index_ML Code.get_datatype_of_constr: "theory -> string -> string option"} - \end{mldecls} - - \begin{description} - - \item @{ML Code.add_eqn}~@{text "thm"}~@{text "thy"} adds function - theorem @{text "thm"} to executable content. - - \item @{ML Code.del_eqn}~@{text "thm"}~@{text "thy"} removes function - theorem @{text "thm"} from executable content, if present. - - \item @{ML Code.add_eqnl}~@{text "(const, lthms)"}~@{text "thy"} adds - suspended code equations @{text lthms} for constant - @{text const} to executable content. - - \item @{ML Code.map_pre}~@{text "f"}~@{text "thy"} changes - the preprocessor simpset. - - \item @{ML Code.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds - function transformer @{text f} (named @{text name}) to executable content; - @{text f} is a transformer of the code equations belonging - to a certain function definition, depending on the - current theory context. Returning @{text NONE} indicates that no - transformation took place; otherwise, the whole process will be iterated - with the new code equations. - - \item @{ML Code.del_functrans}~@{text "name"}~@{text "thy"} removes - function transformer named @{text name} from executable content. - - \item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds - a datatype to executable content, with generation - set @{text cs}. - - \item @{ML Code.get_datatype_of_constr}~@{text "thy"}~@{text "const"} - returns type constructor corresponding to - constructor @{text const}; returns @{text NONE} - if @{text const} is no constructor. - - \end{description} -*} - -subsection {* Auxiliary *} - -text %mlref {* - \begin{mldecls} - @{index_ML Code_Unit.read_const: "theory -> string -> string"} \\ - @{index_ML Code_Unit.head_eqn: "theory -> thm -> string * ((string * sort) list * typ)"} \\ - @{index_ML Code_Unit.rewrite_eqn: "simpset -> thm -> thm"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML Code_Unit.read_const}~@{text thy}~@{text s} - reads a constant as a concrete term expression @{text s}. - - \item @{ML Code_Unit.head_eqn}~@{text thy}~@{text thm} - extracts the constant and its type from a code equation @{text thm}. - - \item @{ML Code_Unit.rewrite_eqn}~@{text ss}~@{text thm} - rewrites a code equation @{text thm} with a simpset @{text ss}; - only arguments and right hand side are rewritten, - not the head of the code equation. - - \end{description} - -*} - -subsection {* Implementing code generator applications *} - -text {* - Implementing code generator applications on top - of the framework set out so far usually not only - involves using those primitive interfaces - but also storing code-dependent data and various - other things. -*} - -subsubsection {* Data depending on the theory's executable content *} - -text {* - Due to incrementality of code generation, changes in the - theory's executable content have to be propagated in a - certain fashion. Additionally, such changes may occur - not only during theory extension but also during theory - merge, which is a little bit nasty from an implementation - point of view. The framework provides a solution - to this technical challenge by providing a functorial - data slot @{ML_functor CodeDataFun}; on instantiation - of this functor, the following types and operations - are required: - - \medskip - \begin{tabular}{l} - @{text "type T"} \\ - @{text "val empty: T"} \\ - @{text "val purge: theory \ string list option \ T \ T"} - \end{tabular} - - \begin{description} - - \item @{text T} the type of data to store. - - \item @{text empty} initial (empty) data. - - \item @{text purge}~@{text thy}~@{text consts} propagates changes in executable content; - @{text consts} indicates the kind - of change: @{ML NONE} stands for a fundamental change - which invalidates any existing code, @{text "SOME consts"} - hints that executable content for constants @{text consts} - has changed. - - \end{description} - - \noindent An instance of @{ML_functor CodeDataFun} provides the following - interface: - - \medskip - \begin{tabular}{l} - @{text "get: theory \ T"} \\ - @{text "change: theory \ (T \ T) \ T"} \\ - @{text "change_yield: theory \ (T \ 'a * T) \ 'a * T"} - \end{tabular} - - \begin{description} - - \item @{text get} retrieval of the current data. - - \item @{text change} update of current data (cached!) - by giving a continuation. - - \item @{text change_yield} update with side result. - - \end{description} -*} - -text {* - \bigskip - - \emph{Happy proving, happy hacking!} -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/Program.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Program.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,534 +0,0 @@ -theory Program -imports Introduction -begin - -section {* Turning Theories into Programs \label{sec:program} *} - -subsection {* The @{text "Isabelle/HOL"} default setup *} - -text {* - We have already seen how by default equations stemming from - @{command definition}/@{command primrec}/@{command fun} - statements are used for code generation. This default behaviour - can be changed, e.g. by providing different code equations. - All kinds of customisation shown in this section is \emph{safe} - in the sense that the user does not have to worry about - correctness -- all programs generatable that way are partially - correct. -*} - -subsection {* Selecting code equations *} - -text {* - Coming back to our introductory example, we - could provide an alternative code equations for @{const dequeue} - explicitly: -*} - -lemma %quote [code]: - "dequeue (AQueue xs []) = - (if xs = [] then (None, AQueue [] []) - else dequeue (AQueue [] (rev xs)))" - "dequeue (AQueue xs (y # ys)) = - (Some y, AQueue xs ys)" - by (cases xs, simp_all) (cases "rev xs", simp_all) - -text {* - \noindent The annotation @{text "[code]"} is an @{text Isar} - @{text attribute} which states that the given theorems should be - considered as code equations for a @{text fun} statement -- - the corresponding constant is determined syntactically. The resulting code: -*} - -text %quote {*@{code_stmts dequeue (consts) dequeue (Haskell)}*} - -text {* - \noindent You may note that the equality test @{term "xs = []"} has been - replaced by the predicate @{term "null xs"}. This is due to the default - setup in the \qn{preprocessor} to be discussed further below (\secref{sec:preproc}). - - Changing the default constructor set of datatypes is also - possible. See \secref{sec:datatypes} for an example. - - As told in \secref{sec:concept}, code generation is based - on a structured collection of code theorems. - For explorative purpose, this collection - may be inspected using the @{command code_thms} command: -*} - -code_thms %quote dequeue - -text {* - \noindent prints a table with \emph{all} code equations - for @{const dequeue}, including - \emph{all} code equations those equations depend - on recursively. - - Similarly, the @{command code_deps} command shows a graph - visualising dependencies between code equations. -*} - -subsection {* @{text class} and @{text instantiation} *} - -text {* - Concerning type classes and code generation, let us examine an example - from abstract algebra: -*} - -class %quote semigroup = - fixes mult :: "'a \ 'a \ 'a" (infixl "\" 70) - assumes assoc: "(x \ y) \ z = x \ (y \ z)" - -class %quote monoid = semigroup + - fixes neutral :: 'a ("\") - assumes neutl: "\ \ x = x" - and neutr: "x \ \ = x" - -instantiation %quote nat :: monoid -begin - -primrec %quote mult_nat where - "0 \ n = (0\nat)" - | "Suc m \ n = n + m \ n" - -definition %quote neutral_nat where - "\ = Suc 0" - -lemma %quote add_mult_distrib: - fixes n m q :: nat - shows "(n + m) \ q = n \ q + m \ q" - by (induct n) simp_all - -instance %quote proof - fix m n q :: nat - show "m \ n \ q = m \ (n \ q)" - by (induct m) (simp_all add: add_mult_distrib) - show "\ \ n = n" - by (simp add: neutral_nat_def) - show "m \ \ = m" - by (induct m) (simp_all add: neutral_nat_def) -qed - -end %quote - -text {* - \noindent We define the natural operation of the natural numbers - on monoids: -*} - -primrec %quote (in monoid) pow :: "nat \ 'a \ 'a" where - "pow 0 a = \" - | "pow (Suc n) a = a \ pow n a" - -text {* - \noindent This we use to define the discrete exponentiation function: -*} - -definition %quote bexp :: "nat \ nat" where - "bexp n = pow n (Suc (Suc 0))" - -text {* - \noindent The corresponding code: -*} - -text %quote {*@{code_stmts bexp (Haskell)}*} - -text {* - \noindent This is a convenient place to show how explicit dictionary construction - manifests in generated code (here, the same example in @{text SML}): -*} - -text %quote {*@{code_stmts bexp (SML)}*} - -text {* - \noindent Note the parameters with trailing underscore (@{verbatim "A_"}) - which are the dictionary parameters. -*} - -subsection {* The preprocessor \label{sec:preproc} *} - -text {* - Before selected function theorems are turned into abstract - code, a chain of definitional transformation steps is carried - out: \emph{preprocessing}. In essence, the preprocessor - consists of two components: a \emph{simpset} and \emph{function transformers}. - - The \emph{simpset} allows to employ the full generality of the Isabelle - simplifier. Due to the interpretation of theorems - as code equations, rewrites are applied to the right - hand side and the arguments of the left hand side of an - equation, but never to the constant heading the left hand side. - An important special case are \emph{inline theorems} which may be - declared and undeclared using the - \emph{code inline} or \emph{code inline del} attribute respectively. - - Some common applications: -*} - -text_raw {* - \begin{itemize} -*} - -text {* - \item replacing non-executable constructs by executable ones: -*} - -lemma %quote [code inline]: - "x \ set xs \ x mem xs" by (induct xs) simp_all - -text {* - \item eliminating superfluous constants: -*} - -lemma %quote [code inline]: - "1 = Suc 0" by simp - -text {* - \item replacing executable but inconvenient constructs: -*} - -lemma %quote [code inline]: - "xs = [] \ List.null xs" by (induct xs) simp_all - -text_raw {* - \end{itemize} -*} - -text {* - \noindent \emph{Function transformers} provide a very general interface, - transforming a list of function theorems to another - list of function theorems, provided that neither the heading - constant nor its type change. The @{term "0\nat"} / @{const Suc} - pattern elimination implemented in - theory @{text Efficient_Nat} (see \secref{eff_nat}) uses this - interface. - - \noindent The current setup of the preprocessor may be inspected using - the @{command print_codesetup} command. - @{command code_thms} provides a convenient - mechanism to inspect the impact of a preprocessor setup - on code equations. - - \begin{warn} - The attribute \emph{code unfold} - associated with the @{text "SML code generator"} also applies to - the @{text "generic code generator"}: - \emph{code unfold} implies \emph{code inline}. - \end{warn} -*} - -subsection {* Datatypes \label{sec:datatypes} *} - -text {* - Conceptually, any datatype is spanned by a set of - \emph{constructors} of type @{text "\ = \ \ \ \\<^isub>1 \ \\<^isub>n"} where @{text - "{\\<^isub>1, \, \\<^isub>n}"} is exactly the set of \emph{all} type variables in - @{text "\"}. The HOL datatype package by default registers any new - datatype in the table of datatypes, which may be inspected using the - @{command print_codesetup} command. - - In some cases, it is appropriate to alter or extend this table. As - an example, we will develop an alternative representation of the - queue example given in \secref{sec:intro}. The amortised - representation is convenient for generating code but exposes its - \qt{implementation} details, which may be cumbersome when proving - theorems about it. Therefore, here a simple, straightforward - representation of queues: -*} - -datatype %quote 'a queue = Queue "'a list" - -definition %quote empty :: "'a queue" where - "empty = Queue []" - -primrec %quote enqueue :: "'a \ 'a queue \ 'a queue" where - "enqueue x (Queue xs) = Queue (xs @ [x])" - -fun %quote dequeue :: "'a queue \ 'a option \ 'a queue" where - "dequeue (Queue []) = (None, Queue [])" - | "dequeue (Queue (x # xs)) = (Some x, Queue xs)" - -text {* - \noindent This we can use directly for proving; for executing, - we provide an alternative characterisation: -*} - -definition %quote AQueue :: "'a list \ 'a list \ 'a queue" where - "AQueue xs ys = Queue (ys @ rev xs)" - -code_datatype %quote AQueue - -text {* - \noindent Here we define a \qt{constructor} @{const "AQueue"} which - is defined in terms of @{text "Queue"} and interprets its arguments - according to what the \emph{content} of an amortised queue is supposed - to be. Equipped with this, we are able to prove the following equations - for our primitive queue operations which \qt{implement} the simple - queues in an amortised fashion: -*} - -lemma %quote empty_AQueue [code]: - "empty = AQueue [] []" - unfolding AQueue_def empty_def by simp - -lemma %quote enqueue_AQueue [code]: - "enqueue x (AQueue xs ys) = AQueue (x # xs) ys" - unfolding AQueue_def by simp - -lemma %quote dequeue_AQueue [code]: - "dequeue (AQueue xs []) = - (if xs = [] then (None, AQueue [] []) - else dequeue (AQueue [] (rev xs)))" - "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)" - unfolding AQueue_def by simp_all - -text {* - \noindent For completeness, we provide a substitute for the - @{text case} combinator on queues: -*} - -definition %quote - aqueue_case_def: "aqueue_case = queue_case" - -lemma %quote aqueue_case [code, code inline]: - "queue_case = aqueue_case" - unfolding aqueue_case_def .. - -lemma %quote case_AQueue [code]: - "aqueue_case f (AQueue xs ys) = f (ys @ rev xs)" - unfolding aqueue_case_def AQueue_def by simp - -text {* - \noindent The resulting code looks as expected: -*} - -text %quote {*@{code_stmts empty enqueue dequeue (SML)}*} - -text {* - \noindent From this example, it can be glimpsed that using own - constructor sets is a little delicate since it changes the set of - valid patterns for values of that type. Without going into much - detail, here some practical hints: - - \begin{itemize} - - \item When changing the constructor set for datatypes, take care - to provide an alternative for the @{text case} combinator - (e.g.~by replacing it using the preprocessor). - - \item Values in the target language need not to be normalised -- - different values in the target language may represent the same - value in the logic. - - \item Usually, a good methodology to deal with the subtleties of - pattern matching is to see the type as an abstract type: provide - a set of operations which operate on the concrete representation - of the type, and derive further operations by combinations of - these primitive ones, without relying on a particular - representation. - - \end{itemize} -*} - - -subsection {* Equality and wellsortedness *} - -text {* - Surely you have already noticed how equality is treated - by the code generator: -*} - -primrec %quote collect_duplicates :: "'a list \ 'a list \ 'a list \ 'a list" where - "collect_duplicates xs ys [] = xs" - | "collect_duplicates xs ys (z#zs) = (if z \ set xs - then if z \ set ys - then collect_duplicates xs ys zs - else collect_duplicates xs (z#ys) zs - else collect_duplicates (z#xs) (z#ys) zs)" - -text {* - \noindent The membership test during preprocessing is rewritten, - resulting in @{const List.member}, which itself - performs an explicit equality check. -*} - -text %quote {*@{code_stmts collect_duplicates (SML)}*} - -text {* - \noindent Obviously, polymorphic equality is implemented the Haskell - way using a type class. How is this achieved? HOL introduces - an explicit class @{class eq} with a corresponding operation - @{const eq_class.eq} such that @{thm eq [no_vars]}. - The preprocessing framework does the rest by propagating the - @{class eq} constraints through all dependent code equations. - For datatypes, instances of @{class eq} are implicitly derived - when possible. For other types, you may instantiate @{text eq} - manually like any other type class. - - Though this @{text eq} class is designed to get rarely in - the way, a subtlety - enters the stage when definitions of overloaded constants - are dependent on operational equality. For example, let - us define a lexicographic ordering on tuples - (also see theory @{theory Product_ord}): -*} - -instantiation %quote "*" :: (order, order) order -begin - -definition %quote [code del]: - "x \ y \ fst x < fst y \ fst x = fst y \ snd x \ snd y" - -definition %quote [code del]: - "x < y \ fst x < fst y \ fst x = fst y \ snd x < snd y" - -instance %quote proof -qed (auto simp: less_eq_prod_def less_prod_def intro: order_less_trans) - -end %quote - -lemma %quote order_prod [code]: - "(x1 \ 'a\order, y1 \ 'b\order) < (x2, y2) \ - x1 < x2 \ x1 = x2 \ y1 < y2" - "(x1 \ 'a\order, y1 \ 'b\order) \ (x2, y2) \ - x1 < x2 \ x1 = x2 \ y1 \ y2" - by (simp_all add: less_prod_def less_eq_prod_def) - -text {* - \noindent Then code generation will fail. Why? The definition - of @{term "op \"} depends on equality on both arguments, - which are polymorphic and impose an additional @{class eq} - class constraint, which the preprocessor does not propagate - (for technical reasons). - - The solution is to add @{class eq} explicitly to the first sort arguments in the - code theorems: -*} - -lemma %quote order_prod_code [code]: - "(x1 \ 'a\{order, eq}, y1 \ 'b\order) < (x2, y2) \ - x1 < x2 \ x1 = x2 \ y1 < y2" - "(x1 \ 'a\{order, eq}, y1 \ 'b\order) \ (x2, y2) \ - x1 < x2 \ x1 = x2 \ y1 \ y2" - by (simp_all add: less_prod_def less_eq_prod_def) - -text {* - \noindent Then code generation succeeds: -*} - -text %quote {*@{code_stmts "op \ \ _ \ _ \ _ \ _ \ bool" (SML)}*} - -text {* - In some cases, the automatically derived code equations - for equality on a particular type may not be appropriate. - As example, watch the following datatype representing - monomorphic parametric types (where type constructors - are referred to by natural numbers): -*} - -datatype %quote monotype = Mono nat "monotype list" -(*<*) -lemma monotype_eq: - "eq_class.eq (Mono tyco1 typargs1) (Mono tyco2 typargs2) \ - eq_class.eq tyco1 tyco2 \ eq_class.eq typargs1 typargs2" by (simp add: eq) -(*>*) - -text {* - \noindent Then code generation for SML would fail with a message - that the generated code contains illegal mutual dependencies: - the theorem @{thm monotype_eq [no_vars]} already requires the - instance @{text "monotype \ eq"}, which itself requires - @{thm monotype_eq [no_vars]}; Haskell has no problem with mutually - recursive @{text instance} and @{text function} definitions, - but the SML serialiser does not support this. - - In such cases, you have to provide your own equality equations - involving auxiliary constants. In our case, - @{const [show_types] list_all2} can do the job: -*} - -lemma %quote monotype_eq_list_all2 [code]: - "eq_class.eq (Mono tyco1 typargs1) (Mono tyco2 typargs2) \ - eq_class.eq tyco1 tyco2 \ list_all2 eq_class.eq typargs1 typargs2" - by (simp add: eq list_all2_eq [symmetric]) - -text {* - \noindent does not depend on instance @{text "monotype \ eq"}: -*} - -text %quote {*@{code_stmts "eq_class.eq :: monotype \ monotype \ bool" (SML)}*} - - -subsection {* Explicit partiality *} - -text {* - Partiality usually enters the game by partial patterns, as - in the following example, again for amortised queues: -*} - -definition %quote strict_dequeue :: "'a queue \ 'a \ 'a queue" where - "strict_dequeue q = (case dequeue q - of (Some x, q') \ (x, q'))" - -lemma %quote strict_dequeue_AQueue [code]: - "strict_dequeue (AQueue xs (y # ys)) = (y, AQueue xs ys)" - "strict_dequeue (AQueue xs []) = - (case rev xs of y # ys \ (y, AQueue [] ys))" - by (simp_all add: strict_dequeue_def dequeue_AQueue split: list.splits) - -text {* - \noindent In the corresponding code, there is no equation - for the pattern @{term "AQueue [] []"}: -*} - -text %quote {*@{code_stmts strict_dequeue (consts) strict_dequeue (Haskell)}*} - -text {* - \noindent In some cases it is desirable to have this - pseudo-\qt{partiality} more explicitly, e.g.~as follows: -*} - -axiomatization %quote empty_queue :: 'a - -definition %quote strict_dequeue' :: "'a queue \ 'a \ 'a queue" where - "strict_dequeue' q = (case dequeue q of (Some x, q') \ (x, q') | _ \ empty_queue)" - -lemma %quote strict_dequeue'_AQueue [code]: - "strict_dequeue' (AQueue xs []) = (if xs = [] then empty_queue - else strict_dequeue' (AQueue [] (rev xs)))" - "strict_dequeue' (AQueue xs (y # ys)) = - (y, AQueue xs ys)" - by (simp_all add: strict_dequeue'_def dequeue_AQueue split: list.splits) - -text {* - Observe that on the right hand side of the definition of @{const - "strict_dequeue'"} the constant @{const empty_queue} occurs - which is unspecified. - - Normally, if constants without any code equations occur in a - program, the code generator complains (since in most cases this is - not what the user expects). But such constants can also be thought - of as function definitions with no equations which always fail, - since there is never a successful pattern match on the left hand - side. In order to categorise a constant into that category - explicitly, use @{command "code_abort"}: -*} - -code_abort %quote empty_queue - -text {* - \noindent Then the code generator will just insert an error or - exception at the appropriate position: -*} - -text %quote {*@{code_stmts strict_dequeue' (consts) empty_queue strict_dequeue' (Haskell)}*} - -text {* - \noindent This feature however is rarely needed in practice. - Note also that the @{text HOL} default setup already declares - @{const undefined} as @{command "code_abort"}, which is most - likely to be used in such situations. -*} - -end - \ No newline at end of file diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/ROOT.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/ROOT.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,11 +0,0 @@ - -(* $Id$ *) - -no_document use_thy "Setup"; -no_document use_thys ["Efficient_Nat"]; - -use_thy "Introduction"; -use_thy "Program"; -use_thy "Adaption"; -use_thy "Further"; -use_thy "ML"; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/Setup.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Setup.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,12 +0,0 @@ -theory Setup -imports Complex_Main -uses "../../../antiquote_setup.ML" "../../../more_antiquote.ML" -begin - -ML {* no_document use_thys - ["Efficient_Nat", "Code_Char_chr", "Product_ord", "~~/src/HOL/Imperative_HOL/Imperative_HOL", - "~~/src/HOL/Reflection/Ferrack"] *} - -ML_val {* Code_Target.code_width := 74 *} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/document/Adaption.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/Adaption.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,679 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Adaption}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Adaption\isanewline -\isakeyword{imports}\ Setup\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -\isanewline -% -\endisadelimtheory -% -\isadeliminvisible -\isanewline -% -\endisadeliminvisible -% -\isataginvisible -\isacommand{setup}\isamarkupfalse% -\ {\isacharverbatimopen}\ Code{\isacharunderscore}Target{\isachardot}extend{\isacharunderscore}target\ {\isacharparenleft}{\isachardoublequote}{\isasymSML}{\isachardoublequote}{\isacharcomma}\ {\isacharparenleft}{\isachardoublequote}SML{\isachardoublequote}{\isacharcomma}\ K\ I{\isacharparenright}{\isacharparenright}\ {\isacharverbatimclose}% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isamarkupsection{Adaption to target languages \label{sec:adaption}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Adapting code generation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The aspects of code generation introduced so far have two aspects - in common: - - \begin{itemize} - \item They act uniformly, without reference to a specific - target language. - \item They are \emph{safe} in the sense that as long as you trust - the code generator meta theory and implementation, you cannot - produce programs that yield results which are not derivable - in the logic. - \end{itemize} - - \noindent In this section we will introduce means to \emph{adapt} the serialiser - to a specific target language, i.e.~to print program fragments - in a way which accommodates \qt{already existing} ingredients of - a target language environment, for three reasons: - - \begin{itemize} - \item improving readability and aesthetics of generated code - \item gaining efficiency - \item interface with language parts which have no direct counterpart - in \isa{HOL} (say, imperative data structures) - \end{itemize} - - \noindent Generally, you should avoid using those features yourself - \emph{at any cost}: - - \begin{itemize} - \item The safe configuration methods act uniformly on every target language, - whereas for adaption you have to treat each target language separate. - \item Application is extremely tedious since there is no abstraction - which would allow for a static check, making it easy to produce garbage. - \item More or less subtle errors can be introduced unconsciously. - \end{itemize} - - \noindent However, even if you ought refrain from setting up adaption - yourself, already the \isa{HOL} comes with some reasonable default - adaptions (say, using target language list syntax). There also some - common adaption cases which you can setup by importing particular - library theories. In order to understand these, we provide some clues here; - these however are not supposed to replace a careful study of the sources.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{The adaption principle% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The following figure illustrates what \qt{adaption} is conceptually - supposed to be: - - \begin{figure}[here] - \begin{tikzpicture}[scale = 0.5] - \tikzstyle water=[color = blue, thick] - \tikzstyle ice=[color = black, very thick, cap = round, join = round, fill = white] - \tikzstyle process=[color = green, semithick, ->] - \tikzstyle adaption=[color = red, semithick, ->] - \tikzstyle target=[color = black] - \foreach \x in {0, ..., 24} - \draw[style=water] (\x, 0.25) sin + (0.25, 0.25) cos + (0.25, -0.25) sin - + (0.25, -0.25) cos + (0.25, 0.25); - \draw[style=ice] (1, 0) -- - (3, 6) node[above, fill=white] {logic} -- (5, 0) -- cycle; - \draw[style=ice] (9, 0) -- - (11, 6) node[above, fill=white] {intermediate language} -- (13, 0) -- cycle; - \draw[style=ice] (15, -6) -- - (19, 6) node[above, fill=white] {target language} -- (23, -6) -- cycle; - \draw[style=process] - (3.5, 3) .. controls (7, 5) .. node[fill=white] {translation} (10.5, 3); - \draw[style=process] - (11.5, 3) .. controls (15, 5) .. node[fill=white] (serialisation) {serialisation} (18.5, 3); - \node (adaption) at (11, -2) [style=adaption] {adaption}; - \node at (19, 3) [rotate=90] {generated}; - \node at (19.5, -5) {language}; - \node at (19.5, -3) {library}; - \node (includes) at (19.5, -1) {includes}; - \node (reserved) at (16.5, -3) [rotate=72] {reserved}; % proper 71.57 - \draw[style=process] - (includes) -- (serialisation); - \draw[style=process] - (reserved) -- (serialisation); - \draw[style=adaption] - (adaption) -- (serialisation); - \draw[style=adaption] - (adaption) -- (includes); - \draw[style=adaption] - (adaption) -- (reserved); - \end{tikzpicture} - \caption{The adaption principle} - \label{fig:adaption} - \end{figure} - - \noindent In the tame view, code generation acts as broker between - \isa{logic}, \isa{intermediate\ language} and - \isa{target\ language} by means of \isa{translation} and - \isa{serialisation}; for the latter, the serialiser has to observe - the structure of the \isa{language} itself plus some \isa{reserved} - keywords which have to be avoided for generated code. - However, if you consider \isa{adaption} mechanisms, the code generated - by the serializer is just the tip of the iceberg: - - \begin{itemize} - \item \isa{serialisation} can be \emph{parametrised} such that - logical entities are mapped to target-specific ones - (e.g. target-specific list syntax, - see also \secref{sec:adaption_mechanisms}) - \item Such parametrisations can involve references to a - target-specific standard \isa{library} (e.g. using - the \isa{Haskell} \verb|Maybe| type instead - of the \isa{HOL} \isa{option} type); - if such are used, the corresponding identifiers - (in our example, \verb|Maybe|, \verb|Nothing| - and \verb|Just|) also have to be considered \isa{reserved}. - \item Even more, the user can enrich the library of the - target-language by providing code snippets - (\qt{\isa{includes}}) which are prepended to - any generated code (see \secref{sec:include}); this typically - also involves further \isa{reserved} identifiers. - \end{itemize} - - \noindent As figure \ref{fig:adaption} illustrates, all these adaption mechanisms - have to act consistently; it is at the discretion of the user - to take care for this.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Common adaption patterns% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} \hyperlink{theory.Main}{\mbox{\isa{Main}}} theory already provides a code - generator setup - which should be suitable for most applications. Common extensions - and modifications are available by certain theories of the \isa{HOL} - library; beside being useful in applications, they may serve - as a tutorial for customising the code generator setup (see below - \secref{sec:adaption_mechanisms}). - - \begin{description} - - \item[\hyperlink{theory.Code-Integer}{\mbox{\isa{Code{\isacharunderscore}Integer}}}] represents \isa{HOL} integers by big - integer literals in target languages. - \item[\hyperlink{theory.Code-Char}{\mbox{\isa{Code{\isacharunderscore}Char}}}] represents \isa{HOL} characters by - character literals in target languages. - \item[\hyperlink{theory.Code-Char-chr}{\mbox{\isa{Code{\isacharunderscore}Char{\isacharunderscore}chr}}}] like \isa{Code{\isacharunderscore}Char}, - but also offers treatment of character codes; includes - \hyperlink{theory.Code-Char}{\mbox{\isa{Code{\isacharunderscore}Char}}}. - \item[\hyperlink{theory.Efficient-Nat}{\mbox{\isa{Efficient{\isacharunderscore}Nat}}}] \label{eff_nat} implements natural numbers by integers, - which in general will result in higher efficiency; pattern - matching with \isa{{\isadigit{0}}} / \isa{Suc} - is eliminated; includes \hyperlink{theory.Code-Integer}{\mbox{\isa{Code{\isacharunderscore}Integer}}} - and \hyperlink{theory.Code-Index}{\mbox{\isa{Code{\isacharunderscore}Index}}}. - \item[\hyperlink{theory.Code-Index}{\mbox{\isa{Code{\isacharunderscore}Index}}}] provides an additional datatype - \isa{index} which is mapped to target-language built-in integers. - Useful for code setups which involve e.g. indexing of - target-language arrays. - \item[\hyperlink{theory.Code-Message}{\mbox{\isa{Code{\isacharunderscore}Message}}}] provides an additional datatype - \isa{message{\isacharunderscore}string} which is isomorphic to strings; - \isa{message{\isacharunderscore}string}s are mapped to target-language strings. - Useful for code setups which involve e.g. printing (error) messages. - - \end{description} - - \begin{warn} - When importing any of these theories, they should form the last - items in an import list. Since these theories adapt the - code generator setup in a non-conservative fashion, - strange effects may occur otherwise. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Parametrising serialisation \label{sec:adaption_mechanisms}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Consider the following function and its corresponding - SML code:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{primrec}\isamarkupfalse% -\ in{\isacharunderscore}interval\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymtimes}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}in{\isacharunderscore}interval\ {\isacharparenleft}k{\isacharcomma}\ l{\isacharparenright}\ n\ {\isasymlongleftrightarrow}\ k\ {\isasymle}\ n\ {\isasymand}\ n\ {\isasymle}\ l{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isataginvisible -% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype nat = Suc of nat | Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype boola = False | True;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun anda x True = x\\ -\hspace*{0pt} ~| anda x False = False\\ -\hspace*{0pt} ~| anda True x = x\\ -\hspace*{0pt} ~| anda False x = False;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\ -\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = False\\ -\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\ -\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = True;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun in{\char95}interval (k,~l) n = anda (less{\char95}eq{\char95}nat k n) (less{\char95}eq{\char95}nat n l);\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Though this is correct code, it is a little bit unsatisfactory: - boolean values and operators are materialised as distinguished - entities with have nothing to do with the SML-built-in notion - of \qt{bool}. This results in less readable code; - additionally, eager evaluation may cause programs to - loop or break which would perfectly terminate when - the existing SML \verb|bool| would be used. To map - the HOL \isa{bool} on SML \verb|bool|, we may use - \qn{custom serialisations}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquotett -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ bool\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}bool{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ True\ \isakeyword{and}\ False\ \isakeyword{and}\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}true{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}false{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}{\isacharunderscore}\ andalso\ {\isacharunderscore}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\begin{isamarkuptext}% -\noindent The \hyperlink{command.code-type}{\mbox{\isa{\isacommand{code{\isacharunderscore}type}}}} command takes a type constructor - as arguments together with a list of custom serialisations. - Each custom serialisation starts with a target language - identifier followed by an expression, which during - code serialisation is inserted whenever the type constructor - would occur. For constants, \hyperlink{command.code-const}{\mbox{\isa{\isacommand{code{\isacharunderscore}const}}}} implements - the corresponding mechanism. Each ``\verb|_|'' in - a serialisation expression is treated as a placeholder - for the type constructor's (the constant's) arguments.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype nat = Suc of nat | Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\ -\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = false\\ -\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\ -\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = true;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun in{\char95}interval (k,~l) n = (less{\char95}eq{\char95}nat k n) andalso (less{\char95}eq{\char95}nat n l);\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This still is not perfect: the parentheses - around the \qt{andalso} expression are superfluous. - Though the serialiser - by no means attempts to imitate the rich Isabelle syntax - framework, it provides some common idioms, notably - associative infixes with precedences which may be used here:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquotett -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}SML\ \isakeyword{infixl}\ {\isadigit{1}}\ {\isachardoublequoteopen}andalso{\isachardoublequoteclose}{\isacharparenright}% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype nat = Suc of nat | Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\ -\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = false\\ -\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\ -\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = true;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun in{\char95}interval (k,~l) n = less{\char95}eq{\char95}nat k n andalso less{\char95}eq{\char95}nat n l;\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The attentive reader may ask how we assert that no generated - code will accidentally overwrite. For this reason the serialiser has - an internal table of identifiers which have to be avoided to be used - for new declarations. Initially, this table typically contains the - keywords of the target language. It can be extended manually, thus avoiding - accidental overwrites, using the \hyperlink{command.code-reserved}{\mbox{\isa{\isacommand{code{\isacharunderscore}reserved}}}} command:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{code{\isacharunderscore}reserved}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymSML}{\isachardoublequoteclose}\ bool\ true\ false\ andalso% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Next, we try to map HOL pairs to SML pairs, using the - infix ``\verb|*|'' type constructor and parentheses:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isataginvisible -% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isadelimquotett -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ {\isacharasterisk}\isanewline -\ \ {\isacharparenleft}SML\ \isakeyword{infix}\ {\isadigit{2}}\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ Pair\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}{\isacharbang}{\isacharparenleft}{\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharcomma}{\isacharslash}\ {\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\begin{isamarkuptext}% -\noindent The initial bang ``\verb|!|'' tells the serialiser - never to put - parentheses around the whole expression (they are already present), - while the parentheses around argument place holders - tell not to put parentheses around the arguments. - The slash ``\verb|/|'' (followed by arbitrary white space) - inserts a space which may be used as a break if necessary - during pretty printing. - - These examples give a glimpse what mechanisms - custom serialisations provide; however their usage - requires careful thinking in order not to introduce - inconsistencies -- or, in other words: - custom serialisations are completely axiomatic. - - A further noteworthy details is that any special - character in a custom serialisation may be quoted - using ``\verb|'|''; thus, in - ``\verb|fn '_ => _|'' the first - ``\verb|_|'' is a proper underscore while the - second ``\verb|_|'' is a placeholder.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{\isa{Haskell} serialisation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -For convenience, the default - \isa{HOL} setup for \isa{Haskell} maps the \isa{eq} class to - its counterpart in \isa{Haskell}, giving custom serialisations - for the class \isa{eq} (by command \hyperlink{command.code-class}{\mbox{\isa{\isacommand{code{\isacharunderscore}class}}}}) and its operation - \isa{eq{\isacharunderscore}class{\isachardot}eq}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquotett -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}class}\isamarkupfalse% -\ eq\isanewline -\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Eq{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}Haskell\ \isakeyword{infixl}\ {\isadigit{4}}\ {\isachardoublequoteopen}{\isacharequal}{\isacharequal}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\begin{isamarkuptext}% -\noindent A problem now occurs whenever a type which - is an instance of \isa{eq} in \isa{HOL} is mapped - on a \isa{Haskell}-built-in type which is also an instance - of \isa{Haskell} \isa{Eq}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{typedecl}\isamarkupfalse% -\ bar\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}x{\isasymColon}bar{\isacharparenright}\ y\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{by}\isamarkupfalse% -\ default\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq{\isacharunderscore}bar{\isacharunderscore}def{\isacharparenright}\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -\isanewline -% -\isadelimquotett -\isanewline -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ bar\isanewline -\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Integer{\isachardoublequoteclose}{\isacharparenright}% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\begin{isamarkuptext}% -\noindent The code generator would produce - an additional instance, which of course is rejected by the \isa{Haskell} - compiler. - To suppress this additional instance, use - \isa{code{\isacharunderscore}instance}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquotett -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}instance}\isamarkupfalse% -\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline -\ \ {\isacharparenleft}Haskell\ {\isacharminus}{\isacharparenright}% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\isamarkupsubsection{Enhancing the target language context \label{sec:include}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In rare cases it is necessary to \emph{enrich} the context of a - target language; this is accomplished using the \hyperlink{command.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}} - command:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquotett -% -\endisadelimquotett -% -\isatagquotett -\isacommand{code{\isacharunderscore}include}\isamarkupfalse% -\ Haskell\ {\isachardoublequoteopen}Errno{\isachardoublequoteclose}\isanewline -{\isacharverbatimopen}errno\ i\ {\isacharequal}\ error\ {\isacharparenleft}{\isachardoublequote}Error\ number{\isacharcolon}\ {\isachardoublequote}\ {\isacharplus}{\isacharplus}\ show\ i{\isacharparenright}{\isacharverbatimclose}\isanewline -\isanewline -\isacommand{code{\isacharunderscore}reserved}\isamarkupfalse% -\ Haskell\ Errno% -\endisatagquotett -{\isafoldquotett}% -% -\isadelimquotett -% -\endisadelimquotett -% -\begin{isamarkuptext}% -\noindent Such named \isa{include}s are then prepended to every generated code. - Inspect such code in order to find out how \hyperlink{command.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}} behaves - with respect to a particular target language.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/document/Codegen.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/Codegen.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1690 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Codegen}% -% -\isadelimtheory -\isanewline -\isanewline -% -\endisadelimtheory -% -\isatagtheory -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\isamarkupchapter{Code generation from Isabelle theories% -} -\isamarkuptrue% -% -\isamarkupsection{Introduction% -} -\isamarkuptrue% -% -\isamarkupsubsection{Motivation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Executing formal specifications as programs is a well-established - topic in the theorem proving community. With increasing - application of theorem proving systems in the area of - software development and verification, its relevance manifests - for running test cases and rapid prototyping. In logical - calculi like constructive type theory, - a notion of executability is implicit due to the nature - of the calculus. In contrast, specifications in Isabelle - can be highly non-executable. In order to bridge - the gap between logic and executable specifications, - an explicit non-trivial transformation has to be applied: - code generation. - - This tutorial introduces a generic code generator for the - Isabelle system \cite{isa-tutorial}. - Generic in the sense that the - \qn{target language} for which code shall ultimately be - generated is not fixed but may be an arbitrary state-of-the-art - functional programming language (currently, the implementation - supports SML \cite{SML}, OCaml \cite{OCaml} and Haskell - \cite{haskell-revised-report}). - We aim to provide a - versatile environment - suitable for software development and verification, - structuring the process - of code generation into a small set of orthogonal principles - while achieving a big coverage of application areas - with maximum flexibility. - - Conceptually the code generator framework is part - of Isabelle's \isa{Pure} meta logic; the object logic - \isa{HOL} which is an extension of \isa{Pure} - already comes with a reasonable framework setup and thus provides - a good working horse for raising code-generation-driven - applications. So, we assume some familiarity and experience - with the ingredients of the \isa{HOL} \emph{Main} theory - (see also \cite{isa-tutorial}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Overview% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The code generator aims to be usable with no further ado - in most cases while allowing for detailed customization. - This manifests in the structure of this tutorial: - we start with a generic example \secref{sec:example} - and introduce code generation concepts \secref{sec:concept}. - Section - \secref{sec:basics} explains how to use the framework naively, - presuming a reasonable default setup. Then, section - \secref{sec:advanced} deals with advanced topics, - introducing further aspects of the code generator framework - in a motivation-driven manner. Last, section \secref{sec:ml} - introduces the framework's internal programming interfaces. - - \begin{warn} - Ultimately, the code generator which this tutorial deals with - is supposed to replace the already established code generator - by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. - So, for the moment, there are two distinct code generators - in Isabelle. - Also note that while the framework itself is - object-logic independent, only \isa{HOL} provides a reasonable - framework setup. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{An example: a simple theory of search trees \label{sec:example}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -When writing executable specifications using \isa{HOL}, - it is convenient to use - three existing packages: the datatype package for defining - datatypes, the function package for (recursive) functions, - and the class package for overloaded definitions. - - We develope a small theory of search trees; trees are represented - as a datatype with key type \isa{{\isacharprime}a} and value type \isa{{\isacharprime}b}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isacharequal}\ Leaf\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}linorder{\isachardoublequoteclose}\ {\isacharprime}b\isanewline -\ \ {\isacharbar}\ Branch\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharprime}a{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent Note that we have constrained the type of keys - to the class of total orders, \isa{linorder}. - - We define \isa{find} and \isa{update} functions:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ find\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isasymColon}linorder{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}find\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ it\ {\isacharequal}\ {\isacharparenleft}if\ it\ {\isacharequal}\ key\ then\ Some\ val\ else\ None{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}find\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ t{\isadigit{2}}{\isacharparenright}\ it\ {\isacharequal}\ {\isacharparenleft}if\ it\ {\isasymle}\ key\ then\ find\ t{\isadigit{1}}\ it\ else\ find\ t{\isadigit{2}}\ it{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{fun}\isamarkupfalse% -\isanewline -\ \ update\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}linorder\ {\isasymtimes}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}\isanewline -\ \ \ \ if\ it\ {\isacharequal}\ key\ then\ Leaf\ key\ entry\isanewline -\ \ \ \ \ \ else\ if\ it\ {\isasymle}\ key\isanewline -\ \ \ \ \ \ then\ Branch\ {\isacharparenleft}Leaf\ it\ entry{\isacharparenright}\ it\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\isanewline -\ \ \ \ \ \ else\ Branch\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ it\ {\isacharparenleft}Leaf\ it\ entry{\isacharparenright}\isanewline -\ \ \ {\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ t{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}\isanewline -\ \ \ \ if\ it\ {\isasymle}\ key\isanewline -\ \ \ \ \ \ then\ {\isacharparenleft}Branch\ {\isacharparenleft}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ t{\isadigit{1}}{\isacharparenright}\ key\ t{\isadigit{2}}{\isacharparenright}\isanewline -\ \ \ \ \ \ else\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ {\isacharparenleft}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ t{\isadigit{2}}{\isacharparenright}{\isacharparenright}\isanewline -\ \ \ {\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent For testing purpose, we define a small example - using natural numbers \isa{nat} (which are a \isa{linorder}) - as keys and list of nats as values:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ example\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat{\isacharcomma}\ nat\ list{\isacharparenright}\ searchtree{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}example\ {\isacharequal}\ update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\ {\isacharparenleft}update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\isanewline -\ \ \ \ {\isacharparenleft}update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\ {\isacharparenleft}Leaf\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent Then we generate code% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ example\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}tree{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent which looks like: - \lstsml{Thy/examples/tree.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Code generation concepts and process \label{sec:concept}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{figure}[h] - \centering - \includegraphics[width=0.7\textwidth]{codegen_process} - \caption{code generator -- processing overview} - \label{fig:process} - \end{figure} - - The code generator employs a notion of executability - for three foundational executable ingredients known - from functional programming: - \emph{defining equations}, \emph{datatypes}, and - \emph{type classes}. A defining equation as a first approximation - is a theorem of the form \isa{f\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n\ {\isasymequiv}\ t} - (an equation headed by a constant \isa{f} with arguments - \isa{t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n} and right hand side \isa{t}). - Code generation aims to turn defining equations - into a functional program by running through - a process (see figure \ref{fig:process}): - - \begin{itemize} - - \item Out of the vast collection of theorems proven in a - \qn{theory}, a reasonable subset modeling - defining equations is \qn{selected}. - - \item On those selected theorems, certain - transformations are carried out - (\qn{preprocessing}). Their purpose is to turn theorems - representing non- or badly executable - specifications into equivalent but executable counterparts. - The result is a structured collection of \qn{code theorems}. - - \item These \qn{code theorems} then are \qn{translated} - into an Haskell-like intermediate - language. - - \item Finally, out of the intermediate language the final - code in the desired \qn{target language} is \qn{serialized}. - - \end{itemize} - - From these steps, only the two last are carried out - outside the logic; by keeping this layer as - thin as possible, the amount of code to trust is - kept to a minimum.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Basics \label{sec:basics}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Invoking the code generator% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Thanks to a reasonable setup of the \isa{HOL} theories, in - most cases code generation proceeds without further ado:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ fac\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}fac\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}fac\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ Suc\ n\ {\isacharasterisk}\ fac\ n{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent This executable specification is now turned to SML code:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ fac\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}fac{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent The \isa{{\isasymEXPORTCODE}} command takes a space-separated list of - constants together with \qn{serialization directives} - These start with a \qn{target language} - identifier, followed by a file specification - where to write the generated code to. - - Internally, the defining equations for all selected - constants are taken, including any transitively required - constants, datatypes and classes, resulting in the following - code: - - \lstsml{Thy/examples/fac.ML} - - The code generator will complain when a required - ingredient does not provide a executable counterpart, - e.g.~generating code - for constants not yielding - a defining equation (e.g.~the Hilbert choice - operation \isa{SOME}):% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ pick{\isacharunderscore}some\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}pick{\isacharunderscore}some\ xs\ {\isacharequal}\ {\isacharparenleft}SOME\ x{\isachardot}\ x\ {\isasymin}\ set\ xs{\isacharparenright}{\isachardoublequoteclose}% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -\isanewline -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ pick{\isacharunderscore}some\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}fail{\isacharunderscore}const{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent will fail.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Theorem selection% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The list of all defining equations in a theory may be inspected - using the \isa{{\isasymPRINTCODESETUP}} command:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{print{\isacharunderscore}codesetup}\isamarkupfalse% -% -\begin{isamarkuptext}% -\noindent which displays a table of constant with corresponding - defining equations (the additional stuff displayed - shall not bother us for the moment). - - The typical \isa{HOL} tools are already set up in a way that - function definitions introduced by \isa{{\isasymDEFINITION}}, - \isa{{\isasymPRIMREC}}, \isa{{\isasymFUN}}, - \isa{{\isasymFUNCTION}}, \isa{{\isasymCONSTDEFS}}, - \isa{{\isasymRECDEF}} are implicitly propagated - to this defining equation table. Specific theorems may be - selected using an attribute: \emph{code func}. As example, - a weight selector function:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ pick\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ list\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}pick\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ n\ {\isacharequal}\ {\isacharparenleft}let\ {\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}\ {\isacharequal}\ x\ in\isanewline -\ \ \ \ if\ n\ {\isacharless}\ k\ then\ v\ else\ pick\ xs\ {\isacharparenleft}n\ {\isacharminus}\ k{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent We want to eliminate the explicit destruction - of \isa{x} to \isa{{\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}pick\ {\isacharparenleft}{\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}{\isacharhash}xs{\isacharparenright}\ n\ {\isacharequal}\ {\isacharparenleft}if\ n\ {\isacharless}\ k\ then\ v\ else\ pick\ xs\ {\isacharparenleft}n\ {\isacharminus}\ k{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ pick\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}pick{\isadigit{1}}{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent This theorem now is used for generating code: - - \lstsml{Thy/examples/pick1.ML} - - \noindent The policy is that \emph{default equations} stemming from - \isa{{\isasymDEFINITION}}, - \isa{{\isasymPRIMREC}}, \isa{{\isasymFUN}}, - \isa{{\isasymFUNCTION}}, \isa{{\isasymCONSTDEFS}}, - \isa{{\isasymRECDEF}} statements are discarded as soon as an - equation is explicitly selected by means of \emph{code func}. - Further applications of \emph{code func} add theorems incrementally, - but syntactic redundancies are implicitly dropped. For example, - using a modified version of the \isa{fac} function - as defining equation, the then redundant (since - syntactically subsumed) original defining equations - are dropped. - - \begin{warn} - The attributes \emph{code} and \emph{code del} - associated with the existing code generator also apply to - the new one: \emph{code} implies \emph{code func}, - and \emph{code del} implies \emph{code func del}. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Type classes% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Type classes enter the game via the Isar class package. - For a short introduction how to use it, see \cite{isabelle-classes}; - here we just illustrate its impact on code generation. - - In a target language, type classes may be represented - natively (as in the case of Haskell). For languages - like SML, they are implemented using \emph{dictionaries}. - Our following example specifies a class \qt{null}, - assigning to each of its inhabitants a \qt{null} value:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{class}\isamarkupfalse% -\ null\ {\isacharequal}\ type\ {\isacharplus}\isanewline -\ \ \isakeyword{fixes}\ null\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline -\isanewline -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ head\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}null\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}head\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ null{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}head\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent We provide some instances for our \isa{null}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{instantiation}\isamarkupfalse% -\ option\ \isakeyword{and}\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}\ null\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isachardoublequoteopen}null\ {\isacharequal}\ None{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isachardoublequoteopen}null\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -\noindent Constructing a dummy example:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isachardoublequoteopen}dummy\ {\isacharequal}\ head\ {\isacharbrackleft}Some\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ None{\isacharbrackright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -Type classes offer a suitable occasion to introduce - the Haskell serializer. Its usage is almost the same - as SML, but, in accordance with conventions - some Haskell systems enforce, each module ends - up in a single file. The module hierarchy is reflected in - the file system, with root directory given as file specification.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ dummy\ \isakeyword{in}\ Haskell\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lsthaskell{Thy/examples/Codegen.hs} - \noindent (we have left out all other modules). - - \medskip - - The whole code in SML with explicit dictionary passing:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ dummy\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}class{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/class.ML} - - \medskip - - \noindent or in OCaml:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ dummy\ \isakeyword{in}\ OCaml\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}class{\isachardot}ocaml{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/class.ocaml} - - \medskip The explicit association of constants - to classes can be inspected using the \isa{{\isasymPRINTCLASSES}} - command.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Recipes and advanced topics \label{sec:advanced}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In this tutorial, we do not attempt to give an exhaustive - description of the code generator framework; instead, - we cast a light on advanced topics by introducing - them together with practically motivated examples. Concerning - further reading, see - - \begin{itemize} - - \item the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} - for exhaustive syntax diagrams. - \item or \cite{Haftmann-Nipkow:2007:codegen} which deals with foundational issues - of the code generator framework. - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Library theories \label{sec:library}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The \isa{HOL} \isa{Main} theory already provides a code - generator setup - which should be suitable for most applications. Common extensions - and modifications are available by certain theories of the \isa{HOL} - library; beside being useful in applications, they may serve - as a tutorial for customizing the code generator setup. - - \begin{description} - - \item[\isa{Code{\isacharunderscore}Integer}] represents \isa{HOL} integers by big - integer literals in target languages. - \item[\isa{Code{\isacharunderscore}Char}] represents \isa{HOL} characters by - character literals in target languages. - \item[\isa{Code{\isacharunderscore}Char{\isacharunderscore}chr}] like \isa{Code{\isacharunderscore}Char}, - but also offers treatment of character codes; includes - \isa{Code{\isacharunderscore}Integer}. - \item[\isa{Efficient{\isacharunderscore}Nat}] \label{eff_nat} implements natural numbers by integers, - which in general will result in higher efficency; pattern - matching with \isa{{\isadigit{0}}} / \isa{Suc} - is eliminated; includes \isa{Code{\isacharunderscore}Integer}. - \item[\isa{Code{\isacharunderscore}Index}] provides an additional datatype - \isa{index} which is mapped to target-language built-in integers. - Useful for code setups which involve e.g. indexing of - target-language arrays. - \item[\isa{Code{\isacharunderscore}Message}] provides an additional datatype - \isa{message{\isacharunderscore}string} which is isomorphic to strings; - \isa{message{\isacharunderscore}string}s are mapped to target-language strings. - Useful for code setups which involve e.g. printing (error) messages. - - \end{description} - - \begin{warn} - When importing any of these theories, they should form the last - items in an import list. Since these theories adapt the - code generator setup in a non-conservative fashion, - strange effects may occur otherwise. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Preprocessing% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Before selected function theorems are turned into abstract - code, a chain of definitional transformation steps is carried - out: \emph{preprocessing}. In essence, the preprocessor - consists of two components: a \emph{simpset} and \emph{function transformers}. - - The \emph{simpset} allows to employ the full generality of the Isabelle - simplifier. Due to the interpretation of theorems - as defining equations, rewrites are applied to the right - hand side and the arguments of the left hand side of an - equation, but never to the constant heading the left hand side. - An important special case are \emph{inline theorems} which may be - declared an undeclared using the - \emph{code inline} or \emph{code inline del} attribute respectively. - Some common applications:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{itemize} -% -\begin{isamarkuptext}% -\item replacing non-executable constructs by executable ones:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ \ \ {\isachardoublequoteopen}x\ {\isasymin}\ set\ xs\ {\isasymlongleftrightarrow}\ x\ mem\ xs{\isachardoublequoteclose}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\item eliminating superfluous constants:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ \ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\item replacing executable but inconvenient constructs:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ \ \ {\isachardoublequoteopen}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongleftrightarrow}\ List{\isachardot}null\ xs{\isachardoublequoteclose}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\end{itemize} -% -\begin{isamarkuptext}% -\emph{Function transformers} provide a very general interface, - transforming a list of function theorems to another - list of function theorems, provided that neither the heading - constant nor its type change. The \isa{{\isadigit{0}}} / \isa{Suc} - pattern elimination implemented in - theory \isa{Efficient{\isacharunderscore}Nat} (see \secref{eff_nat}) uses this - interface. - - \noindent The current setup of the preprocessor may be inspected using - the \isa{{\isasymPRINTCODESETUP}} command. - - \begin{warn} - The attribute \emph{code unfold} - associated with the existing code generator also applies to - the new one: \emph{code unfold} implies \emph{code inline}. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Concerning operational equality% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Surely you have already noticed how equality is treated - by the code generator:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ collect{\isacharunderscore}duplicates\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharparenleft}z{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ z\ {\isasymin}\ set\ xs\isanewline -\ \ \ \ \ \ then\ if\ z\ {\isasymin}\ set\ ys\isanewline -\ \ \ \ \ \ \ \ then\ collect{\isacharunderscore}duplicates\ xs\ ys\ zs\isanewline -\ \ \ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ xs\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs\isanewline -\ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ {\isacharparenleft}z{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -The membership test during preprocessing is rewritten, - resulting in \isa{op\ mem}, which itself - performs an explicit equality check.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ collect{\isacharunderscore}duplicates\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}collect{\isacharunderscore}duplicates{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/collect_duplicates.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -Obviously, polymorphic equality is implemented the Haskell - way using a type class. How is this achieved? HOL introduces - an explicit class \isa{eq} with a corresponding operation - \isa{eq{\isacharunderscore}class{\isachardot}eq} such that \isa{eq{\isacharunderscore}class{\isachardot}eq\ x\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ y{\isacharparenright}}. - The preprocessing framework does the rest. - For datatypes, instances of \isa{eq} are implicitly derived - when possible. For other types, you may instantiate \isa{eq} - manually like any other type class. - - Though this \isa{eq} class is designed to get rarely in - the way, a subtlety - enters the stage when definitions of overloaded constants - are dependent on operational equality. For example, let - us define a lexicographic ordering on tuples:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{instantiation}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}ord{\isacharcomma}\ ord{\isacharparenright}\ ord\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isacharbrackleft}code\ func\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isacharless}\ p{\isadigit{2}}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}let\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isacharbrackleft}code\ func\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isasymle}\ p{\isadigit{2}}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}let\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ ord{\isacharunderscore}prod\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}ord{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}ord{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{unfolding}\isamarkupfalse% -\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp{\isacharunderscore}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Then code generation will fail. Why? The definition - of \isa{op\ {\isasymle}} depends on equality on both arguments, - which are polymorphic and impose an additional \isa{eq} - class constraint, thus violating the type discipline - for class operations. - - The solution is to add \isa{eq} explicitly to the first sort arguments in the - code theorems:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ ord{\isacharunderscore}prod{\isacharunderscore}code\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}ord{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}ord{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{unfolding}\isamarkupfalse% -\ ord{\isacharunderscore}prod\ \isacommand{by}\isamarkupfalse% -\ rule{\isacharplus}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Then code generation succeeds:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}eq{\isacharcomma}\ ord{\isacharbraceright}\ {\isasymtimes}\ {\isacharprime}b{\isasymColon}ord\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}lexicographic{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/lexicographic.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -In general, code theorems for overloaded constants may have more - restrictive sort constraints than the underlying instance relation - between class and type constructor as long as the whole system of - constraints is coregular; code theorems violating coregularity - are rejected immediately. Consequently, it might be necessary - to delete disturbing theorems in the code theorem table, - as we have done here with the original definitions \isa{less{\isacharunderscore}prod{\isacharunderscore}def} - and \isa{less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def}. - - In some cases, the automatically derived defining equations - for equality on a particular type may not be appropriate. - As example, watch the following datatype representing - monomorphic parametric types (where type constructors - are referred to by natural numbers):% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ monotype\ {\isacharequal}\ Mono\ nat\ {\isachardoublequoteopen}monotype\ list{\isachardoublequoteclose}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Then code generation for SML would fail with a message - that the generated code conains illegal mutual dependencies: - the theorem \isa{Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymequiv}\ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ typargs{\isadigit{1}}\ {\isacharequal}\ typargs{\isadigit{2}}} already requires the - instance \isa{monotype\ {\isasymColon}\ eq}, which itself requires - \isa{Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymequiv}\ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ typargs{\isadigit{1}}\ {\isacharequal}\ typargs{\isadigit{2}}}; Haskell has no problem with mutually - recursive \isa{instance} and \isa{function} definitions, - but the SML serializer does not support this. - - In such cases, you have to provide you own equality equations - involving auxiliary constants. In our case, - \isa{list{\isacharunderscore}all{\isadigit{2}}} can do the job:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ monotype{\isacharunderscore}eq{\isacharunderscore}list{\isacharunderscore}all{\isadigit{2}}\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ list{\isacharunderscore}all{\isadigit{2}}\ {\isacharparenleft}op\ {\isacharequal}{\isacharparenright}\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ list{\isacharunderscore}all{\isadigit{2}}{\isacharunderscore}eq\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -does not depend on instance \isa{monotype\ {\isasymColon}\ eq}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isacharequal}\ {\isacharcolon}{\isacharcolon}\ monotype\ {\isasymRightarrow}\ monotype\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}monotype{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/monotype.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Programs as sets of theorems% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -As told in \secref{sec:concept}, code generation is based - on a structured collection of code theorems. - For explorative purpose, this collection - may be inspected using the \isa{{\isasymCODETHMS}} command:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{code{\isacharunderscore}thms}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ mod\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent prints a table with \emph{all} defining equations - for \isa{op\ mod}, including - \emph{all} defining equations those equations depend - on recursivly. \isa{{\isasymCODETHMS}} provides a convenient - mechanism to inspect the impact of a preprocessor setup - on defining equations. - - Similarly, the \isa{{\isasymCODEDEPS}} command shows a graph - visualizing dependencies between defining equations.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Constructor sets for datatypes% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Conceptually, any datatype is spanned by a set of - \emph{constructors} of type \isa{{\isasymtau}\ {\isacharequal}\ {\isasymdots}\ {\isasymRightarrow}\ {\isasymkappa}\ {\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymdots}\ {\isasymalpha}\isactrlisub n} - where \isa{{\isacharbraceleft}{\isasymalpha}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlisub n{\isacharbraceright}} is excactly the set of \emph{all} - type variables in \isa{{\isasymtau}}. The HOL datatype package - by default registers any new datatype in the table - of datatypes, which may be inspected using - the \isa{{\isasymPRINTCODESETUP}} command. - - In some cases, it may be convenient to alter or - extend this table; as an example, we will develope an alternative - representation of natural numbers as binary digits, whose - size does increase logarithmically with its value, not linear - \footnote{Indeed, the \isa{Efficient{\isacharunderscore}Nat} theory (see \ref{eff_nat}) - does something similar}. First, the digit representation:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{definition}\isamarkupfalse% -\ Dig{\isadigit{0}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ n\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ n{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ Dig{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ n\ {\isacharequal}\ Suc\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent We will use these two ">digits"< to represent natural numbers - in binary digits, e.g.:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isadigit{4}}{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}Dig{\isadigit{1}}\ {\isacharparenleft}Dig{\isadigit{0}}\ {\isacharparenleft}Dig{\isadigit{1}}\ {\isacharparenleft}Dig{\isadigit{0}}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ Dig{\isadigit{0}}{\isacharunderscore}def\ Dig{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Of course we also have to provide proper code equations for - the operations, e.g. \isa{op\ {\isacharplus}}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ plus{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isadigit{0}}\ {\isacharplus}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}m\ {\isacharplus}\ {\isadigit{0}}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ n{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ {\isadigit{1}}\ {\isacharequal}\ Dig{\isadigit{1}}\ m{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ {\isadigit{1}}\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ Dig{\isadigit{0}}{\isacharunderscore}def\ Dig{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent We then instruct the code generator to view \isa{{\isadigit{0}}}, - \isa{{\isadigit{1}}}, \isa{Dig{\isadigit{0}}} and \isa{Dig{\isadigit{1}}} as - datatype constructors:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isadigit{0}}{\isasymColon}nat{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isadigit{1}}{\isasymColon}nat{\isachardoublequoteclose}\ Dig{\isadigit{0}}\ Dig{\isadigit{1}}% -\begin{isamarkuptext}% -\noindent For the former constructor \isa{Suc}, we provide a code - equation and remove some parts of the default code generator setup - which are an obstacle here:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ Suc{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}Suc\ n\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{declare}\isamarkupfalse% -\ One{\isacharunderscore}nat{\isacharunderscore}def\ {\isacharbrackleft}code\ inline\ del{\isacharbrackright}\isanewline -\isacommand{declare}\isamarkupfalse% -\ add{\isacharunderscore}Suc{\isacharunderscore}shift\ {\isacharbrackleft}code\ func\ del{\isacharbrackright}% -\begin{isamarkuptext}% -\noindent This yields the following code:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isacharplus}\ {\isasymColon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}nat{\isacharunderscore}binary{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/nat_binary.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\medskip - - From this example, it can be easily glimpsed that using own constructor sets - is a little delicate since it changes the set of valid patterns for values - of that type. Without going into much detail, here some practical hints: - - \begin{itemize} - \item When changing the constuctor set for datatypes, take care to - provide an alternative for the \isa{case} combinator (e.g. by replacing - it using the preprocessor). - \item Values in the target language need not to be normalized -- different - values in the target language may represent the same value in the - logic (e.g. \isa{Dig{\isadigit{1}}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}}). - \item Usually, a good methodology to deal with the subleties of pattern - matching is to see the type as an abstract type: provide a set - of operations which operate on the concrete representation of the type, - and derive further operations by combinations of these primitive ones, - without relying on a particular representation. - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isataginvisible -\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isadigit{0}}{\isacharcolon}{\isacharcolon}nat{\isachardoublequoteclose}\ Suc\isanewline -\isacommand{declare}\isamarkupfalse% -\ plus{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func\ del{\isacharbrackright}\isanewline -\isacommand{declare}\isamarkupfalse% -\ One{\isacharunderscore}nat{\isacharunderscore}def\ {\isacharbrackleft}code\ inline{\isacharbrackright}\isanewline -\isacommand{declare}\isamarkupfalse% -\ add{\isacharunderscore}Suc{\isacharunderscore}shift\ {\isacharbrackleft}code\ func{\isacharbrackright}\ \isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isadigit{0}}\ {\isacharplus}\ n\ {\isacharequal}\ {\isacharparenleft}n\ {\isasymColon}\ nat{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isamarkupsubsection{Customizing serialization% -} -\isamarkuptrue% -% -\isamarkupsubsubsection{Basics% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Consider the following function and its corresponding - SML code:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ in{\isacharunderscore}interval\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymtimes}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}in{\isacharunderscore}interval\ {\isacharparenleft}k{\isacharcomma}\ l{\isacharparenright}\ n\ {\isasymlongleftrightarrow}\ k\ {\isasymle}\ n\ {\isasymand}\ n\ {\isasymle}\ l{\isachardoublequoteclose}% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ in{\isacharunderscore}interval\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}literal{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/bool_literal.ML} - - \noindent Though this is correct code, it is a little bit unsatisfactory: - boolean values and operators are materialized as distinguished - entities with have nothing to do with the SML-builtin notion - of \qt{bool}. This results in less readable code; - additionally, eager evaluation may cause programs to - loop or break which would perfectly terminate when - the existing SML \qt{bool} would be used. To map - the HOL \qt{bool} on SML \qt{bool}, we may use - \qn{custom serializations}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ bool\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}bool{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ True\ \isakeyword{and}\ False\ \isakeyword{and}\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}true{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}false{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}{\isacharunderscore}\ andalso\ {\isacharunderscore}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\begin{isamarkuptext}% -The \isa{{\isasymCODETYPE}} commad takes a type constructor - as arguments together with a list of custom serializations. - Each custom serialization starts with a target language - identifier followed by an expression, which during - code serialization is inserted whenever the type constructor - would occur. For constants, \isa{{\isasymCODECONST}} implements - the corresponding mechanism. Each ``\verb|_|'' in - a serialization expression is treated as a placeholder - for the type constructor's (the constant's) arguments.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ in{\isacharunderscore}interval\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}mlbool{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/bool_mlbool.ML} - - \noindent This still is not perfect: the parentheses - around the \qt{andalso} expression are superfluous. - Though the serializer - by no means attempts to imitate the rich Isabelle syntax - framework, it provides some common idioms, notably - associative infixes with precedences which may be used here:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}SML\ \isakeyword{infixl}\ {\isadigit{1}}\ {\isachardoublequoteopen}andalso{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -\isanewline -\isanewline -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ in{\isacharunderscore}interval\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}infix{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/bool_infix.ML} - - \medskip - - Next, we try to map HOL pairs to SML pairs, using the - infix ``\verb|*|'' type constructor and parentheses:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ {\isacharasterisk}\isanewline -\ \ {\isacharparenleft}SML\ \isakeyword{infix}\ {\isadigit{2}}\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ Pair\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}{\isacharbang}{\isacharparenleft}{\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharcomma}{\isacharslash}\ {\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\begin{isamarkuptext}% -The initial bang ``\verb|!|'' tells the serializer to never put - parentheses around the whole expression (they are already present), - while the parentheses around argument place holders - tell not to put parentheses around the arguments. - The slash ``\verb|/|'' (followed by arbitrary white space) - inserts a space which may be used as a break if necessary - during pretty printing. - - These examples give a glimpse what mechanisms - custom serializations provide; however their usage - requires careful thinking in order not to introduce - inconsistencies -- or, in other words: - custom serializations are completely axiomatic. - - A further noteworthy details is that any special - character in a custom serialization may be quoted - using ``\verb|'|''; thus, in - ``\verb|fn '_ => _|'' the first - ``\verb|_|'' is a proper underscore while the - second ``\verb|_|'' is a placeholder. - - The HOL theories provide further - examples for custom serializations.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Haskell serialization% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -For convenience, the default - HOL setup for Haskell maps the \isa{eq} class to - its counterpart in Haskell, giving custom serializations - for the class (\isa{{\isasymCODECLASS}}) and its operation:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}class}\isamarkupfalse% -\ eq\isanewline -\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Eq{\isachardoublequoteclose}\ \isakeyword{where}\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\ {\isasymequiv}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharequal}{\isacharequal}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}Haskell\ \isakeyword{infixl}\ {\isadigit{4}}\ {\isachardoublequoteopen}{\isacharequal}{\isacharequal}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\begin{isamarkuptext}% -A problem now occurs whenever a type which - is an instance of \isa{eq} in HOL is mapped - on a Haskell-builtin type which is also an instance - of Haskell \isa{Eq}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{typedecl}\isamarkupfalse% -\ bar\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}x{\isasymColon}bar{\isacharparenright}\ y\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ default\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq{\isacharunderscore}bar{\isacharunderscore}def{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -% -\isadelimtt -\isanewline -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ bar\isanewline -\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Integer{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\begin{isamarkuptext}% -The code generator would produce - an additional instance, which of course is rejected. - To suppress this additional instance, use - \isa{{\isasymCODEINSTANCE}}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}instance}\isamarkupfalse% -\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline -\ \ {\isacharparenleft}Haskell\ {\isacharminus}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\isamarkupsubsubsection{Pretty printing% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The serializer provides ML interfaces to set up - pretty serializations for expressions like lists, numerals - and characters; these are - monolithic stubs and should only be used with the - theories introduced in \secref{sec:library}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Cyclic module dependencies% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Sometimes the awkward situation occurs that dependencies - between definitions introduce cyclic dependencies - between modules, which in the Haskell world leaves - you to the mercy of the Haskell implementation you are using, - while for SML code generation is not possible. - - A solution is to declare module names explicitly. - Let use assume the three cyclically dependent - modules are named \emph{A}, \emph{B} and \emph{C}. - Then, by stating% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{code{\isacharunderscore}modulename}\isamarkupfalse% -\ SML\isanewline -\ \ A\ ABC\isanewline -\ \ B\ ABC\isanewline -\ \ C\ ABC% -\begin{isamarkuptext}% -we explicitly map all those modules on \emph{ABC}, - resulting in an ad-hoc merge of this three modules - at serialization time.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Incremental code generation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Code generation is \emph{incremental}: theorems - and abstract intermediate code are cached and extended on demand. - The cache may be partially or fully dropped if the underlying - executable content of the theory changes. - Implementation of caching is supposed to transparently - hid away the details from the user. Anyway, caching - reaches the surface by using a slightly more general form - of the \isa{{\isasymCODETHMS}}, \isa{{\isasymCODEDEPS}} - and \isa{{\isasymEXPORTCODE}} commands: the list of constants - may be omitted. Then, all constants with code theorems - in the current cache are referred to.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{ML interfaces \label{sec:ml}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Since the code generator framework not only aims to provide - a nice Isar interface but also to form a base for - code-generation-based applications, here a short - description of the most important ML interfaces.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Executable theory content: \isa{Code}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -This Pure module implements the core notions of - executable content of a theory.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Managing executable content% -} -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Code.add\_func}\verb|Code.add_func: thm -> theory -> theory| \\ - \indexml{Code.del\_func}\verb|Code.del_func: thm -> theory -> theory| \\ - \indexml{Code.add\_funcl}\verb|Code.add_funcl: string * thm list Susp.T -> theory -> theory| \\ - \indexml{Code.map\_pre}\verb|Code.map_pre: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\ - \indexml{Code.map\_post}\verb|Code.map_post: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\ - \indexml{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> thm list -> thm list option)|\isasep\isanewline% -\verb| -> theory -> theory| \\ - \indexml{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\ - \indexml{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\ - \indexml{Code.get\_datatype}\verb|Code.get_datatype: theory -> string|\isasep\isanewline% -\verb| -> (string * sort) list * (string * typ list) list| \\ - \indexml{Code.get\_datatype\_of\_constr}\verb|Code.get_datatype_of_constr: theory -> string -> string option| - \end{mldecls} - - \begin{description} - - \item \verb|Code.add_func|~\isa{thm}~\isa{thy} adds function - theorem \isa{thm} to executable content. - - \item \verb|Code.del_func|~\isa{thm}~\isa{thy} removes function - theorem \isa{thm} from executable content, if present. - - \item \verb|Code.add_funcl|~\isa{{\isacharparenleft}const{\isacharcomma}\ lthms{\isacharparenright}}~\isa{thy} adds - suspended defining equations \isa{lthms} for constant - \isa{const} to executable content. - - \item \verb|Code.map_pre|~\isa{f}~\isa{thy} changes - the preprocessor simpset. - - \item \verb|Code.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds - function transformer \isa{f} (named \isa{name}) to executable content; - \isa{f} is a transformer of the defining equations belonging - to a certain function definition, depending on the - current theory context. Returning \isa{NONE} indicates that no - transformation took place; otherwise, the whole process will be iterated - with the new defining equations. - - \item \verb|Code.del_functrans|~\isa{name}~\isa{thy} removes - function transformer named \isa{name} from executable content. - - \item \verb|Code.add_datatype|~\isa{cs}~\isa{thy} adds - a datatype to executable content, with generation - set \isa{cs}. - - \item \verb|Code.get_datatype_of_constr|~\isa{thy}~\isa{const} - returns type constructor corresponding to - constructor \isa{const}; returns \isa{NONE} - if \isa{const} is no constructor. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Auxiliary% -} -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\ - \indexml{Code\_Unit.head\_func}\verb|Code_Unit.head_func: thm -> string * ((string * sort) list * typ)| \\ - \indexml{Code\_Unit.rewrite\_func}\verb|Code_Unit.rewrite_func: MetaSimplifier.simpset -> thm -> thm| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Code_Unit.read_const|~\isa{thy}~\isa{s} - reads a constant as a concrete term expression \isa{s}. - - \item \verb|Code_Unit.head_func|~\isa{thm} - extracts the constant and its type from a defining equation \isa{thm}. - - \item \verb|Code_Unit.rewrite_func|~\isa{ss}~\isa{thm} - rewrites a defining equation \isa{thm} with a simpset \isa{ss}; - only arguments and right hand side are rewritten, - not the head of the defining equation. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Implementing code generator applications% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Implementing code generator applications on top - of the framework set out so far usually not only - involves using those primitive interfaces - but also storing code-dependent data and various - other things. - - \begin{warn} - Some interfaces discussed here have not reached - a final state yet. - Changes likely to occur in future. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Data depending on the theory's executable content% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Due to incrementality of code generation, changes in the - theory's executable content have to be propagated in a - certain fashion. Additionally, such changes may occur - not only during theory extension but also during theory - merge, which is a little bit nasty from an implementation - point of view. The framework provides a solution - to this technical challenge by providing a functorial - data slot \verb|CodeDataFun|; on instantiation - of this functor, the following types and operations - are required: - - \medskip - \begin{tabular}{l} - \isa{type\ T} \\ - \isa{val\ empty{\isacharcolon}\ T} \\ - \isa{val\ merge{\isacharcolon}\ Pretty{\isachardot}pp\ {\isasymrightarrow}\ T\ {\isacharasterisk}\ T\ {\isasymrightarrow}\ T} \\ - \isa{val\ purge{\isacharcolon}\ theory\ option\ {\isasymrightarrow}\ CodeUnit{\isachardot}const\ list\ option\ {\isasymrightarrow}\ T\ {\isasymrightarrow}\ T} - \end{tabular} - - \begin{description} - - \item \isa{T} the type of data to store. - - \item \isa{empty} initial (empty) data. - - \item \isa{merge} merging two data slots. - - \item \isa{purge}~\isa{thy}~\isa{consts} propagates changes in executable content; - if possible, the current theory context is handed over - as argument \isa{thy} (if there is no current theory context (e.g.~during - theory merge, \verb|NONE|); \isa{consts} indicates the kind - of change: \verb|NONE| stands for a fundamental change - which invalidates any existing code, \isa{SOME\ consts} - hints that executable content for constants \isa{consts} - has changed. - - \end{description} - - An instance of \verb|CodeDataFun| provides the following - interface: - - \medskip - \begin{tabular}{l} - \isa{get{\isacharcolon}\ theory\ {\isasymrightarrow}\ T} \\ - \isa{change{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ T{\isacharparenright}\ {\isasymrightarrow}\ T} \\ - \isa{change{\isacharunderscore}yield{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T{\isacharparenright}\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T} - \end{tabular} - - \begin{description} - - \item \isa{get} retrieval of the current data. - - \item \isa{change} update of current data (cached!) - by giving a continuation. - - \item \isa{change{\isacharunderscore}yield} update with side result. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\emph{Happy proving, happy hacking!}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/document/Further.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/Further.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,234 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Further}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Further\isanewline -\isakeyword{imports}\ Setup\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupsection{Further issues \label{sec:further}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Further reading% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Do dive deeper into the issue of code generation, you should visit - the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} which - contains exhaustive syntax diagrams.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Modules% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -When invoking the \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command it is possible to leave - out the \hyperlink{keyword.module-name}{\mbox{\isa{\isakeyword{module{\isacharunderscore}name}}}} part; then code is distributed over - different modules, where the module name space roughly is induced - by the \isa{Isabelle} theory name space. - - Then sometimes the awkward situation occurs that dependencies between - definitions introduce cyclic dependencies between modules, which in the - \isa{Haskell} world leaves you to the mercy of the \isa{Haskell} implementation - you are using, while for \isa{SML}/\isa{OCaml} code generation is not possible. - - A solution is to declare module names explicitly. - Let use assume the three cyclically dependent - modules are named \emph{A}, \emph{B} and \emph{C}. - Then, by stating% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{code{\isacharunderscore}modulename}\isamarkupfalse% -\ SML\isanewline -\ \ A\ ABC\isanewline -\ \ B\ ABC\isanewline -\ \ C\ ABC% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -we explicitly map all those modules on \emph{ABC}, - resulting in an ad-hoc merge of this three modules - at serialisation time.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Evaluation oracle% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Code generation may also be used to \emph{evaluate} expressions - (using \isa{SML} as target language of course). - For instance, the \hyperlink{command.value}{\mbox{\isa{\isacommand{value}}}} allows to reduce an expression to a - normal form with respect to the underlying code equations:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{value}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharslash}\ {\isacharparenleft}{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ rat{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent will display \isa{{\isadigit{7}}\ {\isacharslash}\ {\isadigit{2}}}. - - The \hyperlink{method.eval}{\mbox{\isa{eval}}} method tries to reduce a goal by code generation to \isa{True} - and solves it in that case, but fails otherwise:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharslash}\ {\isacharparenleft}{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ rat{\isacharparenright}\ {\isacharequal}\ {\isadigit{7}}\ {\isacharslash}\ {\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ eval% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The soundness of the \hyperlink{method.eval}{\mbox{\isa{eval}}} method depends crucially - on the correctness of the code generator; this is one of the reasons - why you should not use adaption (see \secref{sec:adaption}) frivolously.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Code antiquotation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In scenarios involving techniques like reflection it is quite common - that code generated from a theory forms the basis for implementing - a proof procedure in \isa{SML}. To facilitate interfacing of generated code - with system code, the code generator provides a \isa{code} antiquotation:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{datatype}\isamarkupfalse% -\ form\ {\isacharequal}\ T\ {\isacharbar}\ F\ {\isacharbar}\ And\ form\ form\ {\isacharbar}\ Or\ form\ form\isanewline -\isanewline -\isacommand{ML}\isamarkupfalse% -\ {\isacharverbatimopen}\isanewline -\ \ fun\ eval{\isacharunderscore}form\ % -\isaantiq -code\ T% -\endisaantiq -\ {\isacharequal}\ true\isanewline -\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ % -\isaantiq -code\ F% -\endisaantiq -\ {\isacharequal}\ false\isanewline -\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ {\isacharparenleft}% -\isaantiq -code\ And% -\endisaantiq -\ {\isacharparenleft}p{\isacharcomma}\ q{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ \ \ \ eval{\isacharunderscore}form\ p\ andalso\ eval{\isacharunderscore}form\ q\isanewline -\ \ \ \ {\isacharbar}\ eval{\isacharunderscore}form\ {\isacharparenleft}% -\isaantiq -code\ Or% -\endisaantiq -\ {\isacharparenleft}p{\isacharcomma}\ q{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ \ \ \ eval{\isacharunderscore}form\ p\ orelse\ eval{\isacharunderscore}form\ q{\isacharsemicolon}\isanewline -{\isacharverbatimclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent \isa{code} takes as argument the name of a constant; after the - whole \isa{SML} is read, the necessary code is generated transparently - and the corresponding constant names are inserted. This technique also - allows to use pattern matching on constructors stemming from compiled - \isa{datatypes}. - - For a less simplistic example, theory \hyperlink{theory.Ferrack}{\mbox{\isa{Ferrack}}} is - a good reference.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Imperative data structures% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -If you consider imperative data structures as inevitable for a specific - application, you should consider - \emph{Imperative Functional Programming with Isabelle/HOL} - (\cite{bulwahn-et-al:2008:imperative}); - the framework described there is available in theory \hyperlink{theory.Imperative-HOL}{\mbox{\isa{Imperative{\isacharunderscore}HOL}}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/document/Introduction.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/Introduction.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,390 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Introduction}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Introduction\isanewline -\isakeyword{imports}\ Setup\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Code generation from \isa{Isabelle{\isacharslash}HOL} theories% -} -\isamarkuptrue% -% -\isamarkupsection{Introduction and Overview% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -This tutorial introduces a generic code generator for the - \isa{Isabelle} system. - Generic in the sense that the - \qn{target language} for which code shall ultimately be - generated is not fixed but may be an arbitrary state-of-the-art - functional programming language (currently, the implementation - supports \isa{SML} \cite{SML}, \isa{OCaml} \cite{OCaml} and \isa{Haskell} - \cite{haskell-revised-report}). - - Conceptually the code generator framework is part - of Isabelle's \hyperlink{theory.Pure}{\mbox{\isa{Pure}}} meta logic framework; the logic - \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} which is an extension of \hyperlink{theory.Pure}{\mbox{\isa{Pure}}} - already comes with a reasonable framework setup and thus provides - a good working horse for raising code-generation-driven - applications. So, we assume some familiarity and experience - with the ingredients of the \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} distribution theories. - (see also \cite{isa-tutorial}). - - The code generator aims to be usable with no further ado - in most cases while allowing for detailed customisation. - This manifests in the structure of this tutorial: after a short - conceptual introduction with an example (\secref{sec:intro}), - we discuss the generic customisation facilities (\secref{sec:program}). - A further section (\secref{sec:adaption}) is dedicated to the matter of - \qn{adaption} to specific target language environments. After some - further issues (\secref{sec:further}) we conclude with an overview - of some ML programming interfaces (\secref{sec:ml}). - - \begin{warn} - Ultimately, the code generator which this tutorial deals with - is supposed to replace the existing code generator - by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. - So, for the moment, there are two distinct code generators - in Isabelle. In case of ambiguity, we will refer to the framework - described here as \isa{generic\ code\ generator}, to the - other as \isa{SML\ code\ generator}. - Also note that while the framework itself is - object-logic independent, only \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} provides a reasonable - framework setup. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Code generation via shallow embedding \label{sec:intro}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The key concept for understanding \isa{Isabelle}'s code generation is - \emph{shallow embedding}, i.e.~logical entities like constants, types and - classes are identified with corresponding concepts in the target language. - - Inside \hyperlink{theory.HOL}{\mbox{\isa{HOL}}}, the \hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}} and - \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}/\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}/\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}} declarations form - the core of a functional programming language. The default code generator setup - allows to turn those into functional programs immediately. - This means that \qt{naive} code generation can proceed without further ado. - For example, here a simple \qt{implementation} of amortised queues:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{datatype}\isamarkupfalse% -\ {\isacharprime}a\ queue\ {\isacharequal}\ AQueue\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ empty\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{primrec}\isamarkupfalse% -\ enqueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ AQueue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ ys{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{fun}\isamarkupfalse% -\ dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ option\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ \ {\isacharparenleft}case\ rev\ xs\ of\ y\ {\isacharhash}\ ys\ {\isasymRightarrow}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ ys{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Then we can generate code e.g.~for \isa{SML} as follows:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ empty\ dequeue\ enqueue\ \isakeyword{in}\ SML\isanewline -\ \ \isakeyword{module{\isacharunderscore}name}\ Example\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}example{\isachardot}ML{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent resulting in the following code:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun foldl f a [] = a\\ -\hspace*{0pt} ~| foldl f a (x ::~xs) = foldl f (f a x) xs;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun rev xs = foldl (fn xsa => fn x => x ::~xsa) [] xs;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun list{\char95}case f1 f2 (a ::~lista) = f2 a lista\\ -\hspace*{0pt} ~| list{\char95}case f1 f2 [] = f1;\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype 'a queue = AQueue of 'a list * 'a list;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val empty :~'a queue = AQueue ([],~[])\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun dequeue (AQueue ([],~[])) = (NONE,~AQueue ([],~[]))\\ -\hspace*{0pt} ~| dequeue (AQueue (xs,~y ::~ys)) = (SOME y,~AQueue (xs,~ys))\\ -\hspace*{0pt} ~| dequeue (AQueue (v ::~va,~[])) =\\ -\hspace*{0pt} ~~~let\\ -\hspace*{0pt} ~~~~~val y ::~ys = rev (v ::~va);\\ -\hspace*{0pt} ~~~in\\ -\hspace*{0pt} ~~~~~(SOME y,~AQueue ([],~ys))\\ -\hspace*{0pt} ~~~end;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun enqueue x (AQueue (xs,~ys)) = AQueue (x ::~xs,~ys);\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command takes a space-separated list of - constants for which code shall be generated; anything else needed for those - is added implicitly. Then follows a target language identifier - (\isa{SML}, \isa{OCaml} or \isa{Haskell}) and a freely chosen module name. - A file name denotes the destination to store the generated code. Note that - the semantics of the destination depends on the target language: for - \isa{SML} and \isa{OCaml} it denotes a \emph{file}, for \isa{Haskell} - it denotes a \emph{directory} where a file named as the module name - (with extension \isa{{\isachardot}hs}) is written:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ empty\ dequeue\ enqueue\ \isakeyword{in}\ Haskell\isanewline -\ \ \isakeyword{module{\isacharunderscore}name}\ Example\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This is how the corresponding code in \isa{Haskell} looks like:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}module Example where {\char123}\\ -\hspace*{0pt}\\ -\hspace*{0pt}\\ -\hspace*{0pt}foldla ::~forall a b.~(a -> b -> a) -> a -> [b] -> a;\\ -\hspace*{0pt}foldla f a [] = a;\\ -\hspace*{0pt}foldla f a (x :~xs) = foldla f (f a x) xs;\\ -\hspace*{0pt}\\ -\hspace*{0pt}rev ::~forall a.~[a] -> [a];\\ -\hspace*{0pt}rev xs = foldla ({\char92}~xsa x -> x :~xsa) [] xs;\\ -\hspace*{0pt}\\ -\hspace*{0pt}list{\char95}case ::~forall t a.~t -> (a -> [a] -> t) -> [a] -> t;\\ -\hspace*{0pt}list{\char95}case f1 f2 (a :~list) = f2 a list;\\ -\hspace*{0pt}list{\char95}case f1 f2 [] = f1;\\ -\hspace*{0pt}\\ -\hspace*{0pt}data Queue a = AQueue [a] [a];\\ -\hspace*{0pt}\\ -\hspace*{0pt}empty ::~forall a.~Queue a;\\ -\hspace*{0pt}empty = AQueue [] [];\\ -\hspace*{0pt}\\ -\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\ -\hspace*{0pt}dequeue (AQueue [] []) = (Nothing,~AQueue [] []);\\ -\hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\ -\hspace*{0pt}dequeue (AQueue (v :~va) []) =\\ -\hspace*{0pt} ~let {\char123}\\ -\hspace*{0pt} ~~~(y :~ys) = rev (v :~va);\\ -\hspace*{0pt} ~{\char125}~in (Just y,~AQueue [] ys);\\ -\hspace*{0pt}\\ -\hspace*{0pt}enqueue ::~forall a.~a -> Queue a -> Queue a;\\ -\hspace*{0pt}enqueue x (AQueue xs ys) = AQueue (x :~xs) ys;\\ -\hspace*{0pt}\\ -\hspace*{0pt}{\char125}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This demonstrates the basic usage of the \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command; - for more details see \secref{sec:further}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Code generator architecture \label{sec:concept}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -What you have seen so far should be already enough in a lot of cases. If you - are content with this, you can quit reading here. Anyway, in order to customise - and adapt the code generator, it is inevitable to gain some understanding - how it works. - - \begin{figure}[h] - \begin{tikzpicture}[x = 4.2cm, y = 1cm] - \tikzstyle entity=[rounded corners, draw, thick, color = black, fill = white]; - \tikzstyle process=[ellipse, draw, thick, color = green, fill = white]; - \tikzstyle process_arrow=[->, semithick, color = green]; - \node (HOL) at (0, 4) [style=entity] {\isa{Isabelle{\isacharslash}HOL} theory}; - \node (eqn) at (2, 2) [style=entity] {code equations}; - \node (iml) at (2, 0) [style=entity] {intermediate language}; - \node (seri) at (1, 0) [style=process] {serialisation}; - \node (SML) at (0, 3) [style=entity] {\isa{SML}}; - \node (OCaml) at (0, 2) [style=entity] {\isa{OCaml}}; - \node (further) at (0, 1) [style=entity] {\isa{{\isasymdots}}}; - \node (Haskell) at (0, 0) [style=entity] {\isa{Haskell}}; - \draw [style=process_arrow] (HOL) .. controls (2, 4) .. - node [style=process, near start] {selection} - node [style=process, near end] {preprocessing} - (eqn); - \draw [style=process_arrow] (eqn) -- node (transl) [style=process] {translation} (iml); - \draw [style=process_arrow] (iml) -- (seri); - \draw [style=process_arrow] (seri) -- (SML); - \draw [style=process_arrow] (seri) -- (OCaml); - \draw [style=process_arrow, dashed] (seri) -- (further); - \draw [style=process_arrow] (seri) -- (Haskell); - \end{tikzpicture} - \caption{Code generator architecture} - \label{fig:arch} - \end{figure} - - The code generator employs a notion of executability - for three foundational executable ingredients known - from functional programming: - \emph{code equations}, \emph{datatypes}, and - \emph{type classes}. A code equation as a first approximation - is a theorem of the form \isa{f\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n\ {\isasymequiv}\ t} - (an equation headed by a constant \isa{f} with arguments - \isa{t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n} and right hand side \isa{t}). - Code generation aims to turn code equations - into a functional program. This is achieved by three major - components which operate sequentially, i.e. the result of one is - the input - of the next in the chain, see diagram \ref{fig:arch}: - - \begin{itemize} - - \item Out of the vast collection of theorems proven in a - \qn{theory}, a reasonable subset modelling - code equations is \qn{selected}. - - \item On those selected theorems, certain - transformations are carried out - (\qn{preprocessing}). Their purpose is to turn theorems - representing non- or badly executable - specifications into equivalent but executable counterparts. - The result is a structured collection of \qn{code theorems}. - - \item Before the selected code equations are continued with, - they can be \qn{preprocessed}, i.e. subjected to theorem - transformations. This \qn{preprocessor} is an interface which - allows to apply - the full expressiveness of ML-based theorem transformations - to code generation; motivating examples are shown below, see - \secref{sec:preproc}. - The result of the preprocessing step is a structured collection - of code equations. - - \item These code equations are \qn{translated} to a program - in an abstract intermediate language. Think of it as a kind - of \qt{Mini-Haskell} with four \qn{statements}: \isa{data} - (for datatypes), \isa{fun} (stemming from code equations), - also \isa{class} and \isa{inst} (for type classes). - - \item Finally, the abstract program is \qn{serialised} into concrete - source code of a target language. - - \end{itemize} - - \noindent From these steps, only the two last are carried out outside the logic; by - keeping this layer as thin as possible, the amount of code to trust is - kept to a minimum.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/document/ML.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/ML.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,255 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{ML}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ {\isachardoublequoteopen}ML{\isachardoublequoteclose}\isanewline -\isakeyword{imports}\ Setup\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupsection{ML system interfaces \label{sec:ml}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Since the code generator framework not only aims to provide - a nice Isar interface but also to form a base for - code-generation-based applications, here a short - description of the most important ML interfaces.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Executable theory content: \isa{Code}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -This Pure module implements the core notions of - executable content of a theory.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Managing executable content% -} -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Code.add\_eqn}\verb|Code.add_eqn: thm -> theory -> theory| \\ - \indexml{Code.del\_eqn}\verb|Code.del_eqn: thm -> theory -> theory| \\ - \indexml{Code.add\_eqnl}\verb|Code.add_eqnl: string * (thm * bool) list lazy -> theory -> theory| \\ - \indexml{Code.map\_pre}\verb|Code.map_pre: (simpset -> simpset) -> theory -> theory| \\ - \indexml{Code.map\_post}\verb|Code.map_post: (simpset -> simpset) -> theory -> theory| \\ - \indexml{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline% -\verb| -> theory -> theory| \\ - \indexml{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\ - \indexml{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\ - \indexml{Code.get\_datatype}\verb|Code.get_datatype: theory -> string|\isasep\isanewline% -\verb| -> (string * sort) list * (string * typ list) list| \\ - \indexml{Code.get\_datatype\_of\_constr}\verb|Code.get_datatype_of_constr: theory -> string -> string option| - \end{mldecls} - - \begin{description} - - \item \verb|Code.add_eqn|~\isa{thm}~\isa{thy} adds function - theorem \isa{thm} to executable content. - - \item \verb|Code.del_eqn|~\isa{thm}~\isa{thy} removes function - theorem \isa{thm} from executable content, if present. - - \item \verb|Code.add_eqnl|~\isa{{\isacharparenleft}const{\isacharcomma}\ lthms{\isacharparenright}}~\isa{thy} adds - suspended code equations \isa{lthms} for constant - \isa{const} to executable content. - - \item \verb|Code.map_pre|~\isa{f}~\isa{thy} changes - the preprocessor simpset. - - \item \verb|Code.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds - function transformer \isa{f} (named \isa{name}) to executable content; - \isa{f} is a transformer of the code equations belonging - to a certain function definition, depending on the - current theory context. Returning \isa{NONE} indicates that no - transformation took place; otherwise, the whole process will be iterated - with the new code equations. - - \item \verb|Code.del_functrans|~\isa{name}~\isa{thy} removes - function transformer named \isa{name} from executable content. - - \item \verb|Code.add_datatype|~\isa{cs}~\isa{thy} adds - a datatype to executable content, with generation - set \isa{cs}. - - \item \verb|Code.get_datatype_of_constr|~\isa{thy}~\isa{const} - returns type constructor corresponding to - constructor \isa{const}; returns \isa{NONE} - if \isa{const} is no constructor. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Auxiliary% -} -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\ - \indexml{Code\_Unit.head\_eqn}\verb|Code_Unit.head_eqn: theory -> thm -> string * ((string * sort) list * typ)| \\ - \indexml{Code\_Unit.rewrite\_eqn}\verb|Code_Unit.rewrite_eqn: simpset -> thm -> thm| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Code_Unit.read_const|~\isa{thy}~\isa{s} - reads a constant as a concrete term expression \isa{s}. - - \item \verb|Code_Unit.head_eqn|~\isa{thy}~\isa{thm} - extracts the constant and its type from a code equation \isa{thm}. - - \item \verb|Code_Unit.rewrite_eqn|~\isa{ss}~\isa{thm} - rewrites a code equation \isa{thm} with a simpset \isa{ss}; - only arguments and right hand side are rewritten, - not the head of the code equation. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Implementing code generator applications% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Implementing code generator applications on top - of the framework set out so far usually not only - involves using those primitive interfaces - but also storing code-dependent data and various - other things.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Data depending on the theory's executable content% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Due to incrementality of code generation, changes in the - theory's executable content have to be propagated in a - certain fashion. Additionally, such changes may occur - not only during theory extension but also during theory - merge, which is a little bit nasty from an implementation - point of view. The framework provides a solution - to this technical challenge by providing a functorial - data slot \verb|CodeDataFun|; on instantiation - of this functor, the following types and operations - are required: - - \medskip - \begin{tabular}{l} - \isa{type\ T} \\ - \isa{val\ empty{\isacharcolon}\ T} \\ - \isa{val\ purge{\isacharcolon}\ theory\ {\isasymrightarrow}\ string\ list\ option\ {\isasymrightarrow}\ T\ {\isasymrightarrow}\ T} - \end{tabular} - - \begin{description} - - \item \isa{T} the type of data to store. - - \item \isa{empty} initial (empty) data. - - \item \isa{purge}~\isa{thy}~\isa{consts} propagates changes in executable content; - \isa{consts} indicates the kind - of change: \verb|NONE| stands for a fundamental change - which invalidates any existing code, \isa{SOME\ consts} - hints that executable content for constants \isa{consts} - has changed. - - \end{description} - - \noindent An instance of \verb|CodeDataFun| provides the following - interface: - - \medskip - \begin{tabular}{l} - \isa{get{\isacharcolon}\ theory\ {\isasymrightarrow}\ T} \\ - \isa{change{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ T{\isacharparenright}\ {\isasymrightarrow}\ T} \\ - \isa{change{\isacharunderscore}yield{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T{\isacharparenright}\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T} - \end{tabular} - - \begin{description} - - \item \isa{get} retrieval of the current data. - - \item \isa{change} update of current data (cached!) - by giving a continuation. - - \item \isa{change{\isacharunderscore}yield} update with side result. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\bigskip - - \emph{Happy proving, happy hacking!}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/document/Program.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/Program.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1250 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Program}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Program\isanewline -\isakeyword{imports}\ Introduction\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupsection{Turning Theories into Programs \label{sec:program}% -} -\isamarkuptrue% -% -\isamarkupsubsection{The \isa{Isabelle{\isacharslash}HOL} default setup% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -We have already seen how by default equations stemming from - \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}/\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}/\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}} - statements are used for code generation. This default behaviour - can be changed, e.g. by providing different code equations. - All kinds of customisation shown in this section is \emph{safe} - in the sense that the user does not have to worry about - correctness -- all programs generatable that way are partially - correct.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Selecting code equations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Coming back to our introductory example, we - could provide an alternative code equations for \isa{dequeue} - explicitly:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline -\ \ \ \ \ \ \ else\ dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}cases\ xs{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\ {\isacharparenleft}cases\ {\isachardoublequoteopen}rev\ xs{\isachardoublequoteclose}{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The annotation \isa{{\isacharbrackleft}code{\isacharbrackright}} is an \isa{Isar} - \isa{attribute} which states that the given theorems should be - considered as code equations for a \isa{fun} statement -- - the corresponding constant is determined syntactically. The resulting code:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\ -\hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\ -\hspace*{0pt}dequeue (AQueue xs []) =\\ -\hspace*{0pt} ~(if nulla xs then (Nothing,~AQueue [] [])\\ -\hspace*{0pt} ~~~else dequeue (AQueue [] (rev xs)));% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent You may note that the equality test \isa{xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}} has been - replaced by the predicate \isa{null\ xs}. This is due to the default - setup in the \qn{preprocessor} to be discussed further below (\secref{sec:preproc}). - - Changing the default constructor set of datatypes is also - possible. See \secref{sec:datatypes} for an example. - - As told in \secref{sec:concept}, code generation is based - on a structured collection of code theorems. - For explorative purpose, this collection - may be inspected using the \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} command:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{code{\isacharunderscore}thms}\isamarkupfalse% -\ dequeue% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent prints a table with \emph{all} code equations - for \isa{dequeue}, including - \emph{all} code equations those equations depend - on recursively. - - Similarly, the \hyperlink{command.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}} command shows a graph - visualising dependencies between code equations.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{\isa{class} and \isa{instantiation}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Concerning type classes and code generation, let us examine an example - from abstract algebra:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{class}\isamarkupfalse% -\ semigroup\ {\isacharequal}\isanewline -\ \ \isakeyword{fixes}\ mult\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymotimes}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{class}\isamarkupfalse% -\ monoid\ {\isacharequal}\ semigroup\ {\isacharplus}\isanewline -\ \ \isakeyword{fixes}\ neutral\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\ \ \isakeyword{assumes}\ neutl{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ neutr{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ nat\ {\isacharcolon}{\isacharcolon}\ monoid\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{primrec}\isamarkupfalse% -\ mult{\isacharunderscore}nat\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}{\isadigit{0}}\ {\isasymotimes}\ n\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}Suc\ m\ {\isasymotimes}\ n\ {\isacharequal}\ n\ {\isacharplus}\ m\ {\isasymotimes}\ n{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ neutral{\isacharunderscore}nat\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ add{\isacharunderscore}mult{\isacharunderscore}distrib{\isacharcolon}\isanewline -\ \ \isakeyword{fixes}\ n\ m\ q\ {\isacharcolon}{\isacharcolon}\ nat\isanewline -\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isacharparenleft}n\ {\isacharplus}\ m{\isacharparenright}\ {\isasymotimes}\ q\ {\isacharequal}\ n\ {\isasymotimes}\ q\ {\isacharplus}\ m\ {\isasymotimes}\ q{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ n{\isacharparenright}\ simp{\isacharunderscore}all\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ m\ n\ q\ {\isacharcolon}{\isacharcolon}\ nat\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymotimes}\ q\ {\isacharequal}\ m\ {\isasymotimes}\ {\isacharparenleft}n\ {\isasymotimes}\ q{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ m{\isacharparenright}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ add{\isacharunderscore}mult{\isacharunderscore}distrib{\isacharparenright}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharparenright}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}m\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ m{\isacharparenright}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharparenright}\isanewline -\isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent We define the natural operation of the natural numbers - on monoids:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{primrec}\isamarkupfalse% -\ {\isacharparenleft}\isakeyword{in}\ monoid{\isacharparenright}\ pow\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}pow\ {\isadigit{0}}\ a\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}pow\ {\isacharparenleft}Suc\ n{\isacharparenright}\ a\ {\isacharequal}\ a\ {\isasymotimes}\ pow\ n\ a{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This we use to define the discrete exponentiation function:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{definition}\isamarkupfalse% -\ bexp\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}bexp\ n\ {\isacharequal}\ pow\ n\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The corresponding code:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}module Example where {\char123}\\ -\hspace*{0pt}\\ -\hspace*{0pt}\\ -\hspace*{0pt}data Nat = Suc Nat | Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}class Semigroup a where {\char123}\\ -\hspace*{0pt} ~mult ::~a -> a -> a;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}class (Semigroup a) => Monoid a where {\char123}\\ -\hspace*{0pt} ~neutral ::~a;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}pow ::~forall a.~(Monoid a) => Nat -> a -> a;\\ -\hspace*{0pt}pow Zero{\char95}nat a = neutral;\\ -\hspace*{0pt}pow (Suc n) a = mult a (pow n a);\\ -\hspace*{0pt}\\ -\hspace*{0pt}plus{\char95}nat ::~Nat -> Nat -> Nat;\\ -\hspace*{0pt}plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n);\\ -\hspace*{0pt}plus{\char95}nat Zero{\char95}nat n = n;\\ -\hspace*{0pt}\\ -\hspace*{0pt}neutral{\char95}nat ::~Nat;\\ -\hspace*{0pt}neutral{\char95}nat = Suc Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}mult{\char95}nat ::~Nat -> Nat -> Nat;\\ -\hspace*{0pt}mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat;\\ -\hspace*{0pt}mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\ -\hspace*{0pt}\\ -\hspace*{0pt}instance Semigroup Nat where {\char123}\\ -\hspace*{0pt} ~mult = mult{\char95}nat;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}instance Monoid Nat where {\char123}\\ -\hspace*{0pt} ~neutral = neutral{\char95}nat;\\ -\hspace*{0pt}{\char125};\\ -\hspace*{0pt}\\ -\hspace*{0pt}bexp ::~Nat -> Nat;\\ -\hspace*{0pt}bexp n = pow n (Suc (Suc Zero{\char95}nat));\\ -\hspace*{0pt}\\ -\hspace*{0pt}{\char125}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This is a convenient place to show how explicit dictionary construction - manifests in generated code (here, the same example in \isa{SML}):% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype nat = Suc of nat | Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a semigroup = {\char123}mult :~'a -> 'a -> 'a{\char125};\\ -\hspace*{0pt}fun mult (A{\char95}:'a semigroup) = {\char35}mult A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a monoid = {\char123}Program{\char95}{\char95}semigroup{\char95}monoid :~'a semigroup,~neutral :~'a{\char125};\\ -\hspace*{0pt}fun semigroup{\char95}monoid (A{\char95}:'a monoid) = {\char35}Program{\char95}{\char95}semigroup{\char95}monoid A{\char95};\\ -\hspace*{0pt}fun neutral (A{\char95}:'a monoid) = {\char35}neutral A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun pow A{\char95}~Zero{\char95}nat a = neutral A{\char95}\\ -\hspace*{0pt} ~| pow A{\char95}~(Suc n) a = mult (semigroup{\char95}monoid A{\char95}) a (pow A{\char95}~n a);\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n)\\ -\hspace*{0pt} ~| plus{\char95}nat Zero{\char95}nat n = n;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val neutral{\char95}nat :~nat = Suc Zero{\char95}nat\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat\\ -\hspace*{0pt} ~| mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\ -\hspace*{0pt}\\ -\hspace*{0pt}val semigroup{\char95}nat = {\char123}mult = mult{\char95}nat{\char125}~:~nat semigroup;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val monoid{\char95}nat =\\ -\hspace*{0pt} ~{\char123}Program{\char95}{\char95}semigroup{\char95}monoid = semigroup{\char95}nat,~neutral = neutral{\char95}nat{\char125}~:\\ -\hspace*{0pt} ~nat monoid;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun bexp n = pow monoid{\char95}nat n (Suc (Suc Zero{\char95}nat));\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Note the parameters with trailing underscore (\verb|A_|) - which are the dictionary parameters.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{The preprocessor \label{sec:preproc}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Before selected function theorems are turned into abstract - code, a chain of definitional transformation steps is carried - out: \emph{preprocessing}. In essence, the preprocessor - consists of two components: a \emph{simpset} and \emph{function transformers}. - - The \emph{simpset} allows to employ the full generality of the Isabelle - simplifier. Due to the interpretation of theorems - as code equations, rewrites are applied to the right - hand side and the arguments of the left hand side of an - equation, but never to the constant heading the left hand side. - An important special case are \emph{inline theorems} which may be - declared and undeclared using the - \emph{code inline} or \emph{code inline del} attribute respectively. - - Some common applications:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{itemize} -% -\begin{isamarkuptext}% -\item replacing non-executable constructs by executable ones:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}x\ {\isasymin}\ set\ xs\ {\isasymlongleftrightarrow}\ x\ mem\ xs{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\item eliminating superfluous constants:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\item replacing executable but inconvenient constructs:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongleftrightarrow}\ List{\isachardot}null\ xs{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\end{itemize} -% -\begin{isamarkuptext}% -\noindent \emph{Function transformers} provide a very general interface, - transforming a list of function theorems to another - list of function theorems, provided that neither the heading - constant nor its type change. The \isa{{\isadigit{0}}} / \isa{Suc} - pattern elimination implemented in - theory \isa{Efficient{\isacharunderscore}Nat} (see \secref{eff_nat}) uses this - interface. - - \noindent The current setup of the preprocessor may be inspected using - the \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} command. - \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} provides a convenient - mechanism to inspect the impact of a preprocessor setup - on code equations. - - \begin{warn} - The attribute \emph{code unfold} - associated with the \isa{SML\ code\ generator} also applies to - the \isa{generic\ code\ generator}: - \emph{code unfold} implies \emph{code inline}. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Datatypes \label{sec:datatypes}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Conceptually, any datatype is spanned by a set of - \emph{constructors} of type \isa{{\isasymtau}\ {\isacharequal}\ {\isasymdots}\ {\isasymRightarrow}\ {\isasymkappa}\ {\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymdots}\ {\isasymalpha}\isactrlisub n} where \isa{{\isacharbraceleft}{\isasymalpha}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlisub n{\isacharbraceright}} is exactly the set of \emph{all} type variables in - \isa{{\isasymtau}}. The HOL datatype package by default registers any new - datatype in the table of datatypes, which may be inspected using the - \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} command. - - In some cases, it is appropriate to alter or extend this table. As - an example, we will develop an alternative representation of the - queue example given in \secref{sec:intro}. The amortised - representation is convenient for generating code but exposes its - \qt{implementation} details, which may be cumbersome when proving - theorems about it. Therefore, here a simple, straightforward - representation of queues:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{datatype}\isamarkupfalse% -\ {\isacharprime}a\ queue\ {\isacharequal}\ Queue\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ empty\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ Queue\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{primrec}\isamarkupfalse% -\ enqueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}Queue\ xs{\isacharparenright}\ {\isacharequal}\ Queue\ {\isacharparenleft}xs\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{fun}\isamarkupfalse% -\ dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ option\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}Queue\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}None{\isacharcomma}\ Queue\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}Queue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ x{\isacharcomma}\ Queue\ xs{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This we can use directly for proving; for executing, - we provide an alternative characterisation:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{definition}\isamarkupfalse% -\ AQueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}AQueue\ xs\ ys\ {\isacharequal}\ Queue\ {\isacharparenleft}ys\ {\isacharat}\ rev\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% -\ AQueue% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Here we define a \qt{constructor} \isa{Program{\isachardot}AQueue} which - is defined in terms of \isa{Queue} and interprets its arguments - according to what the \emph{content} of an amortised queue is supposed - to be. Equipped with this, we are able to prove the following equations - for our primitive queue operations which \qt{implement} the simple - queues in an amortised fashion:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ empty{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{unfolding}\isamarkupfalse% -\ AQueue{\isacharunderscore}def\ empty{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ enqueue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ AQueue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ ys{\isachardoublequoteclose}\isanewline -\ \ \isacommand{unfolding}\isamarkupfalse% -\ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ dequeue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline -\ \ \ \ else\ dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{unfolding}\isamarkupfalse% -\ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp{\isacharunderscore}all% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent For completeness, we provide a substitute for the - \isa{case} combinator on queues:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ aqueue{\isacharunderscore}case{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}aqueue{\isacharunderscore}case\ {\isacharequal}\ queue{\isacharunderscore}case{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ aqueue{\isacharunderscore}case\ {\isacharbrackleft}code{\isacharcomma}\ code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}queue{\isacharunderscore}case\ {\isacharequal}\ aqueue{\isacharunderscore}case{\isachardoublequoteclose}\isanewline -\ \ \isacommand{unfolding}\isamarkupfalse% -\ aqueue{\isacharunderscore}case{\isacharunderscore}def\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ case{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}aqueue{\isacharunderscore}case\ f\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ f\ {\isacharparenleft}ys\ {\isacharat}\ rev\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{unfolding}\isamarkupfalse% -\ aqueue{\isacharunderscore}case{\isacharunderscore}def\ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The resulting code looks as expected:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun foldl f a [] = a\\ -\hspace*{0pt} ~| foldl f a (x ::~xs) = foldl f (f a x) xs;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun rev xs = foldl (fn xsa => fn x => x ::~xsa) [] xs;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun null [] = true\\ -\hspace*{0pt} ~| null (x ::~xs) = false;\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype 'a queue = AQueue of 'a list * 'a list;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val empty :~'a queue = AQueue ([],~[])\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun dequeue (AQueue (xs,~y ::~ys)) = (SOME y,~AQueue (xs,~ys))\\ -\hspace*{0pt} ~| dequeue (AQueue (xs,~[])) =\\ -\hspace*{0pt} ~~~(if null xs then (NONE,~AQueue ([],~[]))\\ -\hspace*{0pt} ~~~~~else dequeue (AQueue ([],~rev xs)));\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun enqueue x (AQueue (xs,~ys)) = AQueue (x ::~xs,~ys);\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent From this example, it can be glimpsed that using own - constructor sets is a little delicate since it changes the set of - valid patterns for values of that type. Without going into much - detail, here some practical hints: - - \begin{itemize} - - \item When changing the constructor set for datatypes, take care - to provide an alternative for the \isa{case} combinator - (e.g.~by replacing it using the preprocessor). - - \item Values in the target language need not to be normalised -- - different values in the target language may represent the same - value in the logic. - - \item Usually, a good methodology to deal with the subtleties of - pattern matching is to see the type as an abstract type: provide - a set of operations which operate on the concrete representation - of the type, and derive further operations by combinations of - these primitive ones, without relying on a particular - representation. - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Equality and wellsortedness% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Surely you have already noticed how equality is treated - by the code generator:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{primrec}\isamarkupfalse% -\ collect{\isacharunderscore}duplicates\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharparenleft}z{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ z\ {\isasymin}\ set\ xs\isanewline -\ \ \ \ \ \ then\ if\ z\ {\isasymin}\ set\ ys\isanewline -\ \ \ \ \ \ \ \ then\ collect{\isacharunderscore}duplicates\ xs\ ys\ zs\isanewline -\ \ \ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ xs\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs\isanewline -\ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ {\isacharparenleft}z{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs{\isacharparenright}{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent The membership test during preprocessing is rewritten, - resulting in \isa{op\ mem}, which itself - performs an explicit equality check.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\ -\hspace*{0pt}fun eq (A{\char95}:'a eq) = {\char35}eq A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun eqop A{\char95}~a b = eq A{\char95}~a b;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun member A{\char95}~x [] = false\\ -\hspace*{0pt} ~| member A{\char95}~x (y ::~ys) = eqop A{\char95}~x y orelse member A{\char95}~x ys;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun collect{\char95}duplicates A{\char95}~xs ys [] = xs\\ -\hspace*{0pt} ~| collect{\char95}duplicates A{\char95}~xs ys (z ::~zs) =\\ -\hspace*{0pt} ~~~(if member A{\char95}~z xs\\ -\hspace*{0pt} ~~~~~then (if member A{\char95}~z ys then collect{\char95}duplicates A{\char95}~xs ys zs\\ -\hspace*{0pt} ~~~~~~~~~~~~else collect{\char95}duplicates A{\char95}~xs (z ::~ys) zs)\\ -\hspace*{0pt} ~~~~~else collect{\char95}duplicates A{\char95}~(z ::~xs) (z ::~ys) zs);\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Obviously, polymorphic equality is implemented the Haskell - way using a type class. How is this achieved? HOL introduces - an explicit class \isa{eq} with a corresponding operation - \isa{eq{\isacharunderscore}class{\isachardot}eq} such that \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharequal}\ op\ {\isacharequal}}. - The preprocessing framework does the rest by propagating the - \isa{eq} constraints through all dependent code equations. - For datatypes, instances of \isa{eq} are implicitly derived - when possible. For other types, you may instantiate \isa{eq} - manually like any other type class. - - Though this \isa{eq} class is designed to get rarely in - the way, a subtlety - enters the stage when definitions of overloaded constants - are dependent on operational equality. For example, let - us define a lexicographic ordering on tuples - (also see theory \hyperlink{theory.Product-ord}{\mbox{\isa{Product{\isacharunderscore}ord}}}):% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{instantiation}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}order{\isacharcomma}\ order{\isacharparenright}\ order\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}x\ {\isasymle}\ y\ {\isasymlongleftrightarrow}\ fst\ x\ {\isacharless}\ fst\ y\ {\isasymor}\ fst\ x\ {\isacharequal}\ fst\ y\ {\isasymand}\ snd\ x\ {\isasymle}\ snd\ y{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}x\ {\isacharless}\ y\ {\isasymlongleftrightarrow}\ fst\ x\ {\isacharless}\ fst\ y\ {\isasymor}\ fst\ x\ {\isacharequal}\ fst\ y\ {\isasymand}\ snd\ x\ {\isacharless}\ snd\ y{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -\ {\isacharparenleft}auto\ simp{\isacharcolon}\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}prod{\isacharunderscore}def\ intro{\isacharcolon}\ order{\isacharunderscore}less{\isacharunderscore}trans{\isacharparenright}\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ order{\isacharunderscore}prod\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}order{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}order{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Then code generation will fail. Why? The definition - of \isa{op\ {\isasymle}} depends on equality on both arguments, - which are polymorphic and impose an additional \isa{eq} - class constraint, which the preprocessor does not propagate - (for technical reasons). - - The solution is to add \isa{eq} explicitly to the first sort arguments in the - code theorems:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ order{\isacharunderscore}prod{\isacharunderscore}code\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}order{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}order{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Then code generation succeeds:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\ -\hspace*{0pt}fun eq (A{\char95}:'a eq) = {\char35}eq A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a ord = {\char123}less{\char95}eq :~'a -> 'a -> bool,~less :~'a -> 'a -> bool{\char125};\\ -\hspace*{0pt}fun less{\char95}eq (A{\char95}:'a ord) = {\char35}less{\char95}eq A{\char95};\\ -\hspace*{0pt}fun less (A{\char95}:'a ord) = {\char35}less A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun eqop A{\char95}~a b = eq A{\char95}~a b;\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a preorder = {\char123}Orderings{\char95}{\char95}ord{\char95}preorder :~'a ord{\char125};\\ -\hspace*{0pt}fun ord{\char95}preorder (A{\char95}:'a preorder) = {\char35}Orderings{\char95}{\char95}ord{\char95}preorder A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a order = {\char123}Orderings{\char95}{\char95}preorder{\char95}order :~'a preorder{\char125};\\ -\hspace*{0pt}fun preorder{\char95}order (A{\char95}:'a order) = {\char35}Orderings{\char95}{\char95}preorder{\char95}order A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\ -\hspace*{0pt} ~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\ -\hspace*{0pt} ~~~eqop A1{\char95}~x1 x2 andalso\\ -\hspace*{0pt} ~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2\\ -\hspace*{0pt} ~| less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\ -\hspace*{0pt} ~~~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\ -\hspace*{0pt} ~~~~~eqop A1{\char95}~x1 x2 andalso\\ -\hspace*{0pt} ~~~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2;\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -In some cases, the automatically derived code equations - for equality on a particular type may not be appropriate. - As example, watch the following datatype representing - monomorphic parametric types (where type constructors - are referred to by natural numbers):% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{datatype}\isamarkupfalse% -\ monotype\ {\isacharequal}\ Mono\ nat\ {\isachardoublequoteopen}monotype\ list{\isachardoublequoteclose}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Then code generation for SML would fail with a message - that the generated code contains illegal mutual dependencies: - the theorem \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}{\isacharparenright}\ {\isasymequiv}\ eq{\isacharunderscore}class{\isachardot}eq\ tyco{\isadigit{1}}\ tyco{\isadigit{2}}\ {\isasymand}\ eq{\isacharunderscore}class{\isachardot}eq\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}} already requires the - instance \isa{monotype\ {\isasymColon}\ eq}, which itself requires - \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}{\isacharparenright}\ {\isasymequiv}\ eq{\isacharunderscore}class{\isachardot}eq\ tyco{\isadigit{1}}\ tyco{\isadigit{2}}\ {\isasymand}\ eq{\isacharunderscore}class{\isachardot}eq\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}}; Haskell has no problem with mutually - recursive \isa{instance} and \isa{function} definitions, - but the SML serialiser does not support this. - - In such cases, you have to provide your own equality equations - involving auxiliary constants. In our case, - \isa{list{\isacharunderscore}all{\isadigit{2}}} can do the job:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ monotype{\isacharunderscore}eq{\isacharunderscore}list{\isacharunderscore}all{\isadigit{2}}\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ eq{\isacharunderscore}class{\isachardot}eq\ tyco{\isadigit{1}}\ tyco{\isadigit{2}}\ {\isasymand}\ list{\isacharunderscore}all{\isadigit{2}}\ eq{\isacharunderscore}class{\isachardot}eq\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq\ list{\isacharunderscore}all{\isadigit{2}}{\isacharunderscore}eq\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent does not depend on instance \isa{monotype\ {\isasymColon}\ eq}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype nat = Suc of nat | Zero{\char95}nat;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun null [] = true\\ -\hspace*{0pt} ~| null (x ::~xs) = false;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun eq{\char95}nat (Suc a) Zero{\char95}nat = false\\ -\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat (Suc a) = false\\ -\hspace*{0pt} ~| eq{\char95}nat (Suc nat) (Suc nat') = eq{\char95}nat nat nat'\\ -\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat Zero{\char95}nat = true;\\ -\hspace*{0pt}\\ -\hspace*{0pt}datatype monotype = Mono of nat * monotype list;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun list{\char95}all2 p (x ::~xs) (y ::~ys) = p x y andalso list{\char95}all2 p xs ys\\ -\hspace*{0pt} ~| list{\char95}all2 p xs [] = null xs\\ -\hspace*{0pt} ~| list{\char95}all2 p [] ys = null ys;\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun eq{\char95}monotype (Mono (tyco1,~typargs1)) (Mono (tyco2,~typargs2)) =\\ -\hspace*{0pt} ~eq{\char95}nat tyco1 tyco2 andalso list{\char95}all2 eq{\char95}monotype typargs1 typargs2;\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\isamarkupsubsection{Explicit partiality% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Partiality usually enters the game by partial patterns, as - in the following example, again for amortised queues:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{definition}\isamarkupfalse% -\ strict{\isacharunderscore}dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ q\ {\isacharequal}\ {\isacharparenleft}case\ dequeue\ q\isanewline -\ \ \ \ of\ {\isacharparenleft}Some\ x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ strict{\isacharunderscore}dequeue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ {\isacharparenleft}case\ rev\ xs\ of\ y\ {\isacharhash}\ ys\ {\isasymRightarrow}\ {\isacharparenleft}y{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ ys{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ strict{\isacharunderscore}dequeue{\isacharunderscore}def\ dequeue{\isacharunderscore}AQueue\ split{\isacharcolon}\ list{\isachardot}splits{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent In the corresponding code, there is no equation - for the pattern \isa{Program{\isachardot}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}strict{\char95}dequeue ::~forall a.~Queue a -> (a,~Queue a);\\ -\hspace*{0pt}strict{\char95}dequeue (AQueue xs []) =\\ -\hspace*{0pt} ~let {\char123}\\ -\hspace*{0pt} ~~~(y :~ys) = rev xs;\\ -\hspace*{0pt} ~{\char125}~in (y,~AQueue [] ys);\\ -\hspace*{0pt}strict{\char95}dequeue (AQueue xs (y :~ys)) = (y,~AQueue xs ys);% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent In some cases it is desirable to have this - pseudo-\qt{partiality} more explicitly, e.g.~as follows:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{axiomatization}\isamarkupfalse% -\ empty{\isacharunderscore}queue\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ q\ {\isacharequal}\ {\isacharparenleft}case\ dequeue\ q\ of\ {\isacharparenleft}Some\ x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isacharbar}\ {\isacharunderscore}\ {\isasymRightarrow}\ empty{\isacharunderscore}queue{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ strict{\isacharunderscore}dequeue{\isacharprime}{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ empty{\isacharunderscore}queue\isanewline -\ \ \ \ \ else\ strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ {\isacharparenleft}y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ strict{\isacharunderscore}dequeue{\isacharprime}{\isacharunderscore}def\ dequeue{\isacharunderscore}AQueue\ split{\isacharcolon}\ list{\isachardot}splits{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -Observe that on the right hand side of the definition of \isa{strict{\isacharunderscore}dequeue{\isacharprime}} the constant \isa{empty{\isacharunderscore}queue} occurs - which is unspecified. - - Normally, if constants without any code equations occur in a - program, the code generator complains (since in most cases this is - not what the user expects). But such constants can also be thought - of as function definitions with no equations which always fail, - since there is never a successful pattern match on the left hand - side. In order to categorise a constant into that category - explicitly, use \hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{code{\isacharunderscore}abort}\isamarkupfalse% -\ empty{\isacharunderscore}queue% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Then the code generator will just insert an error or - exception at the appropriate position:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}empty{\char95}queue ::~forall a.~a;\\ -\hspace*{0pt}empty{\char95}queue = error {\char34}empty{\char95}queue{\char34};\\ -\hspace*{0pt}\\ -\hspace*{0pt}strict{\char95}dequeue' ::~forall a.~Queue a -> (a,~Queue a);\\ -\hspace*{0pt}strict{\char95}dequeue' (AQueue xs (y :~ys)) = (y,~AQueue xs ys);\\ -\hspace*{0pt}strict{\char95}dequeue' (AQueue xs []) =\\ -\hspace*{0pt} ~(if nulla xs then empty{\char95}queue\\ -\hspace*{0pt} ~~~else strict{\char95}dequeue' (AQueue [] (rev xs)));% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent This feature however is rarely needed in practice. - Note also that the \isa{HOL} default setup already declares - \isa{undefined} as \hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}, which is most - likely to be used in such situations.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\ \end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/Codegen.hs --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/Codegen.hs Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,23 +0,0 @@ -module Codegen where { - -import qualified Nat; - -class Null a where { - nulla :: a; -}; - -heada :: forall a. (Codegen.Null a) => [a] -> a; -heada (x : xs) = x; -heada [] = Codegen.nulla; - -null_option :: forall a. Maybe a; -null_option = Nothing; - -instance Codegen.Null (Maybe a) where { - nulla = Codegen.null_option; -}; - -dummy :: Maybe Nat.Nat; -dummy = Codegen.heada [Just (Nat.Suc Nat.Zero_nat), Nothing]; - -} diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/Example.hs --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/Example.hs Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,33 +0,0 @@ -{-# OPTIONS_GHC -fglasgow-exts #-} - -module Example where { - - -foldla :: forall a b. (a -> b -> a) -> a -> [b] -> a; -foldla f a [] = a; -foldla f a (x : xs) = foldla f (f a x) xs; - -rev :: forall a. [a] -> [a]; -rev xs = foldla (\ xsa x -> x : xsa) [] xs; - -list_case :: forall t a. t -> (a -> [a] -> t) -> [a] -> t; -list_case f1 f2 (a : list) = f2 a list; -list_case f1 f2 [] = f1; - -data Queue a = AQueue [a] [a]; - -empty :: forall a. Queue a; -empty = AQueue [] []; - -dequeue :: forall a. Queue a -> (Maybe a, Queue a); -dequeue (AQueue [] []) = (Nothing, AQueue [] []); -dequeue (AQueue xs (y : ys)) = (Just y, AQueue xs ys); -dequeue (AQueue (v : va) []) = - let { - (y : ys) = rev (v : va); - } in (Just y, AQueue [] ys); - -enqueue :: forall a. a -> Queue a -> Queue a; -enqueue x (AQueue xs ys) = AQueue (x : xs) ys; - -} diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/arbitrary.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/arbitrary.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,9 +0,0 @@ -structure Codegen = -struct - -val arbitrary_option : 'a option = NONE; - -fun dummy_option [] = arbitrary_option - | dummy_option (x :: xs) = SOME x; - -end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/bool_infix.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/bool_infix.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,19 +0,0 @@ -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -fun less_nat m (Suc n) = less_eq_nat m n - | less_nat n Zero_nat = false -and less_eq_nat (Suc m) n = less_nat m n - | less_eq_nat Zero_nat n = true; - -end; (*struct Nat*) - -structure Codegen = -struct - -fun in_interval (k, l) n = - Nat.less_eq_nat k n andalso Nat.less_eq_nat n l; - -end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/bool_literal.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/bool_literal.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,31 +0,0 @@ -structure HOL = -struct - -datatype boola = False | True; - -fun anda x True = x - | anda x False = False - | anda True x = x - | anda False x = False; - -end; (*struct HOL*) - -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -fun less_nat m (Suc n) = less_eq_nat m n - | less_nat n Zero_nat = HOL.False -and less_eq_nat (Suc m) n = less_nat m n - | less_eq_nat Zero_nat n = HOL.True; - -end; (*struct Nat*) - -structure Codegen = -struct - -fun in_interval (k, l) n = - HOL.anda (Nat.less_eq_nat k n) (Nat.less_eq_nat n l); - -end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/bool_mlbool.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/bool_mlbool.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,19 +0,0 @@ -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -fun less_nat m (Suc n) = less_eq_nat m n - | less_nat n Zero_nat = false -and less_eq_nat (Suc m) n = less_nat m n - | less_eq_nat Zero_nat n = true; - -end; (*struct Nat*) - -structure Codegen = -struct - -fun in_interval (k, l) n = - (Nat.less_eq_nat k n) andalso (Nat.less_eq_nat n l); - -end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/class.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/class.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,24 +0,0 @@ -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -end; (*struct Nat*) - -structure Codegen = -struct - -type 'a null = {null : 'a}; -fun null (A_:'a null) = #null A_; - -fun head A_ (x :: xs) = x - | head A_ [] = null A_; - -val null_option : 'a option = NONE; - -fun null_optiona () = {null = null_option} : ('a option) null; - -val dummy : Nat.nat option = - head (null_optiona ()) [SOME (Nat.Suc Nat.Zero_nat), NONE]; - -end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/class.ocaml --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/class.ocaml Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,24 +0,0 @@ -module Nat = -struct - -type nat = Suc of nat | Zero_nat;; - -end;; (*struct Nat*) - -module Codegen = -struct - -type 'a null = {null : 'a};; -let null _A = _A.null;; - -let rec head _A = function x :: xs -> x - | [] -> null _A;; - -let rec null_option = None;; - -let null_optiona () = ({null = null_option} : ('a option) null);; - -let rec dummy - = head (null_optiona ()) [Some (Nat.Suc Nat.Zero_nat); None];; - -end;; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/collect_duplicates.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/collect_duplicates.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,30 +0,0 @@ -structure HOL = -struct - -type 'a eq = {eq : 'a -> 'a -> bool}; -fun eq (A_:'a eq) = #eq A_; - -fun eqop A_ a = eq A_ a; - -end; (*struct HOL*) - -structure List = -struct - -fun member A_ x (y :: ys) = - (if HOL.eqop A_ y x then true else member A_ x ys) - | member A_ x [] = false; - -end; (*struct List*) - -structure Codegen = -struct - -fun collect_duplicates A_ xs ys (z :: zs) = - (if List.member A_ z xs - then (if List.member A_ z ys then collect_duplicates A_ xs ys zs - else collect_duplicates A_ xs (z :: ys) zs) - else collect_duplicates A_ (z :: xs) (z :: ys) zs) - | collect_duplicates A_ xs ys [] = xs; - -end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/dirty_set.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/dirty_set.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,102 +0,0 @@ -structure ROOT = -struct - -structure Nat = -struct - -datatype nat = Zero_nat | Suc of nat; - -end; (*struct Nat*) - -structure Integer = -struct - -datatype bit = B0 | B1; - -datatype int = Pls | Min | Bit of int * bit | Number_of_int of int; - -fun pred (Bit (k, B0)) = Bit (pred k, B1) - | pred (Bit (k, B1)) = Bit (k, B0) - | pred Min = Bit (Min, B0) - | pred Pls = Min; - -fun uminus_int (Number_of_int w) = Number_of_int (uminus_int w) - | uminus_int (Bit (k, B0)) = Bit (uminus_int k, B0) - | uminus_int (Bit (k, B1)) = Bit (pred (uminus_int k), B1) - | uminus_int Min = Bit (Pls, B1) - | uminus_int Pls = Pls; - -fun succ (Bit (k, B0)) = Bit (k, B1) - | succ (Bit (k, B1)) = Bit (succ k, B0) - | succ Min = Pls - | succ Pls = Bit (Pls, B1); - -fun plus_int (Number_of_int v) (Number_of_int w) = - Number_of_int (plus_int v w) - | plus_int k Min = pred k - | plus_int k Pls = k - | plus_int (Bit (k, B1)) (Bit (l, B1)) = Bit (plus_int k (succ l), B0) - | plus_int (Bit (k, B1)) (Bit (l, B0)) = Bit (plus_int k l, B1) - | plus_int (Bit (k, B0)) (Bit (l, b)) = Bit (plus_int k l, b) - | plus_int Min k = pred k - | plus_int Pls k = k; - -fun minus_int (Number_of_int v) (Number_of_int w) = - Number_of_int (plus_int v (uminus_int w)) - | minus_int z w = plus_int z (uminus_int w); - -fun less_eq_int (Number_of_int k) (Number_of_int l) = less_eq_int k l - | less_eq_int (Bit (k1, B1)) (Bit (k2, B0)) = less_int k1 k2 - | less_eq_int (Bit (k1, v)) (Bit (k2, B1)) = less_eq_int k1 k2 - | less_eq_int (Bit (k1, B0)) (Bit (k2, v)) = less_eq_int k1 k2 - | less_eq_int (Bit (k, v)) Min = less_eq_int k Min - | less_eq_int (Bit (k, B1)) Pls = less_int k Pls - | less_eq_int (Bit (k, B0)) Pls = less_eq_int k Pls - | less_eq_int Min (Bit (k, B1)) = less_eq_int Min k - | less_eq_int Min (Bit (k, B0)) = less_int Min k - | less_eq_int Min Min = true - | less_eq_int Min Pls = true - | less_eq_int Pls (Bit (k, v)) = less_eq_int Pls k - | less_eq_int Pls Min = false - | less_eq_int Pls Pls = true -and less_int (Number_of_int k) (Number_of_int l) = less_int k l - | less_int (Bit (k1, B0)) (Bit (k2, B1)) = less_eq_int k1 k2 - | less_int (Bit (k1, B1)) (Bit (k2, v)) = less_int k1 k2 - | less_int (Bit (k1, v)) (Bit (k2, B0)) = less_int k1 k2 - | less_int (Bit (k, B1)) Min = less_int k Min - | less_int (Bit (k, B0)) Min = less_eq_int k Min - | less_int (Bit (k, v)) Pls = less_int k Pls - | less_int Min (Bit (k, v)) = less_int Min k - | less_int Min Min = false - | less_int Min Pls = true - | less_int Pls (Bit (k, B1)) = less_eq_int Pls k - | less_int Pls (Bit (k, B0)) = less_int Pls k - | less_int Pls Min = false - | less_int Pls Pls = false; - -fun nat_aux n i = - (if less_eq_int i (Number_of_int Pls) then n - else nat_aux (Nat.Suc n) - (minus_int i (Number_of_int (Bit (Pls, B1))))); - -fun nat i = nat_aux Nat.Zero_nat i; - -end; (*struct Integer*) - -structure Codegen = -struct - -val dummy_set : (Nat.nat -> Nat.nat) list = Nat.Suc :: []; - -val foobar_set : Nat.nat list = - Nat.Zero_nat :: - (Nat.Suc Nat.Zero_nat :: - (Integer.nat - (Integer.Number_of_int - (Integer.Bit - (Integer.Bit (Integer.Pls, Integer.B1), Integer.B0))) - :: [])); - -end; (*struct Codegen*) - -end; (*struct ROOT*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/example.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/example.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,27 +0,0 @@ -structure Example = -struct - -fun foldl f a [] = a - | foldl f a (x :: xs) = foldl f (f a x) xs; - -fun rev xs = foldl (fn xsa => fn x => x :: xsa) [] xs; - -fun list_case f1 f2 (a :: lista) = f2 a lista - | list_case f1 f2 [] = f1; - -datatype 'a queue = AQueue of 'a list * 'a list; - -val empty : 'a queue = AQueue ([], []) - -fun dequeue (AQueue ([], [])) = (NONE, AQueue ([], [])) - | dequeue (AQueue (xs, y :: ys)) = (SOME y, AQueue (xs, ys)) - | dequeue (AQueue (v :: va, [])) = - let - val y :: ys = rev (v :: va); - in - (SOME y, AQueue ([], ys)) - end; - -fun enqueue x (AQueue (xs, ys)) = AQueue (x :: xs, ys); - -end; (*struct Example*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/fac.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/fac.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,22 +0,0 @@ -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -val one_nat : nat = Suc Zero_nat; - -fun plus_nat (Suc m) n = plus_nat m (Suc n) - | plus_nat Zero_nat n = n; - -fun times_nat (Suc m) n = plus_nat n (times_nat m n) - | times_nat Zero_nat n = Zero_nat; - -end; (*struct Nat*) - -structure Codegen = -struct - -fun fac (Nat.Suc n) = Nat.times_nat (Nat.Suc n) (fac n) - | fac Nat.Zero_nat = Nat.one_nat; - -end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/integers.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/integers.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,59 +0,0 @@ -structure ROOT = -struct - -structure Integer = -struct - -datatype bit = B0 | B1; - -datatype int = Pls | Min | Bit of int * bit | Number_of_int of int; - -fun pred (Bit (k, B0)) = Bit (pred k, B1) - | pred (Bit (k, B1)) = Bit (k, B0) - | pred Min = Bit (Min, B0) - | pred Pls = Min; - -fun succ (Bit (k, B0)) = Bit (k, B1) - | succ (Bit (k, B1)) = Bit (succ k, B0) - | succ Min = Pls - | succ Pls = Bit (Pls, B1); - -fun plus_int (Number_of_int v) (Number_of_int w) = - Number_of_int (plus_int v w) - | plus_int k Min = pred k - | plus_int k Pls = k - | plus_int (Bit (k, B1)) (Bit (l, B1)) = Bit (plus_int k (succ l), B0) - | plus_int (Bit (k, B1)) (Bit (l, B0)) = Bit (plus_int k l, B1) - | plus_int (Bit (k, B0)) (Bit (l, b)) = Bit (plus_int k l, b) - | plus_int Min k = pred k - | plus_int Pls k = k; - -fun uminus_int (Number_of_int w) = Number_of_int (uminus_int w) - | uminus_int (Bit (k, B0)) = Bit (uminus_int k, B0) - | uminus_int (Bit (k, B1)) = Bit (pred (uminus_int k), B1) - | uminus_int Min = Bit (Pls, B1) - | uminus_int Pls = Pls; - -fun times_int (Number_of_int v) (Number_of_int w) = - Number_of_int (times_int v w) - | times_int (Bit (k, B0)) l = Bit (times_int k l, B0) - | times_int (Bit (k, B1)) l = plus_int (Bit (times_int k l, B0)) l - | times_int Min k = uminus_int k - | times_int Pls w = Pls; - -end; (*struct Integer*) - -structure Codegen = -struct - -fun double_inc k = - Integer.plus_int - (Integer.times_int - (Integer.Number_of_int - (Integer.Bit (Integer.Bit (Integer.Pls, Integer.B1), Integer.B0))) - k) - (Integer.Number_of_int (Integer.Bit (Integer.Pls, Integer.B1))); - -end; (*struct Codegen*) - -end; (*struct ROOT*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/lexicographic.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/lexicographic.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,19 +0,0 @@ -structure HOL = -struct - -type 'a eq = {eq : 'a -> 'a -> bool}; -fun eq (A_:'a eq) = #eq A_; - -type 'a ord = {less_eq : 'a -> 'a -> bool, less : 'a -> 'a -> bool}; -fun less_eq (A_:'a ord) = #less_eq A_; -fun less (A_:'a ord) = #less A_; - -end; (*struct HOL*) - -structure Codegen = -struct - -fun less_eq (A1_, A2_) B_ (x1, y1) (x2, y2) = - HOL.less A2_ x1 x2 orelse HOL.eq A1_ x1 x2 andalso HOL.less_eq B_ y1 y2; - -end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/lookup.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/lookup.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,13 +0,0 @@ -structure ROOT = -struct - -structure Codegen = -struct - -fun lookup ((k, v) :: xs) l = - (if ((k : string) = l) then SOME v else lookup xs l) - | lookup [] l = NONE; - -end; (*struct Codegen*) - -end; (*struct ROOT*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/monotype.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/monotype.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,34 +0,0 @@ -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -fun eq_nat (Suc a) Zero_nat = false - | eq_nat Zero_nat (Suc a) = false - | eq_nat (Suc nat) (Suc nat') = eq_nat nat nat' - | eq_nat Zero_nat Zero_nat = true; - -end; (*struct Nat*) - -structure List = -struct - -fun null (x :: xs) = false - | null [] = true; - -fun list_all2 p (x :: xs) (y :: ys) = p x y andalso list_all2 p xs ys - | list_all2 p xs [] = null xs - | list_all2 p [] ys = null ys; - -end; (*struct List*) - -structure Codegen = -struct - -datatype monotype = Mono of Nat.nat * monotype list; - -fun eq_monotype (Mono (tyco1, typargs1)) (Mono (tyco2, typargs2)) = - Nat.eq_nat tyco1 tyco2 andalso - List.list_all2 eq_monotype typargs1 typargs2; - -end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/nat_binary.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/nat_binary.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,17 +0,0 @@ -structure Nat = -struct - -datatype nat = Dig1 of nat | Dig0 of nat | One_nat | Zero_nat; - -fun plus_nat (Dig1 m) (Dig1 n) = Dig0 (plus_nat (plus_nat m n) One_nat) - | plus_nat (Dig1 m) (Dig0 n) = Dig1 (plus_nat m n) - | plus_nat (Dig0 m) (Dig1 n) = Dig1 (plus_nat m n) - | plus_nat (Dig0 m) (Dig0 n) = Dig0 (plus_nat m n) - | plus_nat (Dig1 m) One_nat = Dig0 (plus_nat m One_nat) - | plus_nat One_nat (Dig1 n) = Dig0 (plus_nat n One_nat) - | plus_nat (Dig0 m) One_nat = Dig1 m - | plus_nat One_nat (Dig0 n) = Dig1 n - | plus_nat m Zero_nat = m - | plus_nat Zero_nat n = n; - -end; (*struct Nat*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/pick1.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/pick1.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,44 +0,0 @@ -structure HOL = -struct - -fun leta s f = f s; - -end; (*struct HOL*) - -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -fun less_nat m (Suc n) = less_eq_nat m n - | less_nat n Zero_nat = false -and less_eq_nat (Suc m) n = less_nat m n - | less_eq_nat Zero_nat n = true; - -fun minus_nat (Suc m) (Suc n) = minus_nat m n - | minus_nat Zero_nat n = Zero_nat - | minus_nat m Zero_nat = m; - -end; (*struct Nat*) - -structure Product_Type = -struct - -fun split f (a, b) = f a b; - -end; (*struct Product_Type*) - -structure Codegen = -struct - -fun pick ((k, v) :: xs) n = - (if Nat.less_nat n k then v else pick xs (Nat.minus_nat n k)) - | pick (x :: xs) n = - let - val a = x; - val (k, v) = a; - in - (if Nat.less_nat n k then v else pick xs (Nat.minus_nat n k)) - end; - -end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/Thy/examples/tree.ML --- a/doc-src/IsarAdvanced/Codegen/Thy/examples/tree.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,95 +0,0 @@ -structure HOL = -struct - -type 'a eq = {eq : 'a -> 'a -> bool}; -fun eq (A_:'a eq) = #eq A_; - -type 'a ord = {less_eq : 'a -> 'a -> bool, less : 'a -> 'a -> bool}; -fun less_eq (A_:'a ord) = #less_eq A_; -fun less (A_:'a ord) = #less A_; - -fun eqop A_ a = eq A_ a; - -end; (*struct HOL*) - -structure Orderings = -struct - -type 'a preorder = {Orderings__ord_preorder : 'a HOL.ord}; -fun ord_preorder (A_:'a preorder) = #Orderings__ord_preorder A_; - -type 'a order = {Orderings__preorder_order : 'a preorder}; -fun preorder_order (A_:'a order) = #Orderings__preorder_order A_; - -type 'a linorder = {Orderings__order_linorder : 'a order}; -fun order_linorder (A_:'a linorder) = #Orderings__order_linorder A_; - -end; (*struct Orderings*) - -structure Nat = -struct - -datatype nat = Suc of nat | Zero_nat; - -fun eq_nat (Suc a) Zero_nat = false - | eq_nat Zero_nat (Suc a) = false - | eq_nat (Suc nat) (Suc nat') = eq_nat nat nat' - | eq_nat Zero_nat Zero_nat = true; - -val eq_nata = {eq = eq_nat} : nat HOL.eq; - -fun less_nat m (Suc n) = less_eq_nat m n - | less_nat n Zero_nat = false -and less_eq_nat (Suc m) n = less_nat m n - | less_eq_nat Zero_nat n = true; - -val ord_nat = {less_eq = less_eq_nat, less = less_nat} : nat HOL.ord; - -val preorder_nat = {Orderings__ord_preorder = ord_nat} : - nat Orderings.preorder; - -val order_nat = {Orderings__preorder_order = preorder_nat} : - nat Orderings.order; - -val linorder_nat = {Orderings__order_linorder = order_nat} : - nat Orderings.linorder; - -end; (*struct Nat*) - -structure Codegen = -struct - -datatype ('a, 'b) searchtree = - Branch of ('a, 'b) searchtree * 'a * ('a, 'b) searchtree | - Leaf of 'a * 'b; - -fun update (A1_, A2_) (it, entry) (Leaf (key, vala)) = - (if HOL.eqop A1_ it key then Leaf (key, entry) - else (if HOL.less_eq - ((Orderings.ord_preorder o Orderings.preorder_order o - Orderings.order_linorder) - A2_) - it key - then Branch (Leaf (it, entry), it, Leaf (key, vala)) - else Branch (Leaf (key, vala), it, Leaf (it, entry)))) - | update (A1_, A2_) (it, entry) (Branch (t1, key, t2)) = - (if HOL.less_eq - ((Orderings.ord_preorder o Orderings.preorder_order o - Orderings.order_linorder) - A2_) - it key - then Branch (update (A1_, A2_) (it, entry) t1, key, t2) - else Branch (t1, key, update (A1_, A2_) (it, entry) t2)); - -val example : (Nat.nat, (Nat.nat list)) searchtree = - update (Nat.eq_nata, Nat.linorder_nat) - (Nat.Suc (Nat.Suc (Nat.Suc (Nat.Suc Nat.Zero_nat))), - [Nat.Suc (Nat.Suc Nat.Zero_nat), Nat.Suc (Nat.Suc Nat.Zero_nat)]) - (update (Nat.eq_nata, Nat.linorder_nat) - (Nat.Suc (Nat.Suc (Nat.Suc Nat.Zero_nat)), - [Nat.Suc (Nat.Suc (Nat.Suc Nat.Zero_nat))]) - (update (Nat.eq_nata, Nat.linorder_nat) - (Nat.Suc (Nat.Suc Nat.Zero_nat), [Nat.Suc (Nat.Suc Nat.Zero_nat)]) - (Leaf (Nat.Suc Nat.Zero_nat, [])))); - -end; (*struct Codegen*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/codegen.tex --- a/doc-src/IsarAdvanced/Codegen/codegen.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,53 +0,0 @@ - -\documentclass[12pt,a4paper,fleqn]{report} -\usepackage{latexsym,graphicx} -\usepackage[refpage]{nomencl} -\usepackage{../../iman,../../extra,../../isar,../../proof} -\usepackage{../../isabelle,../../isabellesym} -\usepackage{style} -\usepackage{pgf} -\usepackage{pgflibraryshapes} -\usepackage{tikz} -\usepackage{../../pdfsetup} - -\hyphenation{Isabelle} -\hyphenation{Isar} -\isadroptag{theory} - -\title{\includegraphics[scale=0.5]{isabelle_isar} - \\[4ex] Code generation from Isabelle/HOL theories} -\author{\emph{Florian Haftmann}} - -\begin{document} - -\maketitle - -\begin{abstract} - This tutorial gives an introduction to a generic code generator framework in Isabelle - for generating executable code in functional programming languages from logical - specifications in Isabelle/HOL. -\end{abstract} - -\thispagestyle{empty}\clearpage - -\pagenumbering{roman} -\clearfirst - -\input{Thy/document/Introduction.tex} -\input{Thy/document/Program.tex} -\input{Thy/document/Adaption.tex} -\input{Thy/document/Further.tex} -\input{Thy/document/ML.tex} - -\begingroup -\bibliographystyle{plain} \small\raggedright\frenchspacing -\bibliography{../../manual} -\endgroup - -\end{document} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: t -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/codegen_process.pdf Binary file doc-src/IsarAdvanced/Codegen/codegen_process.pdf has changed diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/codegen_process.ps --- a/doc-src/IsarAdvanced/Codegen/codegen_process.ps Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,586 +0,0 @@ -%!PS-Adobe-2.0 -%%Creator: dot version 2.2 (Mon Sep 12 23:33:36 UTC 2005) -%%For: (haftmann) Florian Haftmann -%%Title: _anonymous_0 -%%Pages: (atend) -%%BoundingBox: 35 35 451 291 -%%EndComments -save -%%BeginProlog -/DotDict 200 dict def -DotDict begin - -/setupLatin1 { -mark -/EncodingVector 256 array def - EncodingVector 0 - -ISOLatin1Encoding 0 255 getinterval putinterval - -EncodingVector - dup 306 /AE - dup 301 /Aacute - dup 302 /Acircumflex - dup 304 /Adieresis - dup 300 /Agrave - dup 305 /Aring - dup 303 /Atilde - dup 307 /Ccedilla - dup 311 /Eacute - dup 312 /Ecircumflex - dup 313 /Edieresis - dup 310 /Egrave - dup 315 /Iacute - dup 316 /Icircumflex - dup 317 /Idieresis - dup 314 /Igrave - dup 334 /Udieresis - dup 335 /Yacute - dup 376 /thorn - dup 337 /germandbls - dup 341 /aacute - dup 342 /acircumflex - dup 344 /adieresis - dup 346 /ae - dup 340 /agrave - dup 345 /aring - dup 347 /ccedilla - dup 351 /eacute - dup 352 /ecircumflex - dup 353 /edieresis - dup 350 /egrave - dup 355 /iacute - dup 356 /icircumflex - dup 357 /idieresis - dup 354 /igrave - dup 360 /dcroat - dup 361 /ntilde - dup 363 /oacute - dup 364 /ocircumflex - dup 366 /odieresis - dup 362 /ograve - dup 365 /otilde - dup 370 /oslash - dup 372 /uacute - dup 373 /ucircumflex - dup 374 /udieresis - dup 371 /ugrave - dup 375 /yacute - dup 377 /ydieresis - -% Set up ISO Latin 1 character encoding -/starnetISO { - dup dup findfont dup length dict begin - { 1 index /FID ne { def }{ pop pop } ifelse - } forall - /Encoding EncodingVector def - currentdict end definefont -} def -/Times-Roman starnetISO def -/Times-Italic starnetISO def -/Times-Bold starnetISO def -/Times-BoldItalic starnetISO def -/Helvetica starnetISO def -/Helvetica-Oblique starnetISO def -/Helvetica-Bold starnetISO def -/Helvetica-BoldOblique starnetISO def -/Courier starnetISO def -/Courier-Oblique starnetISO def -/Courier-Bold starnetISO def -/Courier-BoldOblique starnetISO def -cleartomark -} bind def - -%%BeginResource: procset graphviz 0 0 -/coord-font-family /Times-Roman def -/default-font-family /Times-Roman def -/coordfont coord-font-family findfont 8 scalefont def - -/InvScaleFactor 1.0 def -/set_scale { - dup 1 exch div /InvScaleFactor exch def - dup scale -} bind def - -% styles -/solid { [] 0 setdash } bind def -/dashed { [9 InvScaleFactor mul dup ] 0 setdash } bind def -/dotted { [1 InvScaleFactor mul 6 InvScaleFactor mul] 0 setdash } bind def -/invis {/fill {newpath} def /stroke {newpath} def /show {pop newpath} def} bind def -/bold { 2 setlinewidth } bind def -/filled { } bind def -/unfilled { } bind def -/rounded { } bind def -/diagonals { } bind def - -% hooks for setting color -/nodecolor { sethsbcolor } bind def -/edgecolor { sethsbcolor } bind def -/graphcolor { sethsbcolor } bind def -/nopcolor {pop pop pop} bind def - -/beginpage { % i j npages - /npages exch def - /j exch def - /i exch def - /str 10 string def - npages 1 gt { - gsave - coordfont setfont - 0 0 moveto - (\() show i str cvs show (,) show j str cvs show (\)) show - grestore - } if -} bind def - -/set_font { - findfont exch - scalefont setfont -} def - -% draw aligned label in bounding box aligned to current point -/alignedtext { % width adj text - /text exch def - /adj exch def - /width exch def - gsave - width 0 gt { - text stringwidth pop adj mul 0 rmoveto - } if - [] 0 setdash - text show - grestore -} def - -/boxprim { % xcorner ycorner xsize ysize - 4 2 roll - moveto - 2 copy - exch 0 rlineto - 0 exch rlineto - pop neg 0 rlineto - closepath -} bind def - -/ellipse_path { - /ry exch def - /rx exch def - /y exch def - /x exch def - matrix currentmatrix - newpath - x y translate - rx ry scale - 0 0 1 0 360 arc - setmatrix -} bind def - -/endpage { showpage } bind def -/showpage { } def - -/layercolorseq - [ % layer color sequence - darkest to lightest - [0 0 0] - [.2 .8 .8] - [.4 .8 .8] - [.6 .8 .8] - [.8 .8 .8] - ] -def - -/layerlen layercolorseq length def - -/setlayer {/maxlayer exch def /curlayer exch def - layercolorseq curlayer 1 sub layerlen mod get - aload pop sethsbcolor - /nodecolor {nopcolor} def - /edgecolor {nopcolor} def - /graphcolor {nopcolor} def -} bind def - -/onlayer { curlayer ne {invis} if } def - -/onlayers { - /myupper exch def - /mylower exch def - curlayer mylower lt - curlayer myupper gt - or - {invis} if -} def - -/curlayer 0 def - -%%EndResource -%%EndProlog -%%BeginSetup -14 default-font-family set_font -1 setmiterlimit -% /arrowlength 10 def -% /arrowwidth 5 def - -% make sure pdfmark is harmless for PS-interpreters other than Distiller -/pdfmark where {pop} {userdict /pdfmark /cleartomark load put} ifelse -% make '<<' and '>>' safe on PS Level 1 devices -/languagelevel where {pop languagelevel}{1} ifelse -2 lt { - userdict (<<) cvn ([) cvn load put - userdict (>>) cvn ([) cvn load put -} if - -%%EndSetup -%%Page: 1 1 -%%PageBoundingBox: 36 36 451 291 -%%PageOrientation: Portrait -gsave -35 35 416 256 boxprim clip newpath -36 36 translate -0 0 1 beginpage -0 0 translate 0 rotate -[ /CropBox [36 36 451 291] /PAGES pdfmark -0.000 0.000 0.000 graphcolor -14.00 /Times-Roman set_font - -% theory -gsave 10 dict begin -newpath 93 254 moveto -1 254 lineto -1 214 lineto -93 214 lineto -closepath -stroke -gsave 10 dict begin -8 237 moveto -(Isabelle/HOL) -[4.56 5.52 6.24 6.96 6.24 3.84 3.84 6.24 3.84 10.08 10.08 8.64] -xshow -16 221 moveto -(Isar theory) -[4.56 5.52 6.24 4.56 3.6 4.08 6.96 6.24 6.96 4.8 6.96] -xshow -end grestore -end grestore - -% selection -gsave 10 dict begin -183 234 38 18 ellipse_path -stroke -gsave 10 dict begin -158 229 moveto -(selection) -[5.52 6.24 3.84 6.24 6.24 3.84 3.84 6.96 6.96] -xshow -end grestore -end grestore - -% theory -> selection -newpath 94 234 moveto -107 234 121 234 135 234 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 135 238 moveto -145 234 lineto -135 231 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 135 238 moveto -145 234 lineto -135 231 lineto -closepath -stroke -end grestore - -% sml -gsave 10 dict begin -newpath 74 144 moveto -20 144 lineto -20 108 lineto -74 108 lineto -closepath -stroke -gsave 10 dict begin -32 121 moveto -(SML) -[7.68 12.48 8.64] -xshow -end grestore -end grestore - -% other -gsave 10 dict begin -gsave 10 dict begin -41 67 moveto -(...) -[3.6 3.6 3.6] -xshow -end grestore -end grestore - -% haskell -gsave 10 dict begin -newpath 77 36 moveto -17 36 lineto -17 0 lineto -77 0 lineto -closepath -stroke -gsave 10 dict begin -25 13 moveto -(Haskell) -[10.08 6.24 5.52 6.72 6.24 3.84 3.84] -xshow -end grestore -end grestore - -% preprocessing -gsave 10 dict begin -183 180 52 18 ellipse_path -stroke -gsave 10 dict begin -143 175 moveto -(preprocessing) -[6.96 4.56 6.24 6.96 4.56 6.96 6.24 6.24 5.52 5.52 3.84 6.96 6.96] -xshow -end grestore -end grestore - -% selection -> preprocessing -newpath 183 216 moveto -183 213 183 211 183 208 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 187 208 moveto -183 198 lineto -180 208 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 187 208 moveto -183 198 lineto -180 208 lineto -closepath -stroke -end grestore - -% def_eqn -gsave 10 dict begin -newpath 403 198 moveto -283 198 lineto -283 162 lineto -403 162 lineto -closepath -stroke -gsave 10 dict begin -291 175 moveto -(defining equations) -[6.96 6.24 4.8 3.84 6.96 3.84 6.96 6.96 3.6 6.24 6.72 6.96 6.24 3.84 3.84 6.96 6.96 5.52] -xshow -end grestore -end grestore - -% preprocessing -> def_eqn -newpath 236 180 moveto -248 180 260 180 273 180 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 273 184 moveto -283 180 lineto -273 177 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 273 184 moveto -283 180 lineto -273 177 lineto -closepath -stroke -end grestore - -% serialization -gsave 10 dict begin -183 72 47 18 ellipse_path -stroke -gsave 10 dict begin -148 67 moveto -(serialization) -[5.52 6.24 4.8 3.84 6.24 3.84 3.84 6.24 6.24 3.84 3.84 6.96 6.96] -xshow -end grestore -end grestore - -% serialization -> sml -newpath 150 85 moveto -129 93 104 103 83 111 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 82 108 moveto -74 115 lineto -85 114 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 82 108 moveto -74 115 lineto -85 114 lineto -closepath -stroke -end grestore - -% serialization -> other -gsave 10 dict begin -dotted -newpath 135 72 moveto -119 72 100 72 84 72 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 84 69 moveto -74 72 lineto -84 76 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 84 69 moveto -74 72 lineto -84 76 lineto -closepath -stroke -end grestore -end grestore - -% serialization -> haskell -newpath 150 59 moveto -131 51 107 42 86 34 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 88 31 moveto -77 30 lineto -85 37 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 88 31 moveto -77 30 lineto -85 37 lineto -closepath -stroke -end grestore - -% translation -gsave 10 dict begin -343 126 43 18 ellipse_path -stroke -gsave 10 dict begin -313 121 moveto -(translation) -[3.84 4.56 6.24 6.96 5.52 3.84 6.24 3.84 3.84 6.96 6.96] -xshow -end grestore -end grestore - -% def_eqn -> translation -newpath 343 162 moveto -343 159 343 157 343 154 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 347 154 moveto -343 144 lineto -340 154 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 347 154 moveto -343 144 lineto -340 154 lineto -closepath -stroke -end grestore - -% iml -gsave 10 dict begin -newpath 413 90 moveto -273 90 lineto -273 54 lineto -413 54 lineto -closepath -stroke -gsave 10 dict begin -280 67 moveto -(intermediate language) -[3.84 6.96 3.84 6.24 4.8 10.8 6.24 6.96 3.84 6.24 3.84 6.24 3.6 3.84 6.24 6.96 6.96 6.96 6.24 6.72 6.24] -xshow -end grestore -end grestore - -% translation -> iml -newpath 343 108 moveto -343 105 343 103 343 100 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 347 100 moveto -343 90 lineto -340 100 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 347 100 moveto -343 90 lineto -340 100 lineto -closepath -stroke -end grestore - -% iml -> serialization -newpath 272 72 moveto -262 72 251 72 241 72 curveto -stroke -gsave 10 dict begin -solid -1 setlinewidth -0.000 0.000 0.000 edgecolor -newpath 241 69 moveto -231 72 lineto -241 76 lineto -closepath -fill -0.000 0.000 0.000 edgecolor -newpath 241 69 moveto -231 72 lineto -241 76 lineto -closepath -stroke -end grestore -endpage -showpage -grestore -%%PageTrailer -%%EndPage: 1 -%%Trailer -%%Pages: 1 -end -restore -%%EOF diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Codegen/style.sty --- a/doc-src/IsarAdvanced/Codegen/style.sty Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,62 +0,0 @@ - -%% toc -\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} -\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} - -%% references -\newcommand{\secref}[1]{\S\ref{#1}} - -%% index -\newcommand{\indexml}[1]{\index{\emph{#1}|bold}} -\newcommand{\indexmltype}[1]{\index{\emph{#1} (type)|bold}} -\newcommand{\indexmlstructure}[1]{\index{\emph{#1} (structure)|bold}} -\newcommand{\indexmlfunctor}[1]{\index{\emph{#1} (functor)|bold}} - -%% logical markup -\newcommand{\strong}[1]{{\bfseries {#1}}} -\newcommand{\qn}[1]{\emph{#1}} - -%% typographic conventions -\newcommand{\qt}[1]{``{#1}''} - -%% verbatim text -\newcommand{\isatypewriter}{\fontsize{9pt}{0pt}\tt\renewcommand{\baselinestretch}{1}\setlength{\baselineskip}{9pt}} - -%% quoted segments -\makeatletter -\isakeeptag{quote} -\newenvironment{quotesegment}{\begin{quote}\isa@parindent\parindent\parindent0pt\isa@parskip\parskip\parskip0pt}{\end{quote}} -\renewcommand{\isatagquote}{\begin{quotesegment}} -\renewcommand{\endisatagquote}{\end{quotesegment}} -\makeatother - -\isakeeptag{quotett} -\renewcommand{\isatagquotett}{\begin{quotesegment}\isabellestyle{tt}\isastyle} -\renewcommand{\endisatagquotett}{\end{quotesegment}\isabellestyle{it}\isastyle} - -%% a trick -\newcommand{\isasymSML}{SML} - -%% presentation -\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} - -\pagestyle{headings} -\binperiod -\underscoreoff - -\renewcommand{\isadigit}[1]{\isamath{#1}} - -%% ml reference -\newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup} - -\isakeeptag{mlref} -\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small} -\renewcommand{\endisatagmlref}{\endgroup} - -\isabellestyle{it} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "implementation" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/IsaMakefile --- a/doc-src/IsarAdvanced/Functions/IsaMakefile Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,33 +0,0 @@ - -## targets - -default: Thy -images: -test: Thy - -all: images test - - -## global settings - -SRC = $(ISABELLE_HOME)/src -OUT = $(ISABELLE_OUTPUT) -LOG = $(OUT)/log - -USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -d false -C false -D document - - -## Thy - -THY = $(LOG)/HOL-Thy.gz - -Thy: $(THY) - -$(THY): Thy/ROOT.ML Thy/Functions.thy - @$(USEDIR) HOL Thy - - -## clean - -clean: - @rm -f $(THY) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/Makefile --- a/doc-src/IsarAdvanced/Functions/Makefile Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,38 +0,0 @@ -# -# $Id$ -# - -## targets - -default: dvi - - -## dependencies - -include ../Makefile.in - -NAME = functions - -FILES = $(NAME).tex Thy/document/Functions.tex intro.tex conclusion.tex \ - style.sty ../../iman.sty ../../extra.sty ../../isar.sty \ - ../../isabelle.sty ../../isabellesym.sty ../../pdfsetup.sty \ - ../../manual.bib ../../proof.sty - -dvi: $(NAME).dvi - -$(NAME).dvi: $(FILES) isabelle_isar.eps - $(LATEX) $(NAME) - $(BIBTEX) $(NAME) - $(LATEX) $(NAME) - $(LATEX) $(NAME) - -pdf: $(NAME).pdf - -$(NAME).pdf: $(FILES) isabelle_isar.pdf - $(PDFLATEX) $(NAME) - $(BIBTEX) $(NAME) - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) - $(FIXBOOKMARKS) $(NAME).out - $(PDFLATEX) $(NAME) - $(PDFLATEX) $(NAME) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/Thy/Functions.thy --- a/doc-src/IsarAdvanced/Functions/Thy/Functions.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1264 +0,0 @@ -(* Title: doc-src/IsarAdvanced/Functions/Thy/Fundefs.thy - Author: Alexander Krauss, TU Muenchen - -Tutorial for function definitions with the new "function" package. -*) - -theory Functions -imports Main -begin - -section {* Function Definitions for Dummies *} - -text {* - In most cases, defining a recursive function is just as simple as other definitions: -*} - -fun fib :: "nat \ nat" -where - "fib 0 = 1" -| "fib (Suc 0) = 1" -| "fib (Suc (Suc n)) = fib n + fib (Suc n)" - -text {* - The syntax is rather self-explanatory: We introduce a function by - giving its name, its type, - and a set of defining recursive equations. - If we leave out the type, the most general type will be - inferred, which can sometimes lead to surprises: Since both @{term - "1::nat"} and @{text "+"} are overloaded, we would end up - with @{text "fib :: nat \ 'a::{one,plus}"}. -*} - -text {* - The function always terminates, since its argument gets smaller in - every recursive call. - Since HOL is a logic of total functions, termination is a - fundamental requirement to prevent inconsistencies\footnote{From the - \qt{definition} @{text "f(n) = f(n) + 1"} we could prove - @{text "0 = 1"} by subtracting @{text "f(n)"} on both sides.}. - Isabelle tries to prove termination automatically when a definition - is made. In \S\ref{termination}, we will look at cases where this - fails and see what to do then. -*} - -subsection {* Pattern matching *} - -text {* \label{patmatch} - Like in functional programming, we can use pattern matching to - define functions. At the moment we will only consider \emph{constructor - patterns}, which only consist of datatype constructors and - variables. Furthermore, patterns must be linear, i.e.\ all variables - on the left hand side of an equation must be distinct. In - \S\ref{genpats} we discuss more general pattern matching. - - If patterns overlap, the order of the equations is taken into - account. The following function inserts a fixed element between any - two elements of a list: -*} - -fun sep :: "'a \ 'a list \ 'a list" -where - "sep a (x#y#xs) = x # a # sep a (y # xs)" -| "sep a xs = xs" - -text {* - Overlapping patterns are interpreted as \qt{increments} to what is - already there: The second equation is only meant for the cases where - the first one does not match. Consequently, Isabelle replaces it - internally by the remaining cases, making the patterns disjoint: -*} - -thm sep.simps - -text {* @{thm [display] sep.simps[no_vars]} *} - -text {* - \noindent The equations from function definitions are automatically used in - simplification: -*} - -lemma "sep 0 [1, 2, 3] = [1, 0, 2, 0, 3]" -by simp - -subsection {* Induction *} - -text {* - - Isabelle provides customized induction rules for recursive - functions. These rules follow the recursive structure of the - definition. Here is the rule @{text sep.induct} arising from the - above definition of @{const sep}: - - @{thm [display] sep.induct} - - We have a step case for list with at least two elements, and two - base cases for the zero- and the one-element list. Here is a simple - proof about @{const sep} and @{const map} -*} - -lemma "map f (sep x ys) = sep (f x) (map f ys)" -apply (induct x ys rule: sep.induct) - -txt {* - We get three cases, like in the definition. - - @{subgoals [display]} -*} - -apply auto -done -text {* - - With the \cmd{fun} command, you can define about 80\% of the - functions that occur in practice. The rest of this tutorial explains - the remaining 20\%. -*} - - -section {* fun vs.\ function *} - -text {* - The \cmd{fun} command provides a - convenient shorthand notation for simple function definitions. In - this mode, Isabelle tries to solve all the necessary proof obligations - automatically. If any proof fails, the definition is - rejected. This can either mean that the definition is indeed faulty, - or that the default proof procedures are just not smart enough (or - rather: not designed) to handle the definition. - - By expanding the abbreviation to the more verbose \cmd{function} command, these proof obligations become visible and can be analyzed or - solved manually. The expansion from \cmd{fun} to \cmd{function} is as follows: - -\end{isamarkuptext} - - -\[\left[\;\begin{minipage}{0.25\textwidth}\vspace{6pt} -\cmd{fun} @{text "f :: \"}\\% -\cmd{where}\\% -\hspace*{2ex}{\it equations}\\% -\hspace*{2ex}\vdots\vspace*{6pt} -\end{minipage}\right] -\quad\equiv\quad -\left[\;\begin{minipage}{0.48\textwidth}\vspace{6pt} -\cmd{function} @{text "("}\cmd{sequential}@{text ") f :: \"}\\% -\cmd{where}\\% -\hspace*{2ex}{\it equations}\\% -\hspace*{2ex}\vdots\\% -\cmd{by} @{text "pat_completeness auto"}\\% -\cmd{termination by} @{text "lexicographic_order"}\vspace{6pt} -\end{minipage} -\right]\] - -\begin{isamarkuptext} - \vspace*{1em} - \noindent Some details have now become explicit: - - \begin{enumerate} - \item The \cmd{sequential} option enables the preprocessing of - pattern overlaps which we already saw. Without this option, the equations - must already be disjoint and complete. The automatic completion only - works with constructor patterns. - - \item A function definition produces a proof obligation which - expresses completeness and compatibility of patterns (we talk about - this later). The combination of the methods @{text "pat_completeness"} and - @{text "auto"} is used to solve this proof obligation. - - \item A termination proof follows the definition, started by the - \cmd{termination} command. This will be explained in \S\ref{termination}. - \end{enumerate} - Whenever a \cmd{fun} command fails, it is usually a good idea to - expand the syntax to the more verbose \cmd{function} form, to see - what is actually going on. - *} - - -section {* Termination *} - -text {*\label{termination} - The method @{text "lexicographic_order"} is the default method for - termination proofs. It can prove termination of a - certain class of functions by searching for a suitable lexicographic - combination of size measures. Of course, not all functions have such - a simple termination argument. For them, we can specify the termination - relation manually. -*} - -subsection {* The {\tt relation} method *} -text{* - Consider the following function, which sums up natural numbers up to - @{text "N"}, using a counter @{text "i"}: -*} - -function sum :: "nat \ nat \ nat" -where - "sum i N = (if i > N then 0 else i + sum (Suc i) N)" -by pat_completeness auto - -text {* - \noindent The @{text "lexicographic_order"} method fails on this example, because none of the - arguments decreases in the recursive call, with respect to the standard size ordering. - To prove termination manually, we must provide a custom wellfounded relation. - - The termination argument for @{text "sum"} is based on the fact that - the \emph{difference} between @{text "i"} and @{text "N"} gets - smaller in every step, and that the recursion stops when @{text "i"} - is greater than @{text "N"}. Phrased differently, the expression - @{text "N + 1 - i"} always decreases. - - We can use this expression as a measure function suitable to prove termination. -*} - -termination sum -apply (relation "measure (\(i,N). N + 1 - i)") - -txt {* - The \cmd{termination} command sets up the termination goal for the - specified function @{text "sum"}. If the function name is omitted, it - implicitly refers to the last function definition. - - The @{text relation} method takes a relation of - type @{typ "('a \ 'a) set"}, where @{typ "'a"} is the argument type of - the function. If the function has multiple curried arguments, then - these are packed together into a tuple, as it happened in the above - example. - - The predefined function @{term[source] "measure :: ('a \ nat) \ ('a \ 'a) set"} constructs a - wellfounded relation from a mapping into the natural numbers (a - \emph{measure function}). - - After the invocation of @{text "relation"}, we must prove that (a) - the relation we supplied is wellfounded, and (b) that the arguments - of recursive calls indeed decrease with respect to the - relation: - - @{subgoals[display,indent=0]} - - These goals are all solved by @{text "auto"}: -*} - -apply auto -done - -text {* - Let us complicate the function a little, by adding some more - recursive calls: -*} - -function foo :: "nat \ nat \ nat" -where - "foo i N = (if i > N - then (if N = 0 then 0 else foo 0 (N - 1)) - else i + foo (Suc i) N)" -by pat_completeness auto - -text {* - When @{text "i"} has reached @{text "N"}, it starts at zero again - and @{text "N"} is decremented. - This corresponds to a nested - loop where one index counts up and the other down. Termination can - be proved using a lexicographic combination of two measures, namely - the value of @{text "N"} and the above difference. The @{const - "measures"} combinator generalizes @{text "measure"} by taking a - list of measure functions. -*} - -termination -by (relation "measures [\(i, N). N, \(i,N). N + 1 - i]") auto - -subsection {* How @{text "lexicographic_order"} works *} - -(*fun fails :: "nat \ nat list \ nat" -where - "fails a [] = a" -| "fails a (x#xs) = fails (x + a) (x # xs)" -*) - -text {* - To see how the automatic termination proofs work, let's look at an - example where it fails\footnote{For a detailed discussion of the - termination prover, see \cite{bulwahnKN07}}: - -\end{isamarkuptext} -\cmd{fun} @{text "fails :: \"nat \ nat list \ nat\""}\\% -\cmd{where}\\% -\hspace*{2ex}@{text "\"fails a [] = a\""}\\% -|\hspace*{1.5ex}@{text "\"fails a (x#xs) = fails (x + a) (x#xs)\""}\\ -\begin{isamarkuptext} - -\noindent Isabelle responds with the following error: - -\begin{isabelle} -*** Unfinished subgoals:\newline -*** (a, 1, <):\newline -*** \ 1.~@{text "\x. x = 0"}\newline -*** (a, 1, <=):\newline -*** \ 1.~False\newline -*** (a, 2, <):\newline -*** \ 1.~False\newline -*** Calls:\newline -*** a) @{text "(a, x # xs) -->> (x + a, x # xs)"}\newline -*** Measures:\newline -*** 1) @{text "\x. size (fst x)"}\newline -*** 2) @{text "\x. size (snd x)"}\newline -*** Result matrix:\newline -*** \ \ \ \ 1\ \ 2 \newline -*** a: ? <= \newline -*** Could not find lexicographic termination order.\newline -*** At command "fun".\newline -\end{isabelle} -*} - - -text {* - The key to this error message is the matrix at the bottom. The rows - of that matrix correspond to the different recursive calls (In our - case, there is just one). The columns are the function's arguments - (expressed through different measure functions, which map the - argument tuple to a natural number). - - The contents of the matrix summarize what is known about argument - descents: The second argument has a weak descent (@{text "<="}) at the - recursive call, and for the first argument nothing could be proved, - which is expressed by @{text "?"}. In general, there are the values - @{text "<"}, @{text "<="} and @{text "?"}. - - For the failed proof attempts, the unfinished subgoals are also - printed. Looking at these will often point to a missing lemma. - -% As a more real example, here is quicksort: -*} -(* -function qs :: "nat list \ nat list" -where - "qs [] = []" -| "qs (x#xs) = qs [y\xs. y < x] @ x # qs [y\xs. y \ x]" -by pat_completeness auto - -termination apply lexicographic_order - -text {* If we try @{text "lexicographic_order"} method, we get the - following error *} -termination by (lexicographic_order simp:l2) - -lemma l: "x \ y \ x < Suc y" by arith - -function - -*) - -section {* Mutual Recursion *} - -text {* - If two or more functions call one another mutually, they have to be defined - in one step. Here are @{text "even"} and @{text "odd"}: -*} - -function even :: "nat \ bool" - and odd :: "nat \ bool" -where - "even 0 = True" -| "odd 0 = False" -| "even (Suc n) = odd n" -| "odd (Suc n) = even n" -by pat_completeness auto - -text {* - To eliminate the mutual dependencies, Isabelle internally - creates a single function operating on the sum - type @{typ "nat + nat"}. Then, @{const even} and @{const odd} are - defined as projections. Consequently, termination has to be proved - simultaneously for both functions, by specifying a measure on the - sum type: -*} - -termination -by (relation "measure (\x. case x of Inl n \ n | Inr n \ n)") auto - -text {* - We could also have used @{text lexicographic_order}, which - supports mutual recursive termination proofs to a certain extent. -*} - -subsection {* Induction for mutual recursion *} - -text {* - - When functions are mutually recursive, proving properties about them - generally requires simultaneous induction. The induction rule @{text "even_odd.induct"} - generated from the above definition reflects this. - - Let us prove something about @{const even} and @{const odd}: -*} - -lemma even_odd_mod2: - "even n = (n mod 2 = 0)" - "odd n = (n mod 2 = 1)" - -txt {* - We apply simultaneous induction, specifying the induction variable - for both goals, separated by \cmd{and}: *} - -apply (induct n and n rule: even_odd.induct) - -txt {* - We get four subgoals, which correspond to the clauses in the - definition of @{const even} and @{const odd}: - @{subgoals[display,indent=0]} - Simplification solves the first two goals, leaving us with two - statements about the @{text "mod"} operation to prove: -*} - -apply simp_all - -txt {* - @{subgoals[display,indent=0]} - - \noindent These can be handled by Isabelle's arithmetic decision procedures. - -*} - -apply arith -apply arith -done - -text {* - In proofs like this, the simultaneous induction is really essential: - Even if we are just interested in one of the results, the other - one is necessary to strengthen the induction hypothesis. If we leave - out the statement about @{const odd} and just write @{term True} instead, - the same proof fails: -*} - -lemma failed_attempt: - "even n = (n mod 2 = 0)" - "True" -apply (induct n rule: even_odd.induct) - -txt {* - \noindent Now the third subgoal is a dead end, since we have no - useful induction hypothesis available: - - @{subgoals[display,indent=0]} -*} - -oops - -section {* General pattern matching *} -text{*\label{genpats} *} - -subsection {* Avoiding automatic pattern splitting *} - -text {* - - Up to now, we used pattern matching only on datatypes, and the - patterns were always disjoint and complete, and if they weren't, - they were made disjoint automatically like in the definition of - @{const "sep"} in \S\ref{patmatch}. - - This automatic splitting can significantly increase the number of - equations involved, and this is not always desirable. The following - example shows the problem: - - Suppose we are modeling incomplete knowledge about the world by a - three-valued datatype, which has values @{term "T"}, @{term "F"} - and @{term "X"} for true, false and uncertain propositions, respectively. -*} - -datatype P3 = T | F | X - -text {* \noindent Then the conjunction of such values can be defined as follows: *} - -fun And :: "P3 \ P3 \ P3" -where - "And T p = p" -| "And p T = p" -| "And p F = F" -| "And F p = F" -| "And X X = X" - - -text {* - This definition is useful, because the equations can directly be used - as simplification rules. But the patterns overlap: For example, - the expression @{term "And T T"} is matched by both the first and - the second equation. By default, Isabelle makes the patterns disjoint by - splitting them up, producing instances: -*} - -thm And.simps - -text {* - @{thm[indent=4] And.simps} - - \vspace*{1em} - \noindent There are several problems with this: - - \begin{enumerate} - \item If the datatype has many constructors, there can be an - explosion of equations. For @{const "And"}, we get seven instead of - five equations, which can be tolerated, but this is just a small - example. - - \item Since splitting makes the equations \qt{less general}, they - do not always match in rewriting. While the term @{term "And x F"} - can be simplified to @{term "F"} with the original equations, a - (manual) case split on @{term "x"} is now necessary. - - \item The splitting also concerns the induction rule @{text - "And.induct"}. Instead of five premises it now has seven, which - means that our induction proofs will have more cases. - - \item In general, it increases clarity if we get the same definition - back which we put in. - \end{enumerate} - - If we do not want the automatic splitting, we can switch it off by - leaving out the \cmd{sequential} option. However, we will have to - prove that our pattern matching is consistent\footnote{This prevents - us from defining something like @{term "f x = True"} and @{term "f x - = False"} simultaneously.}: -*} - -function And2 :: "P3 \ P3 \ P3" -where - "And2 T p = p" -| "And2 p T = p" -| "And2 p F = F" -| "And2 F p = F" -| "And2 X X = X" - -txt {* - \noindent Now let's look at the proof obligations generated by a - function definition. In this case, they are: - - @{subgoals[display,indent=0]}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em} - - The first subgoal expresses the completeness of the patterns. It has - the form of an elimination rule and states that every @{term x} of - the function's input type must match at least one of the patterns\footnote{Completeness could - be equivalently stated as a disjunction of existential statements: -@{term "(\p. x = (T, p)) \ (\p. x = (p, T)) \ (\p. x = (p, F)) \ - (\p. x = (F, p)) \ (x = (X, X))"}, and you can use the method @{text atomize_elim} to get that form instead.}. If the patterns just involve - datatypes, we can solve it with the @{text "pat_completeness"} - method: -*} - -apply pat_completeness - -txt {* - The remaining subgoals express \emph{pattern compatibility}. We do - allow that an input value matches multiple patterns, but in this - case, the result (i.e.~the right hand sides of the equations) must - also be equal. For each pair of two patterns, there is one such - subgoal. Usually this needs injectivity of the constructors, which - is used automatically by @{text "auto"}. -*} - -by auto - - -subsection {* Non-constructor patterns *} - -text {* - Most of Isabelle's basic types take the form of inductive datatypes, - and usually pattern matching works on the constructors of such types. - However, this need not be always the case, and the \cmd{function} - command handles other kind of patterns, too. - - One well-known instance of non-constructor patterns are - so-called \emph{$n+k$-patterns}, which are a little controversial in - the functional programming world. Here is the initial fibonacci - example with $n+k$-patterns: -*} - -function fib2 :: "nat \ nat" -where - "fib2 0 = 1" -| "fib2 1 = 1" -| "fib2 (n + 2) = fib2 n + fib2 (Suc n)" - -(*<*)ML_val "goals_limit := 1"(*>*) -txt {* - This kind of matching is again justified by the proof of pattern - completeness and compatibility. - The proof obligation for pattern completeness states that every natural number is - either @{term "0::nat"}, @{term "1::nat"} or @{term "n + - (2::nat)"}: - - @{subgoals[display,indent=0]} - - This is an arithmetic triviality, but unfortunately the - @{text arith} method cannot handle this specific form of an - elimination rule. However, we can use the method @{text - "atomize_elim"} to do an ad-hoc conversion to a disjunction of - existentials, which can then be solved by the arithmetic decision procedure. - Pattern compatibility and termination are automatic as usual. -*} -(*<*)ML_val "goals_limit := 10"(*>*) -apply atomize_elim -apply arith -apply auto -done -termination by lexicographic_order -text {* - We can stretch the notion of pattern matching even more. The - following function is not a sensible functional program, but a - perfectly valid mathematical definition: -*} - -function ev :: "nat \ bool" -where - "ev (2 * n) = True" -| "ev (2 * n + 1) = False" -apply atomize_elim -by arith+ -termination by (relation "{}") simp - -text {* - This general notion of pattern matching gives you a certain freedom - in writing down specifications. However, as always, such freedom should - be used with care: - - If we leave the area of constructor - patterns, we have effectively departed from the world of functional - programming. This means that it is no longer possible to use the - code generator, and expect it to generate ML code for our - definitions. Also, such a specification might not work very well together with - simplification. Your mileage may vary. -*} - - -subsection {* Conditional equations *} - -text {* - The function package also supports conditional equations, which are - similar to guards in a language like Haskell. Here is Euclid's - algorithm written with conditional patterns\footnote{Note that the - patterns are also overlapping in the base case}: -*} - -function gcd :: "nat \ nat \ nat" -where - "gcd x 0 = x" -| "gcd 0 y = y" -| "x < y \ gcd (Suc x) (Suc y) = gcd (Suc x) (y - x)" -| "\ x < y \ gcd (Suc x) (Suc y) = gcd (x - y) (Suc y)" -by (atomize_elim, auto, arith) -termination by lexicographic_order - -text {* - By now, you can probably guess what the proof obligations for the - pattern completeness and compatibility look like. - - Again, functions with conditional patterns are not supported by the - code generator. -*} - - -subsection {* Pattern matching on strings *} - -text {* - As strings (as lists of characters) are normal datatypes, pattern - matching on them is possible, but somewhat problematic. Consider the - following definition: - -\end{isamarkuptext} -\noindent\cmd{fun} @{text "check :: \"string \ bool\""}\\% -\cmd{where}\\% -\hspace*{2ex}@{text "\"check (''good'') = True\""}\\% -@{text "| \"check s = False\""} -\begin{isamarkuptext} - - \noindent An invocation of the above \cmd{fun} command does not - terminate. What is the problem? Strings are lists of characters, and - characters are a datatype with a lot of constructors. Splitting the - catch-all pattern thus leads to an explosion of cases, which cannot - be handled by Isabelle. - - There are two things we can do here. Either we write an explicit - @{text "if"} on the right hand side, or we can use conditional patterns: -*} - -function check :: "string \ bool" -where - "check (''good'') = True" -| "s \ ''good'' \ check s = False" -by auto - - -section {* Partiality *} - -text {* - In HOL, all functions are total. A function @{term "f"} applied to - @{term "x"} always has the value @{term "f x"}, and there is no notion - of undefinedness. - This is why we have to do termination - proofs when defining functions: The proof justifies that the - function can be defined by wellfounded recursion. - - However, the \cmd{function} package does support partiality to a - certain extent. Let's look at the following function which looks - for a zero of a given function f. -*} - -function (*<*)(domintros, tailrec)(*>*)findzero :: "(nat \ nat) \ nat \ nat" -where - "findzero f n = (if f n = 0 then n else findzero f (Suc n))" -by pat_completeness auto -(*<*)declare findzero.simps[simp del](*>*) - -text {* - \noindent Clearly, any attempt of a termination proof must fail. And without - that, we do not get the usual rules @{text "findzero.simps"} and - @{text "findzero.induct"}. So what was the definition good for at all? -*} - -subsection {* Domain predicates *} - -text {* - The trick is that Isabelle has not only defined the function @{const findzero}, but also - a predicate @{term "findzero_dom"} that characterizes the values where the function - terminates: the \emph{domain} of the function. If we treat a - partial function just as a total function with an additional domain - predicate, we can derive simplification and - induction rules as we do for total functions. They are guarded - by domain conditions and are called @{text psimps} and @{text - pinduct}: -*} - -text {* - \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.psimps}\end{minipage} - \hfill(@{text "findzero.psimps"}) - \vspace{1em} - - \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.pinduct}\end{minipage} - \hfill(@{text "findzero.pinduct"}) -*} - -text {* - Remember that all we - are doing here is use some tricks to make a total function appear - as if it was partial. We can still write the term @{term "findzero - (\x. 1) 0"} and like any other term of type @{typ nat} it is equal - to some natural number, although we might not be able to find out - which one. The function is \emph{underdefined}. - - But it is defined enough to prove something interesting about it. We - can prove that if @{term "findzero f n"} - terminates, it indeed returns a zero of @{term f}: -*} - -lemma findzero_zero: "findzero_dom (f, n) \ f (findzero f n) = 0" - -txt {* \noindent We apply induction as usual, but using the partial induction - rule: *} - -apply (induct f n rule: findzero.pinduct) - -txt {* \noindent This gives the following subgoals: - - @{subgoals[display,indent=0]} - - \noindent The hypothesis in our lemma was used to satisfy the first premise in - the induction rule. However, we also get @{term - "findzero_dom (f, n)"} as a local assumption in the induction step. This - allows to unfold @{term "findzero f n"} using the @{text psimps} - rule, and the rest is trivial. Since the @{text psimps} rules carry the - @{text "[simp]"} attribute by default, we just need a single step: - *} -apply simp -done - -text {* - Proofs about partial functions are often not harder than for total - functions. Fig.~\ref{findzero_isar} shows a slightly more - complicated proof written in Isar. It is verbose enough to show how - partiality comes into play: From the partial induction, we get an - additional domain condition hypothesis. Observe how this condition - is applied when calls to @{term findzero} are unfolded. -*} - -text_raw {* -\begin{figure} -\hrule\vspace{6pt} -\begin{minipage}{0.8\textwidth} -\isabellestyle{it} -\isastyle\isamarkuptrue -*} -lemma "\findzero_dom (f, n); x \ {n ..< findzero f n}\ \ f x \ 0" -proof (induct rule: findzero.pinduct) - fix f n assume dom: "findzero_dom (f, n)" - and IH: "\f n \ 0; x \ {Suc n ..< findzero f (Suc n)}\ \ f x \ 0" - and x_range: "x \ {n ..< findzero f n}" - have "f n \ 0" - proof - assume "f n = 0" - with dom have "findzero f n = n" by simp - with x_range show False by auto - qed - - from x_range have "x = n \ x \ {Suc n ..< findzero f n}" by auto - thus "f x \ 0" - proof - assume "x = n" - with `f n \ 0` show ?thesis by simp - next - assume "x \ {Suc n ..< findzero f n}" - with dom and `f n \ 0` have "x \ {Suc n ..< findzero f (Suc n)}" by simp - with IH and `f n \ 0` - show ?thesis by simp - qed -qed -text_raw {* -\isamarkupfalse\isabellestyle{tt} -\end{minipage}\vspace{6pt}\hrule -\caption{A proof about a partial function}\label{findzero_isar} -\end{figure} -*} - -subsection {* Partial termination proofs *} - -text {* - Now that we have proved some interesting properties about our - function, we should turn to the domain predicate and see if it is - actually true for some values. Otherwise we would have just proved - lemmas with @{term False} as a premise. - - Essentially, we need some introduction rules for @{text - findzero_dom}. The function package can prove such domain - introduction rules automatically. But since they are not used very - often (they are almost never needed if the function is total), this - functionality is disabled by default for efficiency reasons. So we have to go - back and ask for them explicitly by passing the @{text - "(domintros)"} option to the function package: - -\vspace{1ex} -\noindent\cmd{function} @{text "(domintros) findzero :: \"(nat \ nat) \ nat \ nat\""}\\% -\cmd{where}\isanewline% -\ \ \ldots\\ - - \noindent Now the package has proved an introduction rule for @{text findzero_dom}: -*} - -thm findzero.domintros - -text {* - @{thm[display] findzero.domintros} - - Domain introduction rules allow to show that a given value lies in the - domain of a function, if the arguments of all recursive calls - are in the domain as well. They allow to do a \qt{single step} in a - termination proof. Usually, you want to combine them with a suitable - induction principle. - - Since our function increases its argument at recursive calls, we - need an induction principle which works \qt{backwards}. We will use - @{text inc_induct}, which allows to do induction from a fixed number - \qt{downwards}: - - \begin{center}@{thm inc_induct}\hfill(@{text "inc_induct"})\end{center} - - Figure \ref{findzero_term} gives a detailed Isar proof of the fact - that @{text findzero} terminates if there is a zero which is greater - or equal to @{term n}. First we derive two useful rules which will - solve the base case and the step case of the induction. The - induction is then straightforward, except for the unusual induction - principle. - -*} - -text_raw {* -\begin{figure} -\hrule\vspace{6pt} -\begin{minipage}{0.8\textwidth} -\isabellestyle{it} -\isastyle\isamarkuptrue -*} -lemma findzero_termination: - assumes "x \ n" and "f x = 0" - shows "findzero_dom (f, n)" -proof - - have base: "findzero_dom (f, x)" - by (rule findzero.domintros) (simp add:`f x = 0`) - - have step: "\i. findzero_dom (f, Suc i) - \ findzero_dom (f, i)" - by (rule findzero.domintros) simp - - from `x \ n` show ?thesis - proof (induct rule:inc_induct) - show "findzero_dom (f, x)" by (rule base) - next - fix i assume "findzero_dom (f, Suc i)" - thus "findzero_dom (f, i)" by (rule step) - qed -qed -text_raw {* -\isamarkupfalse\isabellestyle{tt} -\end{minipage}\vspace{6pt}\hrule -\caption{Termination proof for @{text findzero}}\label{findzero_term} -\end{figure} -*} - -text {* - Again, the proof given in Fig.~\ref{findzero_term} has a lot of - detail in order to explain the principles. Using more automation, we - can also have a short proof: -*} - -lemma findzero_termination_short: - assumes zero: "x >= n" - assumes [simp]: "f x = 0" - shows "findzero_dom (f, n)" -using zero -by (induct rule:inc_induct) (auto intro: findzero.domintros) - -text {* - \noindent It is simple to combine the partial correctness result with the - termination lemma: -*} - -lemma findzero_total_correctness: - "f x = 0 \ f (findzero f 0) = 0" -by (blast intro: findzero_zero findzero_termination) - -subsection {* Definition of the domain predicate *} - -text {* - Sometimes it is useful to know what the definition of the domain - predicate looks like. Actually, @{text findzero_dom} is just an - abbreviation: - - @{abbrev[display] findzero_dom} - - The domain predicate is the \emph{accessible part} of a relation @{const - findzero_rel}, which was also created internally by the function - package. @{const findzero_rel} is just a normal - inductive predicate, so we can inspect its definition by - looking at the introduction rules @{text findzero_rel.intros}. - In our case there is just a single rule: - - @{thm[display] findzero_rel.intros} - - The predicate @{const findzero_rel} - describes the \emph{recursion relation} of the function - definition. The recursion relation is a binary relation on - the arguments of the function that relates each argument to its - recursive calls. In general, there is one introduction rule for each - recursive call. - - The predicate @{term "accp findzero_rel"} is the accessible part of - that relation. An argument belongs to the accessible part, if it can - be reached in a finite number of steps (cf.~its definition in @{text - "Wellfounded.thy"}). - - Since the domain predicate is just an abbreviation, you can use - lemmas for @{const accp} and @{const findzero_rel} directly. Some - lemmas which are occasionally useful are @{text accpI}, @{text - accp_downward}, and of course the introduction and elimination rules - for the recursion relation @{text "findzero.intros"} and @{text "findzero.cases"}. -*} - -(*lemma findzero_nicer_domintros: - "f x = 0 \ findzero_dom (f, x)" - "findzero_dom (f, Suc x) \ findzero_dom (f, x)" -by (rule accpI, erule findzero_rel.cases, auto)+ -*) - -subsection {* A Useful Special Case: Tail recursion *} - -text {* - The domain predicate is our trick that allows us to model partiality - in a world of total functions. The downside of this is that we have - to carry it around all the time. The termination proof above allowed - us to replace the abstract @{term "findzero_dom (f, n)"} by the more - concrete @{term "(x \ n \ f x = (0::nat))"}, but the condition is still - there and can only be discharged for special cases. - In particular, the domain predicate guards the unfolding of our - function, since it is there as a condition in the @{text psimp} - rules. - - Now there is an important special case: We can actually get rid - of the condition in the simplification rules, \emph{if the function - is tail-recursive}. The reason is that for all tail-recursive - equations there is a total function satisfying them, even if they - are non-terminating. - -% A function is tail recursive, if each call to the function is either -% equal -% -% So the outer form of the -% -%if it can be written in the following -% form: -% {term[display] "f x = (if COND x then BASE x else f (LOOP x))"} - - - The function package internally does the right construction and can - derive the unconditional simp rules, if we ask it to do so. Luckily, - our @{const "findzero"} function is tail-recursive, so we can just go - back and add another option to the \cmd{function} command: - -\vspace{1ex} -\noindent\cmd{function} @{text "(domintros, tailrec) findzero :: \"(nat \ nat) \ nat \ nat\""}\\% -\cmd{where}\isanewline% -\ \ \ldots\\% - - - \noindent Now, we actually get unconditional simplification rules, even - though the function is partial: -*} - -thm findzero.simps - -text {* - @{thm[display] findzero.simps} - - \noindent Of course these would make the simplifier loop, so we better remove - them from the simpset: -*} - -declare findzero.simps[simp del] - -text {* - Getting rid of the domain conditions in the simplification rules is - not only useful because it simplifies proofs. It is also required in - order to use Isabelle's code generator to generate ML code - from a function definition. - Since the code generator only works with equations, it cannot be - used with @{text "psimp"} rules. Thus, in order to generate code for - partial functions, they must be defined as a tail recursion. - Luckily, many functions have a relatively natural tail recursive - definition. -*} - -section {* Nested recursion *} - -text {* - Recursive calls which are nested in one another frequently cause - complications, since their termination proof can depend on a partial - correctness property of the function itself. - - As a small example, we define the \qt{nested zero} function: -*} - -function nz :: "nat \ nat" -where - "nz 0 = 0" -| "nz (Suc n) = nz (nz n)" -by pat_completeness auto - -text {* - If we attempt to prove termination using the identity measure on - naturals, this fails: -*} - -termination - apply (relation "measure (\n. n)") - apply auto - -txt {* - We get stuck with the subgoal - - @{subgoals[display]} - - Of course this statement is true, since we know that @{const nz} is - the zero function. And in fact we have no problem proving this - property by induction. -*} -(*<*)oops(*>*) -lemma nz_is_zero: "nz_dom n \ nz n = 0" - by (induct rule:nz.pinduct) auto - -text {* - We formulate this as a partial correctness lemma with the condition - @{term "nz_dom n"}. This allows us to prove it with the @{text - pinduct} rule before we have proved termination. With this lemma, - the termination proof works as expected: -*} - -termination - by (relation "measure (\n. n)") (auto simp: nz_is_zero) - -text {* - As a general strategy, one should prove the statements needed for - termination as a partial property first. Then they can be used to do - the termination proof. This also works for less trivial - examples. Figure \ref{f91} defines the 91-function, a well-known - challenge problem due to John McCarthy, and proves its termination. -*} - -text_raw {* -\begin{figure} -\hrule\vspace{6pt} -\begin{minipage}{0.8\textwidth} -\isabellestyle{it} -\isastyle\isamarkuptrue -*} - -function f91 :: "nat \ nat" -where - "f91 n = (if 100 < n then n - 10 else f91 (f91 (n + 11)))" -by pat_completeness auto - -lemma f91_estimate: - assumes trm: "f91_dom n" - shows "n < f91 n + 11" -using trm by induct auto - -termination -proof - let ?R = "measure (\x. 101 - x)" - show "wf ?R" .. - - fix n :: nat assume "\ 100 < n" -- "Assumptions for both calls" - - thus "(n + 11, n) \ ?R" by simp -- "Inner call" - - assume inner_trm: "f91_dom (n + 11)" -- "Outer call" - with f91_estimate have "n + 11 < f91 (n + 11) + 11" . - with `\ 100 < n` show "(f91 (n + 11), n) \ ?R" by simp -qed - -text_raw {* -\isamarkupfalse\isabellestyle{tt} -\end{minipage} -\vspace{6pt}\hrule -\caption{McCarthy's 91-function}\label{f91} -\end{figure} -*} - - -section {* Higher-Order Recursion *} - -text {* - Higher-order recursion occurs when recursive calls - are passed as arguments to higher-order combinators such as @{const - map}, @{term filter} etc. - As an example, imagine a datatype of n-ary trees: -*} - -datatype 'a tree = - Leaf 'a -| Branch "'a tree list" - - -text {* \noindent We can define a function which swaps the left and right subtrees recursively, using the - list functions @{const rev} and @{const map}: *} - -fun mirror :: "'a tree \ 'a tree" -where - "mirror (Leaf n) = Leaf n" -| "mirror (Branch l) = Branch (rev (map mirror l))" - -text {* - Although the definition is accepted without problems, let us look at the termination proof: -*} - -termination proof - txt {* - - As usual, we have to give a wellfounded relation, such that the - arguments of the recursive calls get smaller. But what exactly are - the arguments of the recursive calls when mirror is given as an - argument to @{const map}? Isabelle gives us the - subgoals - - @{subgoals[display,indent=0]} - - So the system seems to know that @{const map} only - applies the recursive call @{term "mirror"} to elements - of @{term "l"}, which is essential for the termination proof. - - This knowledge about @{const map} is encoded in so-called congruence rules, - which are special theorems known to the \cmd{function} command. The - rule for @{const map} is - - @{thm[display] map_cong} - - You can read this in the following way: Two applications of @{const - map} are equal, if the list arguments are equal and the functions - coincide on the elements of the list. This means that for the value - @{term "map f l"} we only have to know how @{term f} behaves on - the elements of @{term l}. - - Usually, one such congruence rule is - needed for each higher-order construct that is used when defining - new functions. In fact, even basic functions like @{const - If} and @{const Let} are handled by this mechanism. The congruence - rule for @{const If} states that the @{text then} branch is only - relevant if the condition is true, and the @{text else} branch only if it - is false: - - @{thm[display] if_cong} - - Congruence rules can be added to the - function package by giving them the @{term fundef_cong} attribute. - - The constructs that are predefined in Isabelle, usually - come with the respective congruence rules. - But if you define your own higher-order functions, you may have to - state and prove the required congruence rules yourself, if you want to use your - functions in recursive definitions. -*} -(*<*)oops(*>*) - -subsection {* Congruence Rules and Evaluation Order *} - -text {* - Higher order logic differs from functional programming languages in - that it has no built-in notion of evaluation order. A program is - just a set of equations, and it is not specified how they must be - evaluated. - - However for the purpose of function definition, we must talk about - evaluation order implicitly, when we reason about termination. - Congruence rules express that a certain evaluation order is - consistent with the logical definition. - - Consider the following function. -*} - -function f :: "nat \ bool" -where - "f n = (n = 0 \ f (n - 1))" -(*<*)by pat_completeness auto(*>*) - -text {* - For this definition, the termination proof fails. The default configuration - specifies no congruence rule for disjunction. We have to add a - congruence rule that specifies left-to-right evaluation order: - - \vspace{1ex} - \noindent @{thm disj_cong}\hfill(@{text "disj_cong"}) - \vspace{1ex} - - Now the definition works without problems. Note how the termination - proof depends on the extra condition that we get from the congruence - rule. - - However, as evaluation is not a hard-wired concept, we - could just turn everything around by declaring a different - congruence rule. Then we can make the reverse definition: -*} - -lemma disj_cong2[fundef_cong]: - "(\ Q' \ P = P') \ (Q = Q') \ (P \ Q) = (P' \ Q')" - by blast - -fun f' :: "nat \ bool" -where - "f' n = (f' (n - 1) \ n = 0)" - -text {* - \noindent These examples show that, in general, there is no \qt{best} set of - congruence rules. - - However, such tweaking should rarely be necessary in - practice, as most of the time, the default set of congruence rules - works well. -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/Thy/ROOT.ML --- a/doc-src/IsarAdvanced/Functions/Thy/ROOT.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,4 +0,0 @@ - -(* $Id$ *) - -use_thy "Functions"; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/Thy/document/Functions.tex --- a/doc-src/IsarAdvanced/Functions/Thy/document/Functions.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1985 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Functions}% -% -\isadelimtheory -\isanewline -\isanewline -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Functions\isanewline -\isakeyword{imports}\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupsection{Function Definitions for Dummies% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In most cases, defining a recursive function is just as simple as other definitions:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{fun}\isamarkupfalse% -\ fib\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}fib\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}fib\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}fib\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ fib\ n\ {\isacharplus}\ fib\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -The syntax is rather self-explanatory: We introduce a function by - giving its name, its type, - and a set of defining recursive equations. - If we leave out the type, the most general type will be - inferred, which can sometimes lead to surprises: Since both \isa{{\isadigit{1}}} and \isa{{\isacharplus}} are overloaded, we would end up - with \isa{fib\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ {\isacharprime}a{\isacharcolon}{\isacharcolon}{\isacharbraceleft}one{\isacharcomma}plus{\isacharbraceright}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -The function always terminates, since its argument gets smaller in - every recursive call. - Since HOL is a logic of total functions, termination is a - fundamental requirement to prevent inconsistencies\footnote{From the - \qt{definition} \isa{f{\isacharparenleft}n{\isacharparenright}\ {\isacharequal}\ f{\isacharparenleft}n{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}} we could prove - \isa{{\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}} by subtracting \isa{f{\isacharparenleft}n{\isacharparenright}} on both sides.}. - Isabelle tries to prove termination automatically when a definition - is made. In \S\ref{termination}, we will look at cases where this - fails and see what to do then.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Pattern matching% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\label{patmatch} - Like in functional programming, we can use pattern matching to - define functions. At the moment we will only consider \emph{constructor - patterns}, which only consist of datatype constructors and - variables. Furthermore, patterns must be linear, i.e.\ all variables - on the left hand side of an equation must be distinct. In - \S\ref{genpats} we discuss more general pattern matching. - - If patterns overlap, the order of the equations is taken into - account. The following function inserts a fixed element between any - two elements of a list:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{fun}\isamarkupfalse% -\ sep\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}sep\ a\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}sep\ a\ xs\ \ \ \ \ \ \ {\isacharequal}\ xs{\isachardoublequoteclose}% -\begin{isamarkuptext}% -Overlapping patterns are interpreted as \qt{increments} to what is - already there: The second equation is only meant for the cases where - the first one does not match. Consequently, Isabelle replaces it - internally by the remaining cases, making the patterns disjoint:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{thm}\isamarkupfalse% -\ sep{\isachardot}simps% -\begin{isamarkuptext}% -\begin{isabelle}% -sep\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}\isasep\isanewline% -sep\ a\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\isasep\isanewline% -sep\ a\ {\isacharbrackleft}v{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}v{\isacharbrackright}% -\end{isabelle}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\noindent The equations from function definitions are automatically used in - simplification:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isachardoublequoteopen}sep\ {\isadigit{0}}\ {\isacharbrackleft}{\isadigit{1}}{\isacharcomma}\ {\isadigit{2}}{\isacharcomma}\ {\isadigit{3}}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isadigit{1}}{\isacharcomma}\ {\isadigit{0}}{\isacharcomma}\ {\isadigit{2}}{\isacharcomma}\ {\isadigit{0}}{\isacharcomma}\ {\isadigit{3}}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Induction% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle provides customized induction rules for recursive - functions. These rules follow the recursive structure of the - definition. Here is the rule \isa{sep{\isachardot}induct} arising from the - above definition of \isa{sep}: - - \begin{isabelle}% -{\isasymlbrakk}{\isasymAnd}a\ x\ y\ xs{\isachardot}\ {\isacharquery}P\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharquery}P\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}{\isacharsemicolon}\ {\isasymAnd}a{\isachardot}\ {\isacharquery}P\ a\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\ {\isasymAnd}a\ v{\isachardot}\ {\isacharquery}P\ a\ {\isacharbrackleft}v{\isacharbrackright}{\isasymrbrakk}\isanewline -{\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}a{\isadigit{0}}{\isachardot}{\isadigit{0}}\ {\isacharquery}a{\isadigit{1}}{\isachardot}{\isadigit{0}}% -\end{isabelle} - - We have a step case for list with at least two elements, and two - base cases for the zero- and the one-element list. Here is a simple - proof about \isa{sep} and \isa{map}% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isachardoublequoteopen}map\ f\ {\isacharparenleft}sep\ x\ ys{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}map\ f\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}induct\ x\ ys\ rule{\isacharcolon}\ sep{\isachardot}induct{\isacharparenright}% -\begin{isamarkuptxt}% -We get three cases, like in the definition. - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a\ x\ y\ xs{\isachardot}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharparenleft}y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a{\isachardot}\ map\ f\ {\isacharparenleft}sep\ a\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}a\ v{\isachardot}\ map\ f\ {\isacharparenleft}sep\ a\ {\isacharbrackleft}v{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharbrackleft}v{\isacharbrackright}{\isacharparenright}% -\end{isabelle}% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ auto\ \isanewline -\isacommand{done}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -With the \cmd{fun} command, you can define about 80\% of the - functions that occur in practice. The rest of this tutorial explains - the remaining 20\%.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{fun vs.\ function% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The \cmd{fun} command provides a - convenient shorthand notation for simple function definitions. In - this mode, Isabelle tries to solve all the necessary proof obligations - automatically. If any proof fails, the definition is - rejected. This can either mean that the definition is indeed faulty, - or that the default proof procedures are just not smart enough (or - rather: not designed) to handle the definition. - - By expanding the abbreviation to the more verbose \cmd{function} command, these proof obligations become visible and can be analyzed or - solved manually. The expansion from \cmd{fun} to \cmd{function} is as follows: - -\end{isamarkuptext} - - -\[\left[\;\begin{minipage}{0.25\textwidth}\vspace{6pt} -\cmd{fun} \isa{f\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}}\\% -\cmd{where}\\% -\hspace*{2ex}{\it equations}\\% -\hspace*{2ex}\vdots\vspace*{6pt} -\end{minipage}\right] -\quad\equiv\quad -\left[\;\begin{minipage}{0.48\textwidth}\vspace{6pt} -\cmd{function} \isa{{\isacharparenleft}}\cmd{sequential}\isa{{\isacharparenright}\ f\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}}\\% -\cmd{where}\\% -\hspace*{2ex}{\it equations}\\% -\hspace*{2ex}\vdots\\% -\cmd{by} \isa{pat{\isacharunderscore}completeness\ auto}\\% -\cmd{termination by} \isa{lexicographic{\isacharunderscore}order}\vspace{6pt} -\end{minipage} -\right]\] - -\begin{isamarkuptext} - \vspace*{1em} - \noindent Some details have now become explicit: - - \begin{enumerate} - \item The \cmd{sequential} option enables the preprocessing of - pattern overlaps which we already saw. Without this option, the equations - must already be disjoint and complete. The automatic completion only - works with constructor patterns. - - \item A function definition produces a proof obligation which - expresses completeness and compatibility of patterns (we talk about - this later). The combination of the methods \isa{pat{\isacharunderscore}completeness} and - \isa{auto} is used to solve this proof obligation. - - \item A termination proof follows the definition, started by the - \cmd{termination} command. This will be explained in \S\ref{termination}. - \end{enumerate} - Whenever a \cmd{fun} command fails, it is usually a good idea to - expand the syntax to the more verbose \cmd{function} form, to see - what is actually going on.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Termination% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\label{termination} - The method \isa{lexicographic{\isacharunderscore}order} is the default method for - termination proofs. It can prove termination of a - certain class of functions by searching for a suitable lexicographic - combination of size measures. Of course, not all functions have such - a simple termination argument. For them, we can specify the termination - relation manually.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{The {\tt relation} method% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Consider the following function, which sums up natural numbers up to - \isa{N}, using a counter \isa{i}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ sum\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}sum\ i\ N\ {\isacharequal}\ {\isacharparenleft}if\ i\ {\isachargreater}\ N\ then\ {\isadigit{0}}\ else\ i\ {\isacharplus}\ sum\ {\isacharparenleft}Suc\ i{\isacharparenright}\ N{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ pat{\isacharunderscore}completeness\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent The \isa{lexicographic{\isacharunderscore}order} method fails on this example, because none of the - arguments decreases in the recursive call, with respect to the standard size ordering. - To prove termination manually, we must provide a custom wellfounded relation. - - The termination argument for \isa{sum} is based on the fact that - the \emph{difference} between \isa{i} and \isa{N} gets - smaller in every step, and that the recursion stops when \isa{i} - is greater than \isa{N}. Phrased differently, the expression - \isa{N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i} always decreases. - - We can use this expression as a measure function suitable to prove termination.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{termination}\isamarkupfalse% -\ sum\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% -\begin{isamarkuptxt}% -The \cmd{termination} command sets up the termination goal for the - specified function \isa{sum}. If the function name is omitted, it - implicitly refers to the last function definition. - - The \isa{relation} method takes a relation of - type \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set}, where \isa{{\isacharprime}a} is the argument type of - the function. If the function has multiple curried arguments, then - these are packed together into a tuple, as it happened in the above - example. - - The predefined function \isa{{\isachardoublequote}measure\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set{\isachardoublequote}} constructs a - wellfounded relation from a mapping into the natural numbers (a - \emph{measure function}). - - After the invocation of \isa{relation}, we must prove that (a) - the relation we supplied is wellfounded, and (b) that the arguments - of recursive calls indeed decrease with respect to the - relation: - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ wf\ {\isacharparenleft}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}\ N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharparenright}{\isacharparenright}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}i\ N{\isachardot}\ {\isasymnot}\ N\ {\isacharless}\ i\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isacharparenleft}Suc\ i{\isacharcomma}\ N{\isacharparenright}{\isacharcomma}\ i{\isacharcomma}\ N{\isacharparenright}\ {\isasymin}\ measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}\ N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharparenright}% -\end{isabelle} - - These goals are all solved by \isa{auto}:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ auto\isanewline -\isacommand{done}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Let us complicate the function a little, by adding some more - recursive calls:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ foo\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}foo\ i\ N\ {\isacharequal}\ {\isacharparenleft}if\ i\ {\isachargreater}\ N\ \isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ then\ {\isacharparenleft}if\ N\ {\isacharequal}\ {\isadigit{0}}\ then\ {\isadigit{0}}\ else\ foo\ {\isadigit{0}}\ {\isacharparenleft}N\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ else\ i\ {\isacharplus}\ foo\ {\isacharparenleft}Suc\ i{\isacharparenright}\ N{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ pat{\isacharunderscore}completeness\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -When \isa{i} has reached \isa{N}, it starts at zero again - and \isa{N} is decremented. - This corresponds to a nested - loop where one index counts up and the other down. Termination can - be proved using a lexicographic combination of two measures, namely - the value of \isa{N} and the above difference. The \isa{measures} combinator generalizes \isa{measure} by taking a - list of measure functions.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{termination}\isamarkupfalse% -\ \isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measures\ {\isacharbrackleft}{\isasymlambda}{\isacharparenleft}i{\isacharcomma}\ N{\isacharparenright}{\isachardot}\ N{\isacharcomma}\ {\isasymlambda}{\isacharparenleft}i{\isacharcomma}N{\isacharparenright}{\isachardot}\ N\ {\isacharplus}\ {\isadigit{1}}\ {\isacharminus}\ i{\isacharbrackright}{\isachardoublequoteclose}{\isacharparenright}\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{How \isa{lexicographic{\isacharunderscore}order} works% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -To see how the automatic termination proofs work, let's look at an - example where it fails\footnote{For a detailed discussion of the - termination prover, see \cite{bulwahnKN07}}: - -\end{isamarkuptext} -\cmd{fun} \isa{fails\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymRightarrow}\ nat\ list\ {\isasymRightarrow}\ nat{\isachardoublequote}}\\% -\cmd{where}\\% -\hspace*{2ex}\isa{{\isachardoublequote}fails\ a\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ a{\isachardoublequote}}\\% -|\hspace*{1.5ex}\isa{{\isachardoublequote}fails\ a\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ fails\ {\isacharparenleft}x\ {\isacharplus}\ a{\isacharparenright}\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}{\isachardoublequote}}\\ -\begin{isamarkuptext} - -\noindent Isabelle responds with the following error: - -\begin{isabelle} -*** Unfinished subgoals:\newline -*** (a, 1, <):\newline -*** \ 1.~\isa{{\isasymAnd}x{\isachardot}\ x\ {\isacharequal}\ {\isadigit{0}}}\newline -*** (a, 1, <=):\newline -*** \ 1.~False\newline -*** (a, 2, <):\newline -*** \ 1.~False\newline -*** Calls:\newline -*** a) \isa{{\isacharparenleft}a{\isacharcomma}\ x\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharminus}{\isacharminus}{\isachargreater}{\isachargreater}\ {\isacharparenleft}x\ {\isacharplus}\ a{\isacharcomma}\ x\ {\isacharhash}\ xs{\isacharparenright}}\newline -*** Measures:\newline -*** 1) \isa{{\isasymlambda}x{\isachardot}\ size\ {\isacharparenleft}fst\ x{\isacharparenright}}\newline -*** 2) \isa{{\isasymlambda}x{\isachardot}\ size\ {\isacharparenleft}snd\ x{\isacharparenright}}\newline -*** Result matrix:\newline -*** \ \ \ \ 1\ \ 2 \newline -*** a: ? <= \newline -*** Could not find lexicographic termination order.\newline -*** At command "fun".\newline -\end{isabelle}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -The key to this error message is the matrix at the bottom. The rows - of that matrix correspond to the different recursive calls (In our - case, there is just one). The columns are the function's arguments - (expressed through different measure functions, which map the - argument tuple to a natural number). - - The contents of the matrix summarize what is known about argument - descents: The second argument has a weak descent (\isa{{\isacharless}{\isacharequal}}) at the - recursive call, and for the first argument nothing could be proved, - which is expressed by \isa{{\isacharquery}}. In general, there are the values - \isa{{\isacharless}}, \isa{{\isacharless}{\isacharequal}} and \isa{{\isacharquery}}. - - For the failed proof attempts, the unfinished subgoals are also - printed. Looking at these will often point to a missing lemma. - -% As a more real example, here is quicksort:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Mutual Recursion% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -If two or more functions call one another mutually, they have to be defined - in one step. Here are \isa{even} and \isa{odd}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ even\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ odd\ \ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}even\ {\isadigit{0}}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}odd\ {\isadigit{0}}\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}even\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ odd\ n{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}odd\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ even\ n{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ pat{\isacharunderscore}completeness\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -To eliminate the mutual dependencies, Isabelle internally - creates a single function operating on the sum - type \isa{nat\ {\isacharplus}\ nat}. Then, \isa{even} and \isa{odd} are - defined as projections. Consequently, termination has to be proved - simultaneously for both functions, by specifying a measure on the - sum type:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{termination}\isamarkupfalse% -\ \isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ case\ x\ of\ Inl\ n\ {\isasymRightarrow}\ n\ {\isacharbar}\ Inr\ n\ {\isasymRightarrow}\ n{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -We could also have used \isa{lexicographic{\isacharunderscore}order}, which - supports mutual recursive termination proofs to a certain extent.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Induction for mutual recursion% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -When functions are mutually recursive, proving properties about them - generally requires simultaneous induction. The induction rule \isa{even{\isacharunderscore}odd{\isachardot}induct} - generated from the above definition reflects this. - - Let us prove something about \isa{even} and \isa{odd}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ even{\isacharunderscore}odd{\isacharunderscore}mod{\isadigit{2}}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}odd\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -We apply simultaneous induction, specifying the induction variable - for both goals, separated by \cmd{and}:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}induct\ n\ \isakeyword{and}\ n\ rule{\isacharcolon}\ even{\isacharunderscore}odd{\isachardot}induct{\isacharparenright}% -\begin{isamarkuptxt}% -We get four subgoals, which correspond to the clauses in the - definition of \isa{even} and \isa{odd}: - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ even\ {\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline -\ {\isadigit{2}}{\isachardot}\ odd\ {\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ odd\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}\ {\isasymLongrightarrow}\ even\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline -\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ odd\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{1}}{\isacharparenright}% -\end{isabelle} - Simplification solves the first two goals, leaving us with two - statements about the \isa{mod} operation to prove:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ simp{\isacharunderscore}all% -\begin{isamarkuptxt}% -\begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ odd\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isacharparenright}% -\end{isabelle} - - \noindent These can be handled by Isabelle's arithmetic decision procedures.% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ arith\isanewline -\isacommand{apply}\isamarkupfalse% -\ arith\isanewline -\isacommand{done}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -In proofs like this, the simultaneous induction is really essential: - Even if we are just interested in one of the results, the other - one is necessary to strengthen the induction hypothesis. If we leave - out the statement about \isa{odd} and just write \isa{True} instead, - the same proof fails:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ failed{\isacharunderscore}attempt{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}True{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}induct\ n\ rule{\isacharcolon}\ even{\isacharunderscore}odd{\isachardot}induct{\isacharparenright}% -\begin{isamarkuptxt}% -\noindent Now the third subgoal is a dead end, since we have no - useful induction hypothesis available: - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ even\ {\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline -\ {\isadigit{2}}{\isachardot}\ True\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ True\ {\isasymLongrightarrow}\ even\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Suc\ n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\isanewline -\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ even\ n\ {\isacharequal}\ {\isacharparenleft}n\ mod\ {\isadigit{2}}\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}\ {\isasymLongrightarrow}\ True% -\end{isabelle}% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{oops}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsection{General pattern matching% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\label{genpats}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Avoiding automatic pattern splitting% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Up to now, we used pattern matching only on datatypes, and the - patterns were always disjoint and complete, and if they weren't, - they were made disjoint automatically like in the definition of - \isa{sep} in \S\ref{patmatch}. - - This automatic splitting can significantly increase the number of - equations involved, and this is not always desirable. The following - example shows the problem: - - Suppose we are modeling incomplete knowledge about the world by a - three-valued datatype, which has values \isa{T}, \isa{F} - and \isa{X} for true, false and uncertain propositions, respectively.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ P{\isadigit{3}}\ {\isacharequal}\ T\ {\isacharbar}\ F\ {\isacharbar}\ X% -\begin{isamarkuptext}% -\noindent Then the conjunction of such values can be defined as follows:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{fun}\isamarkupfalse% -\ And\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}And\ T\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And\ p\ T\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And\ p\ F\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And\ F\ p\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And\ X\ X\ {\isacharequal}\ X{\isachardoublequoteclose}% -\begin{isamarkuptext}% -This definition is useful, because the equations can directly be used - as simplification rules. But the patterns overlap: For example, - the expression \isa{And\ T\ T} is matched by both the first and - the second equation. By default, Isabelle makes the patterns disjoint by - splitting them up, producing instances:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{thm}\isamarkupfalse% -\ And{\isachardot}simps% -\begin{isamarkuptext}% -\isa{And\ T\ {\isacharquery}p\ {\isacharequal}\ {\isacharquery}p\isasep\isanewline% -And\ F\ T\ {\isacharequal}\ F\isasep\isanewline% -And\ X\ T\ {\isacharequal}\ X\isasep\isanewline% -And\ F\ F\ {\isacharequal}\ F\isasep\isanewline% -And\ X\ F\ {\isacharequal}\ F\isasep\isanewline% -And\ F\ X\ {\isacharequal}\ F\isasep\isanewline% -And\ X\ X\ {\isacharequal}\ X} - - \vspace*{1em} - \noindent There are several problems with this: - - \begin{enumerate} - \item If the datatype has many constructors, there can be an - explosion of equations. For \isa{And}, we get seven instead of - five equations, which can be tolerated, but this is just a small - example. - - \item Since splitting makes the equations \qt{less general}, they - do not always match in rewriting. While the term \isa{And\ x\ F} - can be simplified to \isa{F} with the original equations, a - (manual) case split on \isa{x} is now necessary. - - \item The splitting also concerns the induction rule \isa{And{\isachardot}induct}. Instead of five premises it now has seven, which - means that our induction proofs will have more cases. - - \item In general, it increases clarity if we get the same definition - back which we put in. - \end{enumerate} - - If we do not want the automatic splitting, we can switch it off by - leaving out the \cmd{sequential} option. However, we will have to - prove that our pattern matching is consistent\footnote{This prevents - us from defining something like \isa{f\ x\ {\isacharequal}\ True} and \isa{f\ x\ {\isacharequal}\ False} simultaneously.}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ And{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}\ {\isasymRightarrow}\ P{\isadigit{3}}{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}And{\isadigit{2}}\ T\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ p\ T\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ p\ F\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ F\ p\ {\isacharequal}\ F{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}And{\isadigit{2}}\ X\ X\ {\isacharequal}\ X{\isachardoublequoteclose}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -\noindent Now let's look at the proof obligations generated by a - function definition. In this case, they are: - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ {\isasymlbrakk}{\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ F{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ \ }{\isasymAnd}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ p{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ x\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ }{\isasymLongrightarrow}\ P\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}T{\isacharcomma}\ pa{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ pa\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ pa\isanewline -\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ F{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline -\ {\isadigit{5}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ pa{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline -\ {\isadigit{6}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ X\isanewline -\ {\isadigit{7}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ pa\isanewline -\ {\isadigit{8}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}pa{\isacharcomma}\ F{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline -\ {\isadigit{9}}{\isachardot}\ {\isasymAnd}p\ pa{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ pa{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ F\isanewline -\ {\isadigit{1}}{\isadigit{0}}{\isachardot}\ {\isasymAnd}p{\isachardot}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}\ {\isasymLongrightarrow}\ p\ {\isacharequal}\ X% -\end{isabelle}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em} - - The first subgoal expresses the completeness of the patterns. It has - the form of an elimination rule and states that every \isa{x} of - the function's input type must match at least one of the patterns\footnote{Completeness could - be equivalently stated as a disjunction of existential statements: -\isa{{\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}T{\isacharcomma}\ p{\isacharparenright}{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ T{\isacharparenright}{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}p{\isacharcomma}\ F{\isacharparenright}{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}p{\isachardot}\ x\ {\isacharequal}\ {\isacharparenleft}F{\isacharcomma}\ p{\isacharparenright}{\isacharparenright}\ {\isasymor}\ x\ {\isacharequal}\ {\isacharparenleft}X{\isacharcomma}\ X{\isacharparenright}}, and you can use the method \isa{atomize{\isacharunderscore}elim} to get that form instead.}. If the patterns just involve - datatypes, we can solve it with the \isa{pat{\isacharunderscore}completeness} - method:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ pat{\isacharunderscore}completeness% -\begin{isamarkuptxt}% -The remaining subgoals express \emph{pattern compatibility}. We do - allow that an input value matches multiple patterns, but in this - case, the result (i.e.~the right hand sides of the equations) must - also be equal. For each pair of two patterns, there is one such - subgoal. Usually this needs injectivity of the constructors, which - is used automatically by \isa{auto}.% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{by}\isamarkupfalse% -\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Non-constructor patterns% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Most of Isabelle's basic types take the form of inductive datatypes, - and usually pattern matching works on the constructors of such types. - However, this need not be always the case, and the \cmd{function} - command handles other kind of patterns, too. - - One well-known instance of non-constructor patterns are - so-called \emph{$n+k$-patterns}, which are a little controversial in - the functional programming world. Here is the initial fibonacci - example with $n+k$-patterns:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ fib{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}fib{\isadigit{2}}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}fib{\isadigit{2}}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}fib{\isadigit{2}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ fib{\isadigit{2}}\ n\ {\isacharplus}\ fib{\isadigit{2}}\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -This kind of matching is again justified by the proof of pattern - completeness and compatibility. - The proof obligation for pattern completeness states that every natural number is - either \isa{{\isadigit{0}}}, \isa{{\isadigit{1}}} or \isa{n\ {\isacharplus}\ {\isadigit{2}}}: - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ {\isasymlbrakk}x\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ x\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}n{\isachardot}\ x\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline -\ {\isadigit{3}}{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline -\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\isanewline -\ {\isadigit{5}}{\isachardot}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline -\ {\isadigit{6}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{1}}\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\isanewline -\ {\isadigit{7}}{\isachardot}\ {\isasymAnd}n\ na{\isachardot}\isanewline -\isaindent{\ {\isadigit{7}}{\isachardot}\ \ \ \ }n\ {\isacharplus}\ {\isadigit{2}}\ {\isacharequal}\ na\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\isanewline -\isaindent{\ {\isadigit{7}}{\isachardot}\ \ \ \ }fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ na\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ na{\isacharparenright}% -\end{isabelle} - - This is an arithmetic triviality, but unfortunately the - \isa{arith} method cannot handle this specific form of an - elimination rule. However, we can use the method \isa{atomize{\isacharunderscore}elim} to do an ad-hoc conversion to a disjunction of - existentials, which can then be solved by the arithmetic decision procedure. - Pattern compatibility and termination are automatic as usual.% -\end{isamarkuptxt}% -\isamarkuptrue% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{apply}\isamarkupfalse% -\ atomize{\isacharunderscore}elim\isanewline -\isacommand{apply}\isamarkupfalse% -\ arith\isanewline -\isacommand{apply}\isamarkupfalse% -\ auto\isanewline -\isacommand{done}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{termination}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ lexicographic{\isacharunderscore}order% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -We can stretch the notion of pattern matching even more. The - following function is not a sensible functional program, but a - perfectly valid mathematical definition:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ ev\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}ev\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}ev\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{apply}\isamarkupfalse% -\ atomize{\isacharunderscore}elim\isanewline -\isacommand{by}\isamarkupfalse% -\ arith{\isacharplus}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{termination}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}relation\ {\isachardoublequoteopen}{\isacharbraceleft}{\isacharbraceright}{\isachardoublequoteclose}{\isacharparenright}\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -This general notion of pattern matching gives you a certain freedom - in writing down specifications. However, as always, such freedom should - be used with care: - - If we leave the area of constructor - patterns, we have effectively departed from the world of functional - programming. This means that it is no longer possible to use the - code generator, and expect it to generate ML code for our - definitions. Also, such a specification might not work very well together with - simplification. Your mileage may vary.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Conditional equations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The function package also supports conditional equations, which are - similar to guards in a language like Haskell. Here is Euclid's - algorithm written with conditional patterns\footnote{Note that the - patterns are also overlapping in the base case}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ gcd\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}gcd\ x\ {\isadigit{0}}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}gcd\ {\isadigit{0}}\ y\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}x\ {\isacharless}\ y\ {\isasymLongrightarrow}\ gcd\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}\ {\isacharequal}\ gcd\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}y\ {\isacharminus}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}{\isasymnot}\ x\ {\isacharless}\ y\ {\isasymLongrightarrow}\ gcd\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}\ {\isacharequal}\ gcd\ {\isacharparenleft}x\ {\isacharminus}\ y{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}atomize{\isacharunderscore}elim{\isacharcomma}\ auto{\isacharcomma}\ arith{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{termination}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ lexicographic{\isacharunderscore}order% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -By now, you can probably guess what the proof obligations for the - pattern completeness and compatibility look like. - - Again, functions with conditional patterns are not supported by the - code generator.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Pattern matching on strings% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -As strings (as lists of characters) are normal datatypes, pattern - matching on them is possible, but somewhat problematic. Consider the - following definition: - -\end{isamarkuptext} -\noindent\cmd{fun} \isa{check\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}string\ {\isasymRightarrow}\ bool{\isachardoublequote}}\\% -\cmd{where}\\% -\hspace*{2ex}\isa{{\isachardoublequote}check\ {\isacharparenleft}{\isacharprime}{\isacharprime}good{\isacharprime}{\isacharprime}{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequote}}\\% -\isa{{\isacharbar}\ {\isachardoublequote}check\ s\ {\isacharequal}\ False{\isachardoublequote}} -\begin{isamarkuptext} - - \noindent An invocation of the above \cmd{fun} command does not - terminate. What is the problem? Strings are lists of characters, and - characters are a datatype with a lot of constructors. Splitting the - catch-all pattern thus leads to an explosion of cases, which cannot - be handled by Isabelle. - - There are two things we can do here. Either we write an explicit - \isa{if} on the right hand side, or we can use conditional patterns:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ check\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}string\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}check\ {\isacharparenleft}{\isacharprime}{\isacharprime}good{\isacharprime}{\isacharprime}{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}s\ {\isasymnoteq}\ {\isacharprime}{\isacharprime}good{\isacharprime}{\isacharprime}\ {\isasymLongrightarrow}\ check\ s\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsection{Partiality% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In HOL, all functions are total. A function \isa{f} applied to - \isa{x} always has the value \isa{f\ x}, and there is no notion - of undefinedness. - This is why we have to do termination - proofs when defining functions: The proof justifies that the - function can be defined by wellfounded recursion. - - However, the \cmd{function} package does support partiality to a - certain extent. Let's look at the following function which looks - for a zero of a given function f.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ findzero\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}findzero\ f\ n\ {\isacharequal}\ {\isacharparenleft}if\ f\ n\ {\isacharequal}\ {\isadigit{0}}\ then\ n\ else\ findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ pat{\isacharunderscore}completeness\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Clearly, any attempt of a termination proof must fail. And without - that, we do not get the usual rules \isa{findzero{\isachardot}simp} and - \isa{findzero{\isachardot}induct}. So what was the definition good for at all?% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Domain predicates% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The trick is that Isabelle has not only defined the function \isa{findzero}, but also - a predicate \isa{findzero{\isacharunderscore}dom} that characterizes the values where the function - terminates: the \emph{domain} of the function. If we treat a - partial function just as a total function with an additional domain - predicate, we can derive simplification and - induction rules as we do for total functions. They are guarded - by domain conditions and are called \isa{psimps} and \isa{pinduct}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\noindent\begin{minipage}{0.79\textwidth}\begin{isabelle}% -findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ {\isacharquery}n{\isacharparenright}\ {\isasymLongrightarrow}\isanewline -findzero\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isacharparenleft}if\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isadigit{0}}\ then\ {\isacharquery}n\ else\ findzero\ {\isacharquery}f\ {\isacharparenleft}Suc\ {\isacharquery}n{\isacharparenright}{\isacharparenright}% -\end{isabelle}\end{minipage} - \hfill(\isa{findzero{\isachardot}psimps}) - \vspace{1em} - - \noindent\begin{minipage}{0.79\textwidth}\begin{isabelle}% -{\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}a{\isadigit{0}}{\isachardot}{\isadigit{0}}{\isacharcomma}\ {\isacharquery}a{\isadigit{1}}{\isachardot}{\isadigit{0}}{\isacharparenright}{\isacharsemicolon}\isanewline -\isaindent{\ }{\isasymAnd}f\ n{\isachardot}\ {\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isacharsemicolon}\ f\ n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isacharquery}P\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ f\ n{\isasymrbrakk}\isanewline -{\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}a{\isadigit{0}}{\isachardot}{\isadigit{0}}\ {\isacharquery}a{\isadigit{1}}{\isachardot}{\isadigit{0}}% -\end{isabelle}\end{minipage} - \hfill(\isa{findzero{\isachardot}pinduct})% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -Remember that all we - are doing here is use some tricks to make a total function appear - as if it was partial. We can still write the term \isa{findzero\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ {\isadigit{1}}{\isacharparenright}\ {\isadigit{0}}} and like any other term of type \isa{nat} it is equal - to some natural number, although we might not be able to find out - which one. The function is \emph{underdefined}. - - But it is defined enough to prove something interesting about it. We - can prove that if \isa{findzero\ f\ n} - terminates, it indeed returns a zero of \isa{f}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ findzero{\isacharunderscore}zero{\isacharcolon}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ n{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -\noindent We apply induction as usual, but using the partial induction - rule:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}induct\ f\ n\ rule{\isacharcolon}\ findzero{\isachardot}pinduct{\isacharparenright}% -\begin{isamarkuptxt}% -\noindent This gives the following subgoals: - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}f\ n{\isachardot}\ {\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isacharsemicolon}\ f\ n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}{\isasymrbrakk}\isanewline -\isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}f\ n{\isachardot}\ }{\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ n{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}% -\end{isabelle} - - \noindent The hypothesis in our lemma was used to satisfy the first premise in - the induction rule. However, we also get \isa{findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}} as a local assumption in the induction step. This - allows to unfold \isa{findzero\ f\ n} using the \isa{psimps} - rule, and the rest is trivial. Since the \isa{psimps} rules carry the - \isa{{\isacharbrackleft}simp{\isacharbrackright}} attribute by default, we just need a single step:% -\end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\isamarkupfalse% -\ simp\isanewline -\isacommand{done}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Proofs about partial functions are often not harder than for total - functions. Fig.~\ref{findzero_isar} shows a slightly more - complicated proof written in Isar. It is verbose enough to show how - partiality comes into play: From the partial induction, we get an - additional domain condition hypothesis. Observe how this condition - is applied when calls to \isa{findzero} are unfolded.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{figure} -\hrule\vspace{6pt} -\begin{minipage}{0.8\textwidth} -\isabellestyle{it} -\isastyle\isamarkuptrue -\isacommand{lemma}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymlbrakk}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isacharsemicolon}\ x\ {\isasymin}\ {\isacharbraceleft}n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ f\ x\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isacharparenleft}induct\ rule{\isacharcolon}\ findzero{\isachardot}pinduct{\isacharparenright}\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ f\ n\ \isacommand{assume}\isamarkupfalse% -\ dom{\isacharcolon}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ IH{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharsemicolon}\ x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharbraceright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ f\ x\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ x{\isacharunderscore}range{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ \isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}f\ n\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{with}\isamarkupfalse% -\ dom\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}findzero\ f\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \ \ \isacommand{with}\isamarkupfalse% -\ x{\isacharunderscore}range\ \isacommand{show}\isamarkupfalse% -\ False\ \isacommand{by}\isamarkupfalse% -\ auto\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ x{\isacharunderscore}range\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isacharequal}\ n\ {\isasymor}\ x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ auto\isanewline -\ \ \isacommand{thus}\isamarkupfalse% -\ {\isachardoublequoteopen}f\ x\ {\isasymnoteq}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{with}\isamarkupfalse% -\ {\isacharbackquoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% -\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{with}\isamarkupfalse% -\ dom\ \isakeyword{and}\ {\isacharbackquoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharbackquoteclose}\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}Suc\ n\ {\isachardot}{\isachardot}{\isacharless}\ findzero\ f\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \ \ \isacommand{with}\isamarkupfalse% -\ IH\ \isakeyword{and}\ {\isacharbackquoteopen}f\ n\ {\isasymnoteq}\ {\isadigit{0}}{\isacharbackquoteclose}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupfalse\isabellestyle{tt} -\end{minipage}\vspace{6pt}\hrule -\caption{A proof about a partial function}\label{findzero_isar} -\end{figure} -% -\isamarkupsubsection{Partial termination proofs% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Now that we have proved some interesting properties about our - function, we should turn to the domain predicate and see if it is - actually true for some values. Otherwise we would have just proved - lemmas with \isa{False} as a premise. - - Essentially, we need some introduction rules for \isa{findzero{\isacharunderscore}dom}. The function package can prove such domain - introduction rules automatically. But since they are not used very - often (they are almost never needed if the function is total), this - functionality is disabled by default for efficiency reasons. So we have to go - back and ask for them explicitly by passing the \isa{{\isacharparenleft}domintros{\isacharparenright}} option to the function package: - -\vspace{1ex} -\noindent\cmd{function} \isa{{\isacharparenleft}domintros{\isacharparenright}\ findzero\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}nat\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequote}}\\% -\cmd{where}\isanewline% -\ \ \ldots\\ - - \noindent Now the package has proved an introduction rule for \isa{findzero{\isacharunderscore}dom}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{thm}\isamarkupfalse% -\ findzero{\isachardot}domintros% -\begin{isamarkuptext}% -\begin{isabelle}% -{\isacharparenleft}{\isadigit{0}}\ {\isacharless}\ {\isacharquery}f\ {\isacharquery}n\ {\isasymLongrightarrow}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ Suc\ {\isacharquery}n{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ {\isacharquery}n{\isacharparenright}% -\end{isabelle} - - Domain introduction rules allow to show that a given value lies in the - domain of a function, if the arguments of all recursive calls - are in the domain as well. They allow to do a \qt{single step} in a - termination proof. Usually, you want to combine them with a suitable - induction principle. - - Since our function increases its argument at recursive calls, we - need an induction principle which works \qt{backwards}. We will use - \isa{inc{\isacharunderscore}induct}, which allows to do induction from a fixed number - \qt{downwards}: - - \begin{center}\isa{{\isasymlbrakk}{\isacharquery}i\ {\isasymle}\ {\isacharquery}j{\isacharsemicolon}\ {\isacharquery}P\ {\isacharquery}j{\isacharsemicolon}\ {\isasymAnd}i{\isachardot}\ {\isasymlbrakk}i\ {\isacharless}\ {\isacharquery}j{\isacharsemicolon}\ {\isacharquery}P\ {\isacharparenleft}Suc\ i{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ i{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ {\isacharquery}i}\hfill(\isa{inc{\isacharunderscore}induct})\end{center} - - Figure \ref{findzero_term} gives a detailed Isar proof of the fact - that \isa{findzero} terminates if there is a zero which is greater - or equal to \isa{n}. First we derive two useful rules which will - solve the base case and the step case of the induction. The - induction is then straightforward, except for the unusual induction - principle.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{figure} -\hrule\vspace{6pt} -\begin{minipage}{0.8\textwidth} -\isabellestyle{it} -\isastyle\isamarkuptrue -\isacommand{lemma}\isamarkupfalse% -\ findzero{\isacharunderscore}termination{\isacharcolon}\isanewline -\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}x\ {\isasymge}\ n{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{shows}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isacharminus}\ \isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ base{\isacharcolon}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ findzero{\isachardot}domintros{\isacharparenright}\ {\isacharparenleft}simp\ add{\isacharcolon}{\isacharbackquoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}{\isacharbackquoteclose}{\isacharparenright}\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ step{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}i{\isachardot}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ Suc\ i{\isacharparenright}\ \isanewline -\ \ \ \ {\isasymLongrightarrow}\ findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ i{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ findzero{\isachardot}domintros{\isacharparenright}\ simp\isanewline -\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ {\isacharbackquoteopen}x\ {\isasymge}\ n{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% -\ {\isacharquery}thesis\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isacharparenleft}induct\ rule{\isacharcolon}inc{\isacharunderscore}induct{\isacharparenright}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ x{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ base{\isacharparenright}\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ i\ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ Suc\ i{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{thus}\isamarkupfalse% -\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ i{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ step{\isacharparenright}\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupfalse\isabellestyle{tt} -\end{minipage}\vspace{6pt}\hrule -\caption{Termination proof for \isa{findzero}}\label{findzero_term} -\end{figure} -% -\begin{isamarkuptext}% -Again, the proof given in Fig.~\ref{findzero_term} has a lot of - detail in order to explain the principles. Using more automation, we - can also have a short proof:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ findzero{\isacharunderscore}termination{\isacharunderscore}short{\isacharcolon}\isanewline -\ \ \isakeyword{assumes}\ zero{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isachargreater}{\isacharequal}\ n{\isachardoublequoteclose}\ \isanewline -\ \ \isakeyword{assumes}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{shows}\ {\isachardoublequoteopen}findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{using}\isamarkupfalse% -\ zero\isanewline -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ rule{\isacharcolon}inc{\isacharunderscore}induct{\isacharparenright}\ {\isacharparenleft}auto\ intro{\isacharcolon}\ findzero{\isachardot}domintros{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent It is simple to combine the partial correctness result with the - termination lemma:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ findzero{\isacharunderscore}total{\isacharunderscore}correctness{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}f\ x\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}findzero\ f\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}blast\ intro{\isacharcolon}\ findzero{\isacharunderscore}zero\ findzero{\isacharunderscore}termination{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Definition of the domain predicate% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Sometimes it is useful to know what the definition of the domain - predicate looks like. Actually, \isa{findzero{\isacharunderscore}dom} is just an - abbreviation: - - \begin{isabelle}% -findzero{\isacharunderscore}dom\ {\isasymequiv}\ accp\ findzero{\isacharunderscore}rel% -\end{isabelle} - - The domain predicate is the \emph{accessible part} of a relation \isa{findzero{\isacharunderscore}rel}, which was also created internally by the function - package. \isa{findzero{\isacharunderscore}rel} is just a normal - inductive predicate, so we can inspect its definition by - looking at the introduction rules \isa{findzero{\isacharunderscore}rel{\isachardot}intros}. - In our case there is just a single rule: - - \begin{isabelle}% -{\isacharquery}f\ {\isacharquery}n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ findzero{\isacharunderscore}rel\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ Suc\ {\isacharquery}n{\isacharparenright}\ {\isacharparenleft}{\isacharquery}f{\isacharcomma}\ {\isacharquery}n{\isacharparenright}% -\end{isabelle} - - The predicate \isa{findzero{\isacharunderscore}rel} - describes the \emph{recursion relation} of the function - definition. The recursion relation is a binary relation on - the arguments of the function that relates each argument to its - recursive calls. In general, there is one introduction rule for each - recursive call. - - The predicate \isa{findzero{\isacharunderscore}dom} is the accessible part of - that relation. An argument belongs to the accessible part, if it can - be reached in a finite number of steps (cf.~its definition in \isa{Accessible{\isacharunderscore}Part{\isachardot}thy}). - - Since the domain predicate is just an abbreviation, you can use - lemmas for \isa{accp} and \isa{findzero{\isacharunderscore}rel} directly. Some - lemmas which are occasionally useful are \isa{accpI}, \isa{accp{\isacharunderscore}downward}, and of course the introduction and elimination rules - for the recursion relation \isa{findzero{\isachardot}intros} and \isa{findzero{\isachardot}cases}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{A Useful Special Case: Tail recursion% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The domain predicate is our trick that allows us to model partiality - in a world of total functions. The downside of this is that we have - to carry it around all the time. The termination proof above allowed - us to replace the abstract \isa{findzero{\isacharunderscore}dom\ {\isacharparenleft}f{\isacharcomma}\ n{\isacharparenright}} by the more - concrete \isa{n\ {\isasymle}\ x\ {\isasymand}\ f\ x\ {\isacharequal}\ {\isadigit{0}}}, but the condition is still - there and can only be discharged for special cases. - In particular, the domain predicate guards the unfolding of our - function, since it is there as a condition in the \isa{psimp} - rules. - - Now there is an important special case: We can actually get rid - of the condition in the simplification rules, \emph{if the function - is tail-recursive}. The reason is that for all tail-recursive - equations there is a total function satisfying them, even if they - are non-terminating. - -% A function is tail recursive, if each call to the function is either -% equal -% -% So the outer form of the -% -%if it can be written in the following -% form: -% {term[display] "f x = (if COND x then BASE x else f (LOOP x))"} - - - The function package internally does the right construction and can - derive the unconditional simp rules, if we ask it to do so. Luckily, - our \isa{findzero} function is tail-recursive, so we can just go - back and add another option to the \cmd{function} command: - -\vspace{1ex} -\noindent\cmd{function} \isa{{\isacharparenleft}domintros{\isacharcomma}\ tailrec{\isacharparenright}\ findzero\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}nat\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequote}}\\% -\cmd{where}\isanewline% -\ \ \ldots\\% - - - \noindent Now, we actually get unconditional simplification rules, even - though the function is partial:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{thm}\isamarkupfalse% -\ findzero{\isachardot}simps% -\begin{isamarkuptext}% -\begin{isabelle}% -findzero\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isacharparenleft}if\ {\isacharquery}f\ {\isacharquery}n\ {\isacharequal}\ {\isadigit{0}}\ then\ {\isacharquery}n\ else\ findzero\ {\isacharquery}f\ {\isacharparenleft}Suc\ {\isacharquery}n{\isacharparenright}{\isacharparenright}% -\end{isabelle} - - \noindent Of course these would make the simplifier loop, so we better remove - them from the simpset:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{declare}\isamarkupfalse% -\ findzero{\isachardot}simps{\isacharbrackleft}simp\ del{\isacharbrackright}% -\begin{isamarkuptext}% -Getting rid of the domain conditions in the simplification rules is - not only useful because it simplifies proofs. It is also required in - order to use Isabelle's code generator to generate ML code - from a function definition. - Since the code generator only works with equations, it cannot be - used with \isa{psimp} rules. Thus, in order to generate code for - partial functions, they must be defined as a tail recursion. - Luckily, many functions have a relatively natural tail recursive - definition.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Nested recursion% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Recursive calls which are nested in one another frequently cause - complications, since their termination proof can depend on a partial - correctness property of the function itself. - - As a small example, we define the \qt{nested zero} function:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ nz\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}nz\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}nz\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ nz\ {\isacharparenleft}nz\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ pat{\isacharunderscore}completeness\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -If we attempt to prove termination using the identity measure on - naturals, this fails:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{termination}\isamarkupfalse% -\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}n{\isachardot}\ n{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\ \ \isacommand{apply}\isamarkupfalse% -\ auto% -\begin{isamarkuptxt}% -We get stuck with the subgoal - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ nz{\isacharunderscore}dom\ n\ {\isasymLongrightarrow}\ nz\ n\ {\isacharless}\ Suc\ n% -\end{isabelle} - - Of course this statement is true, since we know that \isa{nz} is - the zero function. And in fact we have no problem proving this - property by induction.% -\end{isamarkuptxt}% -\isamarkuptrue% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isacommand{lemma}\isamarkupfalse% -\ nz{\isacharunderscore}is{\isacharunderscore}zero{\isacharcolon}\ {\isachardoublequoteopen}nz{\isacharunderscore}dom\ n\ {\isasymLongrightarrow}\ nz\ n\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ rule{\isacharcolon}nz{\isachardot}pinduct{\isacharparenright}\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -We formulate this as a partial correctness lemma with the condition - \isa{nz{\isacharunderscore}dom\ n}. This allows us to prove it with the \isa{pinduct} rule before we have proved termination. With this lemma, - the termination proof works as expected:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{termination}\isamarkupfalse% -\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}relation\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}n{\isachardot}\ n{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\ {\isacharparenleft}auto\ simp{\isacharcolon}\ nz{\isacharunderscore}is{\isacharunderscore}zero{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -As a general strategy, one should prove the statements needed for - termination as a partial property first. Then they can be used to do - the termination proof. This also works for less trivial - examples. Figure \ref{f91} defines the 91-function, a well-known - challenge problem due to John McCarthy, and proves its termination.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{figure} -\hrule\vspace{6pt} -\begin{minipage}{0.8\textwidth} -\isabellestyle{it} -\isastyle\isamarkuptrue -\isacommand{function}\isamarkupfalse% -\ f{\isadigit{9}}{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}f{\isadigit{9}}{\isadigit{1}}\ n\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharless}\ n\ then\ n\ {\isacharminus}\ {\isadigit{1}}{\isadigit{0}}\ else\ f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ pat{\isacharunderscore}completeness\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}estimate{\isacharcolon}\ \isanewline -\ \ \isakeyword{assumes}\ trm{\isacharcolon}\ {\isachardoublequoteopen}f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}dom\ n{\isachardoublequoteclose}\ \isanewline -\ \ \isakeyword{shows}\ {\isachardoublequoteopen}n\ {\isacharless}\ f{\isadigit{9}}{\isadigit{1}}\ n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{using}\isamarkupfalse% -\ trm\ \isacommand{by}\isamarkupfalse% -\ induct\ auto% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{termination}\isamarkupfalse% -\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{let}\isamarkupfalse% -\ {\isacharquery}R\ {\isacharequal}\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{1}}\ {\isacharminus}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}wf\ {\isacharquery}R{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ n\ {\isacharcolon}{\isacharcolon}\ nat\ \isacommand{assume}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isasymnot}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharless}\ n{\isachardoublequoteclose}\ % -\isamarkupcmt{Assumptions for both calls% -} -\isanewline -\isanewline -\ \ \isacommand{thus}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharcomma}\ n{\isacharparenright}\ {\isasymin}\ {\isacharquery}R{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\ % -\isamarkupcmt{Inner call% -} -\isanewline -\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ inner{\isacharunderscore}trm{\isacharcolon}\ {\isachardoublequoteopen}f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}dom\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\ % -\isamarkupcmt{Outer call% -} -\isanewline -\ \ \isacommand{with}\isamarkupfalse% -\ f{\isadigit{9}}{\isadigit{1}}{\isacharunderscore}estimate\ \isacommand{have}\isamarkupfalse% -\ {\isachardoublequoteopen}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}\ {\isacharless}\ f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% -\isanewline -\ \ \isacommand{with}\isamarkupfalse% -\ {\isacharbackquoteopen}{\isasymnot}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharless}\ n{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharparenleft}f{\isadigit{9}}{\isadigit{1}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\isacharcomma}\ n{\isacharparenright}\ {\isasymin}\ {\isacharquery}R{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupfalse\isabellestyle{tt} -\end{minipage} -\vspace{6pt}\hrule -\caption{McCarthy's 91-function}\label{f91} -\end{figure} -% -\isamarkupsection{Higher-Order Recursion% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Higher-order recursion occurs when recursive calls - are passed as arguments to higher-order combinators such as \isa{map}, \isa{filter} etc. - As an example, imagine a datatype of n-ary trees:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ {\isacharprime}a\ tree\ {\isacharequal}\ \isanewline -\ \ Leaf\ {\isacharprime}a\ \isanewline -{\isacharbar}\ Branch\ {\isachardoublequoteopen}{\isacharprime}a\ tree\ list{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent We can define a function which swaps the left and right subtrees recursively, using the - list functions \isa{rev} and \isa{map}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{fun}\isamarkupfalse% -\ mirror\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ tree{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}mirror\ {\isacharparenleft}Leaf\ n{\isacharparenright}\ {\isacharequal}\ Leaf\ n{\isachardoublequoteclose}\isanewline -{\isacharbar}\ {\isachardoublequoteopen}mirror\ {\isacharparenleft}Branch\ l{\isacharparenright}\ {\isacharequal}\ Branch\ {\isacharparenleft}rev\ {\isacharparenleft}map\ mirror\ l{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -Although the definition is accepted without problems, let us look at the termination proof:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{termination}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -% -\begin{isamarkuptxt}% -As usual, we have to give a wellfounded relation, such that the - arguments of the recursive calls get smaller. But what exactly are - the arguments of the recursive calls when mirror is given as an - argument to map? Isabelle gives us the - subgoals - - \begin{isabelle}% -\ {\isadigit{1}}{\isachardot}\ wf\ {\isacharquery}R\isanewline -\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}l\ x{\isachardot}\ x\ {\isasymin}\ set\ l\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ Branch\ l{\isacharparenright}\ {\isasymin}\ {\isacharquery}R% -\end{isabelle} - - So the system seems to know that \isa{map} only - applies the recursive call \isa{mirror} to elements - of \isa{l}, which is essential for the termination proof. - - This knowledge about map is encoded in so-called congruence rules, - which are special theorems known to the \cmd{function} command. The - rule for map is - - \begin{isabelle}% -{\isasymlbrakk}{\isacharquery}xs\ {\isacharequal}\ {\isacharquery}ys{\isacharsemicolon}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ set\ {\isacharquery}ys\ {\isasymLongrightarrow}\ {\isacharquery}f\ x\ {\isacharequal}\ {\isacharquery}g\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ map\ {\isacharquery}f\ {\isacharquery}xs\ {\isacharequal}\ map\ {\isacharquery}g\ {\isacharquery}ys% -\end{isabelle} - - You can read this in the following way: Two applications of \isa{map} are equal, if the list arguments are equal and the functions - coincide on the elements of the list. This means that for the value - \isa{map\ f\ l} we only have to know how \isa{f} behaves on - the elements of \isa{l}. - - Usually, one such congruence rule is - needed for each higher-order construct that is used when defining - new functions. In fact, even basic functions like \isa{If} and \isa{Let} are handled by this mechanism. The congruence - rule for \isa{If} states that the \isa{then} branch is only - relevant if the condition is true, and the \isa{else} branch only if it - is false: - - \begin{isabelle}% -{\isasymlbrakk}{\isacharquery}b\ {\isacharequal}\ {\isacharquery}c{\isacharsemicolon}\ {\isacharquery}c\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isacharequal}\ {\isacharquery}u{\isacharsemicolon}\ {\isasymnot}\ {\isacharquery}c\ {\isasymLongrightarrow}\ {\isacharquery}y\ {\isacharequal}\ {\isacharquery}v{\isasymrbrakk}\isanewline -{\isasymLongrightarrow}\ {\isacharparenleft}if\ {\isacharquery}b\ then\ {\isacharquery}x\ else\ {\isacharquery}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ {\isacharquery}c\ then\ {\isacharquery}u\ else\ {\isacharquery}v{\isacharparenright}% -\end{isabelle} - - Congruence rules can be added to the - function package by giving them the \isa{fundef{\isacharunderscore}cong} attribute. - - The constructs that are predefined in Isabelle, usually - come with the respective congruence rules. - But if you define your own higher-order functions, you may have to - state and prove the required congruence rules yourself, if you want to use your - functions in recursive definitions.% -\end{isamarkuptxt}% -\isamarkuptrue% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Congruence Rules and Evaluation Order% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Higher order logic differs from functional programming languages in - that it has no built-in notion of evaluation order. A program is - just a set of equations, and it is not specified how they must be - evaluated. - - However for the purpose of function definition, we must talk about - evaluation order implicitly, when we reason about termination. - Congruence rules express that a certain evaluation order is - consistent with the logical definition. - - Consider the following function.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{function}\isamarkupfalse% -\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}f\ n\ {\isacharequal}\ {\isacharparenleft}n\ {\isacharequal}\ {\isadigit{0}}\ {\isasymor}\ f\ {\isacharparenleft}n\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -For this definition, the termination proof fails. The default configuration - specifies no congruence rule for disjunction. We have to add a - congruence rule that specifies left-to-right evaluation order: - - \vspace{1ex} - \noindent \isa{{\isasymlbrakk}{\isacharquery}P\ {\isacharequal}\ {\isacharquery}P{\isacharprime}{\isacharsemicolon}\ {\isasymnot}\ {\isacharquery}P{\isacharprime}\ {\isasymLongrightarrow}\ {\isacharquery}Q\ {\isacharequal}\ {\isacharquery}Q{\isacharprime}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isacharquery}P\ {\isasymor}\ {\isacharquery}Q{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharquery}P{\isacharprime}\ {\isasymor}\ {\isacharquery}Q{\isacharprime}{\isacharparenright}}\hfill(\isa{disj{\isacharunderscore}cong}) - \vspace{1ex} - - Now the definition works without problems. Note how the termination - proof depends on the extra condition that we get from the congruence - rule. - - However, as evaluation is not a hard-wired concept, we - could just turn everything around by declaring a different - congruence rule. Then we can make the reverse definition:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ disj{\isacharunderscore}cong{\isadigit{2}}{\isacharbrackleft}fundef{\isacharunderscore}cong{\isacharbrackright}{\isacharcolon}\ \isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}{\isasymnot}\ Q{\isacharprime}\ {\isasymLongrightarrow}\ P\ {\isacharequal}\ P{\isacharprime}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}Q\ {\isacharequal}\ Q{\isacharprime}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}P{\isacharprime}\ {\isasymor}\ Q{\isacharprime}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{fun}\isamarkupfalse% -\ f{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}f{\isacharprime}\ n\ {\isacharequal}\ {\isacharparenleft}f{\isacharprime}\ {\isacharparenleft}n\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}\ {\isasymor}\ n\ {\isacharequal}\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent These examples show that, in general, there is no \qt{best} set of - congruence rules. - - However, such tweaking should rarely be necessary in - practice, as most of the time, the default set of congruence rules - works well.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/Thy/document/session.tex --- a/doc-src/IsarAdvanced/Functions/Thy/document/session.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,6 +0,0 @@ -\input{Functions.tex} - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/conclusion.tex --- a/doc-src/IsarAdvanced/Functions/conclusion.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,7 +0,0 @@ -\section{Conclusion} - -\fixme{} - - - - diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/functions.tex --- a/doc-src/IsarAdvanced/Functions/functions.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,96 +0,0 @@ - -%% $Id$ - -\documentclass[a4paper,fleqn]{article} - -\usepackage{latexsym,graphicx} -\usepackage[refpage]{nomencl} -\usepackage{../../iman,../../extra,../../isar,../../proof} -\usepackage{../../isabelle,../../isabellesym} -\usepackage{style} -\usepackage{mathpartir} -\usepackage{amsthm} -\usepackage{../../pdfsetup} - -\newcommand{\cmd}[1]{\isacommand{#1}} - -\newcommand{\isasymINFIX}{\cmd{infix}} -\newcommand{\isasymLOCALE}{\cmd{locale}} -\newcommand{\isasymINCLUDES}{\cmd{includes}} -\newcommand{\isasymDATATYPE}{\cmd{datatype}} -\newcommand{\isasymAXCLASS}{\cmd{axclass}} -\newcommand{\isasymFIXES}{\cmd{fixes}} -\newcommand{\isasymASSUMES}{\cmd{assumes}} -\newcommand{\isasymDEFINES}{\cmd{defines}} -\newcommand{\isasymNOTES}{\cmd{notes}} -\newcommand{\isasymSHOWS}{\cmd{shows}} -\newcommand{\isasymCLASS}{\cmd{class}} -\newcommand{\isasymINSTANCE}{\cmd{instance}} -\newcommand{\isasymLEMMA}{\cmd{lemma}} -\newcommand{\isasymPROOF}{\cmd{proof}} -\newcommand{\isasymQED}{\cmd{qed}} -\newcommand{\isasymFIX}{\cmd{fix}} -\newcommand{\isasymASSUME}{\cmd{assume}} -\newcommand{\isasymSHOW}{\cmd{show}} -\newcommand{\isasymNOTE}{\cmd{note}} -\newcommand{\isasymCODEGEN}{\cmd{code\_gen}} -\newcommand{\isasymPRINTCODETHMS}{\cmd{print\_codethms}} -\newcommand{\isasymFUN}{\cmd{fun}} -\newcommand{\isasymFUNCTION}{\cmd{function}} -\newcommand{\isasymPRIMREC}{\cmd{primrec}} -\newcommand{\isasymRECDEF}{\cmd{recdef}} - -\newcommand{\qt}[1]{``#1''} -\newcommand{\qtt}[1]{"{}{#1}"{}} -\newcommand{\qn}[1]{\emph{#1}} -\newcommand{\strong}[1]{{\bfseries #1}} -\newcommand{\fixme}[1][!]{\strong{FIXME: #1}} - -\newtheorem{exercise}{Exercise}{\bf}{\itshape} -%\newtheorem*{thmstar}{Theorem}{\bf}{\itshape} - -\hyphenation{Isabelle} -\hyphenation{Isar} - -\isadroptag{theory} -\title{Defining Recursive Functions in Isabelle/HOL} -\author{Alexander Krauss} - -\isabellestyle{tt} -\renewcommand{\isastyletxt}{\isastyletext}% use same formatting for txt and text - -\begin{document} - -\date{\ \\} -\maketitle - -\begin{abstract} - This tutorial describes the use of the new \emph{function} package, - which provides general recursive function definitions for Isabelle/HOL. - We start with very simple examples and then gradually move on to more - advanced topics such as manual termination proofs, nested recursion, - partiality, tail recursion and congruence rules. -\end{abstract} - -%\thispagestyle{empty}\clearpage - -%\pagenumbering{roman} -%\clearfirst - -\input{intro.tex} -\input{Thy/document/Functions.tex} -%\input{conclusion.tex} - -\begingroup -%\tocentry{\bibname} -\bibliographystyle{plain} \small\raggedright\frenchspacing -\bibliography{../../manual} -\endgroup - -\end{document} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: t -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/intro.tex --- a/doc-src/IsarAdvanced/Functions/intro.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,55 +0,0 @@ -\section{Introduction} - -Starting from Isabelle 2007, new facilities for recursive -function definitions~\cite{krauss2006} are available. They provide -better support for general recursive definitions than previous -packages. But despite all tool support, function definitions can -sometimes be a difficult thing. - -This tutorial is an example-guided introduction to the practical use -of the package and related tools. It should help you get started with -defining functions quickly. For the more difficult definitions we will -discuss what problems can arise, and how they can be solved. - -We assume that you have mastered the fundamentals of Isabelle/HOL -and are able to write basic specifications and proofs. To start out -with Isabelle in general, consult the Isabelle/HOL tutorial -\cite{isa-tutorial}. - - - -\paragraph{Structure of this tutorial.} -Section 2 introduces the syntax and basic operation of the \cmd{fun} -command, which provides full automation with reasonable default -behavior. The impatient reader can stop after that -section, and consult the remaining sections only when needed. -Section 3 introduces the more verbose \cmd{function} command which -gives fine-grained control. This form should be used -whenever the short form fails. -After that we discuss more specialized issues: -termination, mutual, nested and higher-order recursion, partiality, pattern matching -and others. - - -\paragraph{Some background.} -Following the LCF tradition, the package is realized as a definitional -extension: Recursive definitions are internally transformed into a -non-recursive form, such that the function can be defined using -standard definition facilities. Then the recursive specification is -derived from the primitive definition. This is a complex task, but it -is fully automated and mostly transparent to the user. Definitional -extensions are valuable because they are conservative by construction: -The \qt{new} concept of general wellfounded recursion is completely reduced -to existing principles. - - - - -The new \cmd{function} command, and its short form \cmd{fun} have mostly -replaced the traditional \cmd{recdef} command \cite{slind-tfl}. They solve -a few of technical issues around \cmd{recdef}, and allow definitions -which were not previously possible. - - - - diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/isabelle_isar.eps --- a/doc-src/IsarAdvanced/Functions/isabelle_isar.eps Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,6488 +0,0 @@ -%!PS-Adobe-2.0 EPSF-1.2 -%%Title: isabelle_any -%%Creator: FreeHand 5.5 -%%CreationDate: 24.09.1998 21:04 Uhr -%%BoundingBox: 0 0 202 178 -%%FHPathName:MacSystem:Home:Markus:TUM:Isabelle Logo:export:isabelle_any -%ALDOriginalFile:MacSystem:Home:Markus:TUM:Isabelle Logo:export:isabelle_any -%ALDBoundingBox: -153 -386 442 456 -%%FHPageNum:1 -%%DocumentSuppliedResources: procset Altsys_header 4 0 -%%ColorUsage: Color -%%DocumentProcessColors: Cyan Magenta Yellow Black -%%DocumentNeededResources: font Symbol -%%+ font ZapfHumanist601BT-Bold -%%DocumentFonts: Symbol -%%+ ZapfHumanist601BT-Bold -%%DocumentNeededFonts: Symbol -%%+ ZapfHumanist601BT-Bold -%%EndComments -%!PS-AdobeFont-1.0: ZapfHumanist601BT-Bold 003.001 -%%CreationDate: Mon Jun 22 16:09:28 1992 -%%VMusage: 35200 38400 -% Bitstream Type 1 Font Program -% Copyright 1990-1992 as an unpublished work by Bitstream Inc., Cambridge, MA. -% All rights reserved. -% Confidential and proprietary to Bitstream Inc. -% U.S. GOVERNMENT RESTRICTED RIGHTS -% This software typeface product is provided with RESTRICTED RIGHTS. Use, -% duplication or disclosure by the Government is subject to restrictions -% as set forth in the license agreement and in FAR 52.227-19 (c) (2) (May, 1987), -% when applicable, or the applicable provisions of the DOD FAR supplement -% 252.227-7013 subdivision (a) (15) (April, 1988) or subdivision (a) (17) -% (April, 1988). Contractor/manufacturer is Bitstream Inc., -% 215 First Street, Cambridge, MA 02142. -% Bitstream is a registered trademark of Bitstream Inc. -11 dict begin -/FontInfo 9 dict dup begin - /version (003.001) readonly def - /Notice (Copyright 1990-1992 as an unpublished work by Bitstream Inc. All rights reserved. Confidential.) readonly def - /FullName (Zapf Humanist 601 Bold) readonly def - /FamilyName (Zapf Humanist 601) readonly def - /Weight (Bold) readonly def - /ItalicAngle 0 def - /isFixedPitch false def - /UnderlinePosition -136 def - /UnderlineThickness 85 def -end readonly def -/FontName /ZapfHumanist601BT-Bold def -/PaintType 0 def -/FontType 1 def -/FontMatrix [0.001 0 0 0.001 0 0] readonly def -/Encoding StandardEncoding def -/FontBBox {-167 -275 1170 962} readonly def -/UniqueID 15530396 def -currentdict end -currentfile eexec -a2951840838a4133839ca9d22e2b99f2b61c767cd675080aacfcb24e19cd -1336739bb64994c56737090b4cec92c9945ff0745ef7ffc61bb0a9a3b849 -e7e98740e56c0b5af787559cc6956ab31e33cf8553d55c0b0e818ef5ec6b -f48162eac42e7380ca921dae1c82b38fd6bcf2001abb5d001a56157094cf -e27d8f4eac9693e88372d20358b47e0c3876558ebf757a1fbc5c1cddf62b -3c57bf727ef1c4879422c142a084d1c7462ac293e097fabe3a3ecfcd8271 -f259833bac7912707218ec9a3063bf7385e02d8c1058ac06df00b33b8c01 -8768b278010eb4dd58c7ba59321899741cb7215d8a55bee8d3398c887f02 -e1f4869387f89141de693fcb429c7884c22dcdeddcaa62b7f5060249dfab -cfc351201f7d188b6ed68a228abda4d33b3d269ac09cde172bc045e67449 -c0f25d224efbe8c9f9d2968a01edbfb039123c365ed0db58ad38aabe015b -8881191dd23092f6d53d5c1cd68ebd038e098d32cb24b433b7d5d89c28ee -05ea0b6070bb785a2974b5a160ee4cf8b6d8c73445d36720af0530441cd9 -61bc0c367f1af1ec1c5ab7255ddda153c1868aba58cd5b44835535d85326 -5d7fed5ff7118adb5d5b76cc3b72e5ff27e21eb857261b3afb7688fca12d -1663b0d8bdc1dd47a84b65b47d3e76d3b8fa8b319f17e1bb22b45a7482fd -f9ad1b6129e09ae47f15cd2447484cd2d64f59ab0f2f876c81e7d87ccdf9 -005aa8fc093d02db51a075d571e925f2d309a1c535a1e59d34215c6cd33e -3c38997b4956461f376399901a8d0943dca6a335baac93fc8482c0659f04 -329c6f040b35828ea6dd1bd1858f2a9be4ef77731b5b75a1c536c6bc9479 -0821e5d88f6e2981835dbfd65ec254ebcf2cf49c917d121cd3bbb476a12b -69c15f17d9c17bb15ad1e7d31d2afcf58c8f0ad526d68615a0f1ac3b1d1c -d3beafeea3cf56c8f8a66367c70df9159f0b1b3157ccfd010045c4718e0e -625c0891e85790c9b97b85517c74c9d55eaca31a01cddc64898bf0eeadf3 -53391a185e507dcb0a6f52661a56357ac818dfc740a540aadf02f4e7a79d -8008a77cd30abde337025b01217d6a68c306abe145b7260c3478fa5f366f -b2d37259ead8a8ec2db2f09ae0eb3a682d27b0d73a60935f80254c24426a -003a87a29a4370cbf1b2ef1e19ad8466ec725fd5b463d06011a5e0da2258 -ff6c1483c4532bc21f2ed9b99b929b2800ddefc1a98d12ba085adc210bac -e0274b69e24d16af015a51ca73edf779a7534a887aa780337ad966839881 -edc22ba72038aa1a96a6deba24ad676795da711b92f8cf5f54cb4322ec04 -40ef9e15b11e3005f3ff69376ecb29bb66e8fc1b685f2b05fb162fcb35aa -d7eb2a8ec39b97ab1ff05ef02f8dbbc12085a5cd252cc4010fab7f63dfd5 -7fa1be86f724d37db5faef17ae8e8e537444e8e9db917b270344183473af -7f47d5703a160d8ef1e772438620d3336b2fbcf6433612e4b5e64fae0329 -8a3c4010c17d93f13ba66d549c69dd58c7d26ddc90285fed831918563a16 -2a7ac2511e2f18c9eb3df905a9dcba65a31cc1c39f07458abb11b4c60346 -aea19070e126982f1dde336e79be0ecd69a8afbe2493d064d4d2ff38788b -b3038125961302db9761403c3b8019ec641e107425002205a77ae2ae0f4f -7550d623dd03f0ec0199f42a9a8b89514e2e21baca9b3c8c96ca48cbf9dc -ee6d6241d713e014b49e83ad85e62a6b2f70b58e4cc72d41ea6fcbdd3b5c -890c8af0d24200658773b1628c6cc9aaaabb08865ee4c7ff2c513ad7aa23 -155a321213fa94731683a3e282a0e60aa3c87aade3da231465bdd9097f2c -89a1af8e5b9649143f2d9482546501ea97e8bea2f5d5eea97d4f19bb6835 -3138d3babb2461c08d537491aaede1f23d734c6f099eb5bef6e2ffaaf138 -e5ab71b8b41599091037e440127a4eaedf208c20c8a2fc62eadab191d1ab -4d5531f826aa6b9fff2797a7f54673e0a3fae09a93a0dfafb8b11d60dc69 -5acf9b7e1a47c31d0b5a0b85b7b50cddff5ac831651d9c7469c2139c7a89 -7d2f868f36c65156921803eccfdbdd1618595ab6d2a9230ef523a1b5ee51 -f2a0d200fc0e94aff7f546593ff2a3eb865d129895af01b8ab6e4616fe20 -9123b6e2b7e0817adc3cdb78ae8b0b1d75f2986ebd8fb24c4de92ac9e8c3 -6afa520636bcad2e6a03d11c268d97fa578561f6e7523e042c4cc73a0eac -7a841907450e83d8e7a8de4db5085f6e8b25dc85b59e018380f4b9523a7f -02cbeec989f0221b7681ec427519062b429dcd8fc2e4f81173519f88e2e4 -3798b90a84060920f6ae789afd6a9182e7fec87cd2d4235d37a65c3f3bcc -c742c89cbe5f3e2ba6c4f40ebba162e12569d20684cc167685285a087e7a -0a995fe1939bf25c09553512ba2cf77ef21d2ef8da1c73ba6e5826f4099e -27d8bc7b3545fc592b75228e70127c15a382774357457cd4586d80dc0bd6 -065aee32acfd5c0523303cece85a3dbf46853b917618c0330146f527c15b -dbb9f6526964368b2b8593eed1551dad75659565d54c6a0a52da7a8e366f -dd009ef853491c0fb995e19933cba1dbdc8902721c3ea6017ffdd5851cb8 -3c8bada46075ac121afe13a70e87529e40693157adcc999ed4657e017adf -f7dbac4bc0d204f204c6f47b769aaf714f9ec1d25226f24d0a1b53e28ac5 -374ab99755852c1431b2486df5fd637e2005a25303345a1c95a15a1189ba -f6f6883de1ad46d48427b137c2003d210ab2b2f5680f2633939f289d7553 -eb943adf8127f1c3ee7d6453b5566393700ad74ab86eb9a89f8b0380af55 -6b62f51b7dbd0c5dcc9a9fb658944d7ad5845d58dedc2d38200d0ef7cb0f -76041dc104ef3ab89c1dc2f6a75635d48051c8a7dd9f5e60253a53957ec8 -9d1266566d7ed20d79dfc2807b397d7cf056bdaccdb72528a88aa4987682 -c909b2fe1e35a71c2f29e89a2bf32173967e79610367ce4574ba6a1cc031 -cfb176fc0313f74f91a866ef9954b95b29caf917a6b919586f54d23cb7ce -23305886ae7760ebd6263df0d3c511ac7afc361df78bc2621f66d3268b99 -078fa59124f0eb9476496c938eb4584e87455dc6f2faa999e938460b31c6 -28021c652acfa12d4556aa4302bbcd043e60389480b796c3fc0b2e51b81e -c2afa4a34335318a1c5a842dcaa120df414acba2e79ab5cc61c99e98108c -5cb907a96b30d731131782f9df79aabfc16a20ace8852d047497982e11c8 -26321addf679de8a96a2d18743f6f2c3b2bc397370b10ad273fcfb76b48b -9dad27cf89ca2c370149cd48ab956e9bbce01cbf8e1f0c661b99cf19b36e -87b6165dd85ae3f3674525e17d85971093e110520d17f2d6253611e35ec9 -e17000e48e2926454c1e8eb4098e0115b49536f2c33505eb5e574f7a414b -e176398c5ddf6d846ea5ddf2a5e94c0422e0115c57a8c9e56bf8ba962c82 -698c96bd6138baaca7347e44352cc62df4eeba364954ad921a5a43280980 -264df4a7fb29d005423179f7bd1d98b4280d62ce23c927551f1ffc2b8f17 -0a9c23656c0c91b640cdcfdbd88089ffb28d3ac68bad25dbbed82c083696 -1f9f86a6183cc1061ffdb32279796569d05b31c946955402d0be1fb9f2bf -304d1ad8e1e357be49e6e2ee67f7f1e7bc699d46a5f7853fe659ba2e1930 -0d3e3ea658b9862701dcab08fdd23bf1d751777f25efbe9e02d12b5612b3 -c3fc2275190346b94ec4024e4ade08e54d75c0b4c48f4956b9182e8ce997 -74b54da4a9318c099d89f1ce3b6803a95f48b9fb8b845372be25e54478e8 -49e4707ea03a36e134efa661e4e6250d89649ae074cfd9d6b9e2071c1099 -3b8a5a5ebc3e1cb228c97565aef7f254e3f90af8a3dd281c83792755719d -c6a5b3bab4aa6be5afe9624050eee8dfb13b018f4088c932cd48ace38dfe -b1b4218dba8f7fada6628076acf1b54db0c95d4fb12232f1fa9d1ba848f9 -fe80c65b75d6946c00fe78839142c5302707d9269b24565dbcc551aca060 -b4d0f99b961dd3cc795a982063ac42e9fc81fc98add42744a9f92e74b00d -637ee4606ea2099b6c763493c3159f8e52a90dafca682010c0e92bc9038a -10abb066c75c8d97f7ad6fb1a37136e52cf2093c4fa485fe12adad10e4d0 -83b78b023628ddc5326cbf8392516027a9f3de4945f93488e4a1997efd2a -22c2c136dbac1bdb829e082beac48cdd06dcb17bacf09451c7b636bd49a8 -fc60cb1d2292250bea78e1dd276725ab4c526b66ddabf4e1b2bf0a2571df -2490df70735f5da321fac74fe4fab54444e53dace9832cff38a70c58104a -4f0c0545dcf7a1a9ecb54e0e32d6d8195d30b2c98a567741fcf325a4ddeb -244a1a35676e246ddc835fac13b569f35f22ec93191eca3efbe17ff9a950 -d08f863710b2bbecec969068c498eb2338b68f3fc3f5447449fe4de2b065 -e068ecd9255d369b2bb6c4b0b7423774bed69294758aca2bdb1a8d5bf618 -d3fa09462f7838e8a79b7a53bebe6dacb0a1561eaa074bc6f069f6a06fb2 -c4a5cb13e2172bce9be659a82665da5cded73da84322bb16aa6e19ac1958 -7515cb5d2b65e65e803f76700ce5efd3df7fe4ed431fae0e8e286b1d5056 -a0d18df926b2be7a93c503ab903abd4788680a6201fdc299f2cb5d6a9b6e -2048109c8d1fb633a54128938594b2cce86a7e0185e7d59e6536584039ec -9e30ff7be6ddba9fdba82de7415fdc47de84d97afb1aa3ba495bd91dee9d -f3b21ee1164987dd8510925087cd0919f1085cba6e4dd3c7384d26667f94 -ad7f736a60d8bd76dfaa4b53c23536fc309ff2317c48ee0107ff2ca4d1b3 -f78c5a27b901c931128bdb636094ef0cd543a5b62b7dbe10ed83aed5780c -a16067a4a7e8b7c5bf8a8e822149bc1b5bcdabe13a7f6aa6eaeff24a42f4 -a58a2b70f545103540169137fda9abb589f734b6776cb50402d6123ce802 -10dce05e3697a98c9411cf60a02c278c91e03d540b936cd00c668960e357 -1aeaf4d94cfb496b259ec0d8fdba9199fb46634ff177bc8d310ea1314eef -d46c927a981c58e88743ed4e07d80fe841edee812e3053412bf2e710146c -b25dec8ea70c38bb1f6e4db3c2e8ba521963c1584eeb60ea1e9555058f13 -e98307c13cbd15c26b611f543149b1ddf88dd6296ae703f58daeb67f1b03 -ab5b24c72d5770cb9d8ed242c4faaad1dd940ada00e98ff3a61799d13355 -aba916910aa9a6e5ee8af438d0ba8235346fcd139b9d2cb7db7bd3f298a3 -94ff0aff3b9042f32a004e042c346a5ea35917f229848a9c5a32909b0090 -4aa715640277a6ada99f8b2927fda22899ff1347f42bac73e2bd4bbf3945 -55fd7dd30d5c6dadf5c259fdb2455d8c607de1c5da588e20159f18e4e8da -b714e532d888a0386c52c2b85964251af003ac9d10c0c8b9b3465e1dde48 -2e74a29e17a7cf6c9a43b5af1646f0b8508f98e9a1279ec3908073d89dcb -aa093e8dd1004c1ecccce0803095b0069d4be7a1eb14b02efc37d137dfe3 -f0190bc9628069abc257f45d0e050e60c7f5281277937dd986fcd5b94a2b -845a1a75addd74a142800f18ef408c08a2c2ad16a93298f148c0ae7d2990 -ded147f92f053809a60d4f992a227167aad5b5eb2bbe8a4a77dc77a08b72 -6acb34422e2532eec7e274e4a42d71ee584f5f3df5a3a5b379974ede73ab -5f1b0a2dbfcc8cfac4747ca26eb6030dc2f85a104b754432977850c093b9 -97ed90af711b544ff682e7b1eac82b2b7b44014b09c17ecf084c994a095d -9eeef5391305caf093b62ac9916f982a436d521fcf4d75c5b8e4d92266fd -e99a58aa39d7693ecd1796b8851761d64bbca39a6d5a0b4533ae47123327 -f98d0ad0e8b36625cc3647b55459552906d8a1d5766845ffac101980efcf -79657e365510be5db557cefef21193ca3cf3dad175ee2e7ae91d174fdc06 -2ff5c51ffe2f021122e96df042019d3a1883e662537ec1b69c11fbb6e750 -0132eabf5803c816153ecbff60ca3b3b39708c26cb1751afb2e65d8e5f4a -c4397a88fb1f112672fcdd24e4ba545c5b2a7968c17b62f8e2530a8acbff -cfca82c64b7abcab84e2c4a0a7ced67b15669301fe9ff2c756e70ff7ce33 -497be6acc4ac5617e1f043bd8a87416299a39bf17fc31c02d72d75fdc2a1 -e60669fa4d5e4a49d9afea2f53f4626680e9c0dfca223529efa415c4343a -b6067aa800c484457ea050eaaa5d3fafeedd0eec72f327e02c6b3912b5a8 -c404de4839c9c4a99da42681cde43316606a34c7d2f02269de1aab776857 -e668f35946af4d618d36d444bdc02b1f63ea25b6260b4fb606ac8575b5c9 -782a5de4037350d5753b1537537ccb008c454eeb264e6cd4727c999e240e -0ac89e95a896b67d54910a3531345f64198ad394b5ceb52881f1dd9e6beb -95862dc188d45b3e46aacb5fe40097947dab9bc3c1ee46bfc9b1b3ed6167 -efd0d65ceb043d7b24c1456676e4baa47b1209a315f199bb3a91f4374cd9 -cc0b40d3f09f19f8dd8a46915eee55eeeeb3c7b8f437106ee491ef0f4ff9 -2c5c6f779e0fbe7bd5293964bb645ca362b106abeb773571d9ae83b786a3 -d5a4ea3ea970daadc46cc5e6037f76fd20e0fffc47cf4e7af9522b91f96b -3297720fd45d9bc2200622ad2ca9445556c8a8202b1991bc63da360d021d -55be2528e043f803e08da99b91ab9cfc5e65b2655d78206b4aecd445a7b0 -1caa0d06b0a55e8f04b70b60b04a860c8e1ab439f4910051e3f7441b47c7 -8aa3ab8519f181a9e833f3242fa58d02ed76bf0031f71f9def8484ecced8 -b6e41aca56176b6b32a2443d12492c8a0f5ba8a3e227219dfd1dd23fcb48 -fcfd255dbbf3e198874e607399db8d8498e719f00e9ed8bdd96c88817606 -357a0063c23854e64ad4e59ddd5060845b2c4cddd00c40081458f8ee02c7 -303c11747bd104440046bf2d09794fca2c4beb23ed1b66d9ccb9a4dd57ad -a24943461ecc00704c916bdc621bfddb17913dfb0f3513b65f3ab015786a -caa51ee9546bc8ddf87e2e104137e35ddf8f8d23724e9a53824169bc7cfa -99562656e6f1c888d4dbff0b269c5d1e733e5f212d91297610201eb43249 -35e336dd0052738db2d64f3e89429903bb5c1810009cf766e9a06223dd2f -219b706394a121dc029af55c6ada9052af59682ef7c51e121cf16f0319ac -0aa9512ef900c548d673fe361da19052808797e958209072e145d46ec8cc -a89fafd76630eff30ae979973bdf0f8c9e469d8edd3b1c93731c72f976b7 -d81142bc15c376403f967affaa5f482efd57c6f91970729d16db851f0ed3 -ea7d82f409307b5b436886c1beda94a1fa3ab1b60686f6574c844fb2c0b3 -a07174dc4f27b4fed2f8bd4d5436be4b343e5efdf0400d235bd910255341 -a20770804a26f8437e9bce6da8e9f8258a343c7aee291f1510be306ae67a -ab1d7696453530c02fd153bbe49dbf62baad6146029cbd1656cdb76c952c -b93edfee76fe33832930be59636bb947e8ad285f20f663cccf484fba97d6 -7446c7b6c6f5857428bb1737d9ae801df75d9cb4d7bd59ef7a4cbadde928 -38f15d232005585d2e40483d2d3e08cc8f398bb43afedb84343c3ba3835d -0ba82a86dce859cf655f85e63e41365e0dbefcf511b9a27a2b6e66b2ad3a -c657902842287a317e46ceaa93b5088f09d53a65815b44538af90ad3b06b -4e5e2dc509f02e30a01e05201c67d4d39582bbe64e20b669f5fd787909a3 -30fc50a95b31426bbb57a4fbf9feacdc31f98bcf50da7e50c2bfc169c6fd -ccf213cdb878653bcea372968ea6e31fd30dd55434cc91c0af22179ce669 -a05493f195e12432c6173ae2ac3c94fb83f38210014a9f969ea2b44e99f5 -e5a7317e848d429ad62167a4fc5001149676c0c28cdf59b8d1c5a582f516 -3eee855312777fee6dacbf993f5c058f355dbde6552dc960d336eee445dd -11d53fd21b745d1e5ec317efbbef25e070d0a36797a6081c356ac2328e64 -e5c55fbc81dc75d9c1575548ece74b8307eef485aa8e28859a2e0435c831 -23a600efb323c362fe9f02407a5411c41a69566cd50add324b63ab939980 -b9d7a929ae4887163cfa7acbfc9fabaab8987a1f6906b9881491cd055b94 -485c968479dbb05b34ed0cd6844729a692978c6928c3392e33e8324ded88 -814cacdac8128e1425c0091a13558100d7cdbed5992795d94d39c32f32dc -621ab6f3b75187a66741f61d6a9c91d791b1cfc3d0e94d4a76302e0c3f2e -cbdc51f14f3251aa5c8bb989f0e13ee500b7b7f2f1e52ca970ad8a7b4b99 -57e93126254297380d67179deb8ff1e99d5cdf7a35c5bb9fa7b402e63234 -78640344e1f10c378ad23c5cd1aa18e1e0b308db70d3a624a455f8e291a2 -ee102ad10776208c2d546cb76d89ca8103a8b95f8acc2d2bdc9791324915 -6c9e05429091071f0c5b76d82c8d1c8a69d840fd460922cd2090624bc218 -0c9391005926a25042a55e322060807363462e1cdeab309033124ba3a884 -1db13f39edae04ec52cde9dbde64ddda1ad805141d4230ec76bd81fd98d7 -0d90fa1aaa26ea551bf687ddd6cdcf3de5a446b266c68434f07d9c0b382d -5816c4e22f22cc03ff78064c6dffb12315c6bcbbf5dc510f5aaabf23471a -234efceeb4aa2f9af9ea787c014c5587ef162fc5b35e8f4c23b168c6e247 -41d33dcc11d2a56d3ba9d8eed6e79aebf9f0faf1a3aeb89d792d69041f0b -b8fadfc0aa090effc6ae5e2f13cdbf54b5bed69b039eef2627769613b6f1 -aefe9b66747fe8feaf7455796740f411a770d4a1764f0483719584880f45 -430e38d3af184145892a08b2add234a3f3ee4ccfc9f6995c02392adafccd -722f366d748cfe9373fbf5f878ed47e9d221fd156bb28369df9e7d2b79da -76120d135ebaf36cff93beb7e313c2b2de7477176fc19609a1b906c995cd -defef08899265b6b8aefb44da1aadefd1c523dce5ca1b84c0c652b3009fd -057789892d4d31764f181754b2e0a62c465587585509989a219711a5e4e2 -5b3b340ca8fdd3f04fef204b1b722b2f6c2ccb00c3cf1a94ba9bdfbfeda9 -e2a062c6f1ced3b8aae5dae32ade1fca1001f98d0ad0e8b36625cc3647b5 -5459552906d8a788eb8bc734ccb65fe9582c71df94fd95d22c5323de235c -28220fb9a2ccb37362174d8cd5922c9e5a87b51d0668555100a33e33750e -f1f795cbed962494a994be7ce8cf71fc58ff4204551b1615ed27cf088171 -fd000b72462b67935961e7c6c3a05bfd67b9ba094ea2c16fdf486da912e1 -e97bfd1c17934535e551cede20c001b5d2adb2be4cbad7d6ba0bdeae4b1a -a739f90293e67ecbdeea4d35825e092697fb05b215083e3f3d6be260790e -2a175fd44eb1c4c16759504827a6eb58a838c4d65fec6eef108495577019 -15740cac164111892e8d1cc447cd208e243a89ab847d8ebf4fb98bff49e7 -a3453facf3b0e8cb67590f390173ddba68324531d2e426aed152e12301d7 -538c1f3c0048a9cc00c009a1a9138460082123209c1e007266fbf236eb72 -21f87d4ca38a0b699e84ca230ffb5095f90a6528bf2a9118f95ac9ab8d2d -ed9eed9b8b27be894b717469758c8d94fa89acc64f530f432d0e5f16c922 -36d6a63410f099c9e909450fd731d698ef658d8ffc1de14817b850814f68 -1a4a9be5cc7a71c381974c249f0b209bfdc2e97f9540c96f57bb4d283622 -00969b82011315289e6a025b137030a0af3b4b42b00fed7cec49df43c59e -3b2495a036dd1b17a8e6adae63bfbbd48807c44b5bbf71813355e1b0e58e -22b6fb88005fc55565be49c17244901b73ef02fc4eb7669be5af22d89c0d -dff0fc6821d810d13e5821d48d4a71d7e463d5b60bc279d0dcf5f8da3a95 -905b56d6f2be95e6d4243b1048e3b662e62401ffaa3bc3f5f90b0854b8a3 -8c38039f61fcb359b06bbb7d59e3b39a295dccd6db9a8b83a6f64ef8dc94 -a77123dd164cfd1c46f1ee51aa19c3d6e7db92a298d10159f2b5eff2caf9 -dc93a6d267fb65bd900d6adf0c6be598050b6d3a9b3a322ab3c9e880d774 -1f58016ff97e5f606b5dbd72ba99252c669209bb556dd5be84fdd7c1ce92 -8a3b3d3aab8d37e6b740227563bb4d60f6bb04052356e1a48d2079feca44 -7ea17fd06f208426d045dee660d1d6460455f8d20dbc5ae64550bbdf60d7 -27d96cc9afef842a8c8c78ea2257e6c6d0d207c80cfe399e8874c693274e -d2c2022d303ca50a70624b07434fb85040a76a823f446c7454dab4f9c05f -10274eb5ba164aa3649d1bc90694316ba5cb3e7df4442e777124cff7ebef -53df2320a0c441ab61666493cb43da46d5711c21699de85bc74359444da2 -e3e397d4c16234f81531505b621aa242a6698886f82b447104b1f1062f60 -b5c87cea9151bb3c627bfa4532b06fd147c556ed8d61ae30a8719dfb8705 -f8a6c74368381403640cc57026d3790c49e2bbd1c0e48285ec6ba44de678 -e3a1394d659c412f09644b83ee1a333a1f51ad8deb4e6d77b3b226ac2c4f -fe653411a7976ae7c4a3cb7df309788da6b483f8a7bab4a6990db74362f5 -bc41d545a320389b2599fd726e426ed9fa2916ece67b058f6a269544e517 -128bda38d117f402409d0d8f8c88ed509aa2ba882e0c579b45af4be80770 -22d7269684eaf0f9afc3054316da6611e3fd260d67fb6fe52c9ade5dda24 -a0050a819ed21342aac9d25194778beb3145f56a66980f620998923521ea -3f957b6ed0c5470734af9f416a16427dd03eff9a0e023452097d4ef936d5 -49a90823cef6de340a1ee02a52851b310cbcf41ae274947a62f9d1d8702a -669023e3caf967204a340694b45fecbda4bf9552f6bdc62d43b3b2c3d571 -9983c182453e22ee34241ab908e667115f7988174684cd70084aefc55caa -f5352a88e9dac45d1ea0e032af61fe9a9118a3931b2050fc6db66ab96a39 -74353b597f34dfd9f72150de23285eda5e555a607d198c291965a7233715 -3f4946a57af0b440ff8567b01a6f46c6d32fea5f8bf57d89dccbab7da882 -ee6c9260e89443b1d7db099477492bd0468850df3db668d741123e7ebe3d -c21748ab4c5cbeb5de33b8963aecafe76bba0c4f6ed8e8263a116ed85e58 -fb71ec4ab0071301be7c7d3afd5fa6ad46c0232807bb7fe129e44bfd16e9 -fd0c8bb5e7cdd86a78b5fb0669093c22eda9151d85b6f58a9c8ead3727c0 -09850bd31a8b4a873d0a506240bb2aeccb8dcb6369532f21d9b967aa8443 -fd6d77cb2d65c4678a5fad188db85940f0a187aa1031dcf5b8e0d0cbfb6d -b3b96fedec5b249b7a69de9b42dfa605bd622de7a220cce9b66e9f3394d6 -13487dc5e82c1e619079cd057b1e19ac05ebdfd7c8bf01c6c66fab49e0b6 -613df9e42beae2f7b9407a2bff8896d8035cea0fd5c11bc5889cb3d90876 -61766138d2625f42d0244adca65d1bc73989328c0eea0b97c7c766285ab3 -351ce2b183f774488a8806c33178090a3808f0ce5e339b87cf7add933301 -ca486742831ca751f0626864ce13172829a8419af5c78794a0eaa17b5bcd -fcb684f7d4bb7af15deb432e44dc7dedf56eb8bea08b46f1e8123a49a349 -a7cbccf833a528f5e22d2d463040e09b91e543a2f33077b3e7b9ecc64f14 -306186cdae1fc317a6ced7e9b4d51a10bbbcf2fadff876b4d9082e3f4aef -dfef230e4232572f4fa33a6e065f6895aa2ea96c5659cb579b023179f0fe -de7ba64bbd9362a7b2b8c4eaec254915629e81d01c839096339b99bc9e25 -84536955feaa52fa20666f65bafd9b2e69c3e8c15d24fa407e7d881679b1 -789a0e2a695d13553c92c0214c9b7562cd6a9a3d77c8b0c2196cef76dc51 -d855c1dac37f96eae4cc7bf07e17dc7c08333d7af33c8b2965ea1f23446b -3c96c52b30ea628ad572694d145b58a606f90b278290297aa372cff56b6f -56f4aad6612eb7c7bd07db4f7d1a70d8044d16d0b5c1605ee02a852ffdb4 -450147b3f9b87d72dc431b34fcdc899462dcc1b6bb6ab1758b6a589e91e5 -8f5196251d00133b43749b7a11fb67a22664c5e38e336dbdeb5509c2d9d6 -2642c07275949df0e2db59314ae0fb34641fc171d3fe1289f919136d853c -d9048ee9db50c699c49e27a8df199590bbc65b23b55bb387eed0c73f2db5 -1cb091f8c22af83103f214199e371f7de1df23f757817200be30610004df -81fe8ed6eba79e856fca21a126ca326ad2f313c16e15754663ad6a065e08 -4050ff005fc899d6e233691b918a093b5f1ffda8839ab23ae66b1bb7b953 -0a7f896ec55de6fb9faf1b49656ff2e57488cd7f1c44114c75f9d571461f -767a6040ffa14e9fb43096f164d60ca530d7cca76d526d1999ac1b52a793 -28651112a65db1f2564ecf90ea6bf2c9ecf515640719c3fb5e36cfc58591 -e227793f39b9d3a9025cb10f324a95c29c488724aa74812366ff0b118fc7 -19f9fd0f202a040be47ec99b46b4dfc3d2a17902a5779c8d52b27231a1bb -5cd794c838daddc3e6824ca8297ba669a818c239b389400faf17aa04b802 -f763029edb9784dfdc42f223e6496a938e613463bf9bbbd59d63300a9ad7 -4e71865cac4b4e81a5864388c3886e70799c8989188341f7d17cb514cd99 -3b211883f171ec6402cc361885f4f4b110757bb3e52941a94bfaebb2faa0 -3e32eb72e25e31abdde82c2a9015478afa0f434ae3f8b97a4bef598d6eda -44ffe1915c26ee0e8339d2d45a6a080550f538ded5542c8b96ca2f596979 -8bb6223e460e857516ab5a3323136ee8fc4b0556a7c39d0cf7acb45e48be -4ae9db325e4750b73289e36a61b301795bdb2ca2a8b933be1c09fd0cd2cb -8677df171d36ef1519a2269b21e4103b2ee151c513df3e10b2a216d6fb22 -18bf2005fa7e0f0563ad96661a7f55e1b5b991f8ca285651b2683c6a7c9d -2d1941374989b06f2e9b42a6af60193dc758dd8e9fcfc7c1aa06eab47e81 -bd79660666defac0c6b9e484df9c17a61ce7a61ef73150e8cd406af6da17 -4d9c2392cc420eddda40f975ffbeacad8ce1b4e14bee29ba8552ff03376f -c034784b38dc1d0ab7bc53943d2545b03d39797af8d58d6dffce56a353d9 -bebc833f04db321ca8642bbb7fcc63ed2349ffa08a33a5d0d78f4fd2c5ea -4258e4671e362036f1f67fcef9d878ae2c203fd9c05200c59cc98633e65a -99d912ec51d6f74500d5358b70e799a6817f59adfc43365d7bba1fd6766c -5c8e76248daf3f01e7a8950fe875d657397797a45e7f99a92887300b6806 -b86db61e03c4c09d6cf507800aeead874a94e6f665746752937214302045 -0b19cfa8db69230517183a03a16e5503882ea1e419c333d3e3b73cef6762 -873ac06bec34c3f736494483442619f5bbadd86f128a5a40b854051893ea -8d31dd6656777ad4ac2572d17c6fb21385b053495d1270e65d78334a4115 -2787ea89b86f97e72718905a11e9c5664837701a3c1c65ccaf26aebe8dab -c1207d5da2079c37883d9235708f370203b3b2a8ec3a5bb35fab93dae115 -aef626dc44b67ca56fac18caf1c22e6fbab93564829a75776630b9c42513 -721ca0fbb0b402f4d1db8f701d2b29fa60162feaa8a167eb3113c6f57036 -e8361357913eb24dd38dc6d3bf4c3176a07ffc75cecf8e5940a310f79a8e -f590844383d631796ade04a91144d073a9413cff34fb454f1fd75cfbe5e6 -525c3bd36ddab80138f6c19aad7417d47df1f1e0fc958fb190a8205b5321 -7c43a4dcb0599be404473d6faebe7240dc402a0e0caa21b56a601b154524 -f44988e5074c71ae8e1948bb2a2ce72fc24cf3b1813cf7408a6b097aff22 -f9d285134d09b7053464259531eb7b270cd5f39f81bbf41a36420f61e5f6 -b429036bbf20e27af1a437becd74c5bbc25ee2519402454fc94d430636e1 -736fe65a643d9b9d21c9a54eac5a8fed51ff60a47b85a0e9423e330e00cf -220c23e056d20aec2fca3e6bc7a61a8366eb940c9bc99fb90e8704e27655 -20335a983eccc7e20b13745c4b4f30a842f1ba64745718c152697c688c73 -6cffcf5cc8eb5756201560413117a45ad3d264291cd51404f98448d31474 -d47d17d201def12867ba679f0e2605de8f3e8135ed0234890cffa68848f0 -6de427741b34c2ea654251ae8450a152538eb806ace3ecfe86d8c4a137ec -c98c6d6cbdc191a5f8f5b5972c70b4896960037b6d4c7c63586a52d5eb59 -47af8c192eb980d0801fa670bb1d08740819f9da1dd9e153010bf9580a1d -0925d8327ea1b88db8d934f40266ddf93e5ea137f267847d826cd7999b33 -c795d0ac05abe2ec1770dd98eea67912f1939118defc9b379e237d6477bc -91ad08e0046b0836fafa1272b0213dce990c90815f5b30d0eb103ac9539c -2f7bd2280264cd95b4be84cbc5139a7628ed211905dcb92cbc3180ac9e6b -b9ecc3cb08608b2395827d5729781dea49d328ba0c1b4cf2cec9f6bbc822 -1f2bbbb9d88f9e7682b9ecc06b9705faa8a90a51678183db1e24cc2c4307 -e16b3c20f08f179ec69df7a8c4261427f5886f9179c493bf2d0ef36640d7 -79925585724aba69df6d1b4f0bd2a356eedfd74a88bea667b102420c2300 -ec420e99b9ce8be1472b617e1255a7f43a0b52f11657f1a4dbb624a24886 -9604fe2062b98f5787d010723e520a4f42a0c3943e151ee627f3d5db90e0 -7747e1a88a53c4784c8d2b042b9c23c9e436d7d88343171161a364cd8961 -37a19582a00d774ef01c7c3fc9e9c7be5074c858d2bacd707a6a4f322027 -137d6ca0421ed9f9c7e7229e867678e5272cfc7156a419e893404ad7dabf -a5d8b6fd0787cb4fe1a901c34dd931f1b64f0c470ff807005fb66350d0ea -eb84ebef2c2399cd14a4454ea5004bddd99988b39c4134b92121ec77faee -55cc716eecc58b594b39c41dcab308efa4458ed71943ec5805dcd0194ddc -1ba04a5d3d42d07ac62a907ea25cd2a7e77aba470324d41dc1a3fe088388 -787b3312f472cb4f23a414fa5f7c7a0cc5d121d7642b5b3f0cf7ca2173af -3f878f374938251feb3ce5ddd2d7703fc79a130978ac516daf70ae903799 -28bea3a4296f48725d578d2e8fb0f932e398404fa8a242024bc011c0ae81 -7b92bb104712253a5d89c543a744332069e33ca08bd133211d233ef799f2 -fed6a20a9073021e505def8b79e1279dacc062cfd4dddc2e8e0a7fda5dd6 -bb5a745f99cccb7ec1df532308da3da0f236c74639c280ea649b2f7ec27d -24221470b642567f3b2e1cd0b3ffa65c5ac986b557aa9b444bf470380435 -abae9b51c6da7ff753810ca7938d8a1c47d2b41fafd236cb5998f3ef365e -1f700bb257679ba3a82e235a3e97a667a6ad94412839c96dcd49dd86ccbb -6df8ad01756b311e9fd57ccd2eb2f19f035e214804e2b77769319a5389c2 -35f3ca2a73c616c9ef0984abcba167d7d652b330c68f4f6378aba69628b4 -2d59eaa2a7e4c782f6eb96f6758d17d35650b15cb5de9bf973b3b6f67c1d -f3285be8322fc2b44359640a3ba5d6d7b96142583a00a9a0ef84fbf14046 -09ad55b2aefe8c5c8f58ed21623bf765f81dbb6cca6d2a51fb7730a14839 -392cad6b47f5e03448350ab36a37d9ff2b9dab69be5196511072b10cc91f -2e6b5160b2b1bd112e6c02d14063a9bb46977b0d4bc79b921fd942f916c9 -c5708e0d133c8309de2f6ee0b1afc996c889c36de20fbbbfd32878f477cd -7735c7c3fa59e9c46e654ea20b4381d9f6c6431082e6918d532bcd539284 -af0333a783c9e7fd4fa1e4da5ce8fea2ea4037644a24532d65fa5c1ee982 -89e4b9abaf71a35d308a9b8c337f70babc5fc8dbb0327143707ca5b675c5 -2d3cf09f7a4f667fcda03d8c82d157e661517787ce6bfb35ea772de13c66 -2bd24b74ff9ab0fbcf6635d8e06b54b5b3125d17ae13d175cb7922338ec8 -9d1159fea2110995ce48f7d2b094f06d11d59b3a64a44a83d48c78855e47 -21243e82d9858401b094a236fa0a90d61863931c30d13b9bf33a35ac0d11 -a999f2b4dfba6fc187f8c235a5217d777a5a97112e7db6a8a4b06b07d9c9 -f41820e233c8b58b9e47ac56ad1ddcc0b35dd03976bc776c6ac3692ec0ca -f8c75ea7825bc84156468ca7b269d890ec9d4a365b0b31d2f6530185d5e0 -2acc3ce14eea55ebb5667067825a8682e135d23c78863d32065ddcf1a755 -e0de6dea7220d1a28416b96db40b1e9f159aeb070c9a9515f301f162b0cf -e32c6c89287de6e2b40458e3393826189a10af8517ff5a10c41c9d05d999 -aa9305a2ee8e7fe46076bc9c5722ee0a140a144ae383e84a8abe70af5d29 -96a0a896cd499caa0ed7867e7c3aac563763216e7769d12218b584d853ec -01db93ca22d0c8d6b286b20b6b26d6ef19f2cebe7030ecaa68d069fac7a0 -09d61770b5e8f83024a99142f59d88297cb8d093992c3c6c11b043b151e8 -20df640407d8bc829bfc196bf2901e63c6f16102d03ffb7c54a7a560f5f9 -5cf8379f4a2eccdcb604bd553e6157b4381940d1b3c768dbfbf2618812f5 -7fbe744b3d8ad680dd9223d8bf2412ecbb614d05b485e3b4669d22b417f5 -02cce2d705c208b15fa83b5be77ccfc1c840f385a58ae49fbe6ab4e53912 -473630e0cfecefab95ebc632a2b10a2103bfe801ca0302542080cfb4cf4d -4c241b1a6c8d28114516e3f1bf39dc02db73e6d9a797279acfd79b02a71b -ae34860dd0e11b18954129f8dd57c039bb7063a4c92f0f6a1e25f4ae59d6 -6c1cc6b73a79d6a56f7f2a8a64d571caa8a760f4f485d770d000ddf393ba -784bb27b781c47678dd78ae9b5d5e8b57d163c42c7a55e4aae22061686bf -aebcede728ff2f65e75955585208c176d100912836b5200a79062d4f09b1 -ba9465b0e937e289160ec543a4cedbbe0cdb5ecfbb4838138ee9e1ac757d -3c5f04fb6b510b389e2f521759e403bfc8ec6bd79e2d40bdd81901c10dd7 -4620acaac9108940daf03af23f09d3c8b785db562b05e597056406557857 -e96fc8bea53c2c2ccd0ea6572abb0acacfe29e737173d665ab6dc2995f60 -807aaa4073a183aed23c26c67eb137c937999fafc63b66a021125e4ee5c1 -a745ad1fff2bd828dcef392052965ce0e9af7a2c88d730fef69da91083fd -83d9fe9f73d42a8dbdcaba85b0fa93b210dbf49cdcbf5d4b69e07375fab1 -a39038cc51f66f0b10eebe0cc61f697f7025d9755830b2d65f1ad0db91ef -ebbfb578053de329935bb28d6ed6c12f748a2f70458990f04d56c35557e3 -8bc5d2e5de7f52bcf00c3bcce091aaa8852d53ac686f8f407baf3f7c8968 -69f3b62f44a5e2291aff9d30d7b5c663658a41add74562dbb0f1062f564a -9b907846291700151de04c1a55cb945eaa2e7a709218ec56d1becce1c0b7 -dc41d5f016ae8080c3b07311590a0def35337fc3c844c0ccd04926be9fec -509b1255ef12f368d20601b1ac8c68b0a935f987a21de0f8191604e921ea -0c04b00dc188fd73499852dbcccd4119ef799472b353be7f7dcc904ddfdb -920839f3d4a13bb1796f2dc886f31217845f8d7a543aabbc720311fd0e6d -a31ad3daa06d5e7e6270a34304f35ef170a7abe733428e96b0522fddbb5d -eb35aacec147067fe066c9ef145246fa3d444d176c274b91fddb8a7bd7ff -7cc7693c25895bf931eb321dc9d79f662a17691f9bd1662fecbcecf6d1f9 -cd8ddcda56d19811f05fa48bcb492feb355b0ec7c04d6046549c56f7799c -2cd0d9dade8809de7d510702e525ad9cc82c41b4fb36218e3d72e905c507 -159076a9c0e4a008ccca17bd594c69f5eee656426f865fc1988d677b72ce -b710b29a0aa8f8337552ae30e93bf7c6e5d013555872dba4737dc5f08c0f -efd428c66fc8da675373f13f89102688977e18e14dedd7f3b676256b0263 -b66b013617d9a026794b0d6040c23c5506a98530249633a6beec46117c96 -ec036eaf6439e25b8e57754af5ebaaf9b57880ad4fc93f002fb03e9fda21 -df4acb78296b0c49a5a852c134c3b10755177a0dbd6c54ea7a2b9bdac62b -5d7f3da649df856478e4baf97899e0f891a96536c283f5c81200c51c6ab6 -77285450c7f7e96836b6da5660f6cb76782ddfc64b6fc348ebc3ba4a46f7 -19176296d8c5a31132b3fa7d935a5d777c1dc84d669d564cb4fd689a38ce -680d0b3b130caea0be43864826d0d154019fd0d865f1c389cd367cb5248e -24640eb6f66603e50581f6fb5aca6cfec1d6dbf4196da10a5e1ebb14e4ca -0251c4c8412cc1673d6e7a9666b04b090567efa0b830d2362fd384cb0303 -8a40290597bdaffe429bb89fb66b9dfcfa92f39d92a8baba7266d144ac04 -f069093ebb3fcea961ba4497d3628ad207e0c8c4fac0e5f3f2a663a8d05d -b6dc33b890ae13d84dce64b495d24cc749b121659373ca31cee09bff2e9e -e5b62e89d5faa4482a75f341dd172500a54b98fc108a69a3ea94db696513 -d4c7691e0095ed3900cd4489ab008b5460b34ae8dedf3721c60de7086605 -6c391137cf23255c565bf11403bdeecf8bf39ad5e4317a4bb37003b2e7c1 -400c3b8ed7f63719bddf07908dc2decdb0f68e8ef722851c4420303f6de1 -b5efc9b2598732fd1f2cbe45a504bd7fbfdafeade3add7274a1e875aba3c -4e0abfc6444944b79f95b5009560818f7a0599e5bab4405378fadfe084f1 -653e5a0166714047e8bd4e4cb116596d8089bae9147ec1d62cd94491af75 -a1743d58bafa11b63b447c954a8d7fe11d39d969feac8fa93c614f97807d -ac62cb7a84a974a0fa555a2e3f0ef662706efcb828ef72e2ea83b29e212d -f89ffecabcb08dbb7119203c4c5db823bf4e8b698b763fbd4d21e57940d9 -1754959d21f3f649d856ac6615eac692ebcbac555f772eb6ba3cece5ebfb -cfcc2f3d8dcad7edc697df93aef762cd47cc3ba9e2cdd10940be676efe7a -a3749170edb47b7562805e3f8bd978b18057c9110ff8d19b466ea238af32 -993e2d3021745b238021f824d887d2e01a7ff12fc6f084b35292f4864579 -406c0f61d0ac7cdf7e4770b424e2ccc22353e6c82bf8ff172973df267ded -bdaabc2a742beea02e35b9b253f98de9ca131f802deee2905ca1a6dc4608 -19a59b4a4265c723007d0215fc8ac2a91ec5f86cd6aac1e370a297103c3a -3cff58c7ae201cbaaa8a12c93e95e73974f9abcd678451b1db02ebb2e10c -c5abfa573a2ea4219fd1851765649318bb556b728d432ec05a86e9894aad -9cdca63d08642655801bb37f28b6e11b958e8e800c8d521ca4aa045fe9ab -ac02dc015d18b1901d519181ef60227170a07f3328a6d5fe4c5aedb35fc1 -3dbe86564a9b1dd4c7ec648880360cdd1742ed4ac409450f1d9681cb5e46 -5edd1de2a2c7f8ed63436f98e849504ae71bb872683ae107ad5df3ca0b47 -a5b79513e02d7c540257d465ae4521cb3449d79c931e2ce8c5b0a0a4ac88 -cef7b9e5f92bf721ad51682d6b6f6c14747f78eaac1891fe29aed4eaf177 -e3d2fc655ae889c0c30a3575a76c52e95db2f6a4d8ffee9518391954b92d -39dae4e97c4022031f8ab390b66ada6dc9ab2de4d1dddf26ac4032981a69 -08f73d34b4849ae28832cddc0dcd116a47d9262b0f93c24fbfdf8a78e6ae -ae3357f3fb89530854257a9db773a1acf5271fc4ca04a06b46dbe661ca11 -9f45e0080cd129e1a7c23a33f1c48af960761b117d9d91fa5a0ed3e47865 -b774a322f7dddfda2960b91fa7ba20c8f9eb213251299ae328b28ef54b0f -55fd54f8047c555e4045cbd70964e1c953e471408e4f25fe8ca7009bfe44 -0244b1e30dff518ea7ce5078027baba4e07ecf0ebecb497b4bd88f1ff72e -b261f6dffec0ed895e237b5608d31ef479e8c9ae9003039a5fe67252ee39 -774e1501100c0fcf154f5c5c81c70539e03118ab91f4ce247f6132d46346 -bbbb126c09d7459c1977e6e367a0c83d14edf7dea081e5f795a7c831fd1b -325b33674ec9c2b68029a0e600746329ea2e1b9bdd5cb2b140468e53c108 -8e8f2567425443f8146ec37101fa4dfccb0e032fff6cdfd76382463551b1 -ae8ca6cbff0e34a3f75ad400a9573217f8cbb00a6d59ff46e48421e97091 -cb17f53f20ebeb89609ea55ed6ba4101f2f3ceccbc7ade21202439ef91d8 -a9a783c22de7e6601b50c4342e094d0eff223494489fa92150425da1b432 -908423fb3f41e0b115ec1ba592a4f920d15610b9fb33f9912aba67912d05 -1ee00a13282c1909a3a56c4ed06f2f4d1739dc296b7492aad0446f87a416 -c6db4d42b504dec3a6756f3d0845ab2d2e151aa5fde12b31a9c3b5ae1cc9 -d97192bc048f00dead66940004281c4d5a92c20b1f77795cb4f98b8eaa7c -be16f9b9d4a34a1a53e0a0deadb4fb4b20d9e8064d3412ea8d2ebd259b8f -2f04bf4bf11a5ab7883c99943d762549c3d5866bb6ed85a0e862eafbcfc7 -03bf4b77cecc0d65bce4df33e0d65456397f231f8cbf66672457cf539817 -6aa5292fae24695009e55904a04588659a3a23fa11989b925705ab45f954 -6f862b0e176fddf75b70d9ef7389f750becbffae25d58a1252cc04a79e13 -fbb6a666fd87cec5562c3e14fd78ad05be28ff3871d6fceff5aa8965bb65 -67ec76d105a6348e915b27767f5010011e80e0e2f9c34742a4eeba369e66 -8faf086a45ac9bcdd76c758db01a78602412a4244c759ece0b963d9ea58b -0efbf4376bf115288803a54cfcf78584c8af80da2a3324096463e3898285 -57de6c6354444b12a74d5e66053f6907c48522cae9e93bccdb4632131add -52eb374213888125de71994c31dba481b70b2e4c1f10b865d58ef09fc9dd -2ca7f69bd2855895256caa5dd6bf7d4d8b341d677c56ca08fd7ba37485b1 -444af8be0dcdb233a512088936ab4d7fc8c03139df396b7408747b142782 -d9406db0dcd31368d2f23ddef61b0da3c0704e9049ccf7f904548c3ca963 -76eadf1ccf77f94c157f5b84f74b0c43466134876a90c5fdc2c53af70c3f -f5c2d13cb665fed9016454bac1a629361c8ea62f4b2399233e8587db6e75 -a9cde3530f20a68ec155d275a4aa6f63aa5cd115244643b54911c954feca -d57be2a6c40f1bac38e393969617b066f7d94e8b18dd80fccd0168d4a385 -f2f1489d1dd41b68d47e5ec66ec568333d1f584e3dca90f1367a990630d0 -14355be7dc45378aa111c319838edd441f15e125f928e044640f25ffdcc5 -c116c3f6ce0d4d3195187b22200808366eca9b508ec45e664e562186efec -a97b22835d384758849605a01973cd9ffc1657b124950c9d9fa3e18b1a20 -7156c4f96f08b87824373c2865845d17a0dda71b1d69f5331c5676d0648b -ca80a7958a2aa034d7e1e9fafead9248e6e64f9ec327c60ae4f724e1fb95 -8a71e82ac3842768b27b506b5982311557432dc3f270ae6eab23a42fef70 -dd0d407a02cbadeb7b8b74a2523cf46a5f61e52b053c2007f75ae053a96d -e00646662d027d93f950e516cddff40501c76cd0d7cf76c66b7bcd1998d2 -7a19f52635c8e27511324aabbb641dd524d11d48a946937b7fa0d89a5dbc -4b582d921811b3fd84c2a432dacb67d684a77ac08845e078e2417c7d9e08 -bd555c5265024aeb55fef4579b46f8c5e79770432c5349d5a65a47ce9338 -e1b599328bb1dff2a838f732852f3debf4bb9b828f9274d03d7cf813b123 -687c5e78a26310d87870bfcb0a76bf32aa20e46f6b2826912e562f503aed -11e427b7765cd2a68da2ec0609259ff14f57c07963d075e96f8bd2eab9a0 -dc32714dd8905f2627c6d6f33563436bda2d7fa9a976f88947b84c72f454 -bf0b66ca84470375d2ff252b4a2df52ab613d0c8ef0465ff1d809ca82025 -c2122a8f44c56ebfa25690bf6a05675ebb8634ddfd24c3734fe8cb32d6d6 -c69c72a4951cb959175770b4286d383e7a3f158450945c8a2ccf7e54fb19 -aa8d2d98a07f0c55f834f2728d89f82a598269750115a02287c4d415cdaa -14e1d9e7032684002f90603c0108dd26b40fb569bb21cc63d0da7e9e1873 -9df0a9c85bc340d2b0940860d95571dc244628c59bab449f057e409e58ca -cc3369f4baa8e53c6765a55620e78341dae06e5cdf2fa5e5ba58634b29ee -ddfee7f78672e55f18a7debbc30862f278f83f4cc123ab591371f548fbf9 -bd24b3453b9b57051c2e67edff2104f3a05a9f0cb7efd81c1b1b0a2bbe95 -21854902526e5d4fa1b3be270811b972e8726623410cec7911c07f871428 -1caaead97c503714eaadb14ae5923f020093722df1b9d9c055d7d5f95af2 -a9fbc5ab6f6c2bd655f685534d7dc5fbb5ebded6ccdcf369bd83c644dc62 -84c2810495888e9d8f464a42228cdc231d5b561c6b210bc493fc1e7bfd66 -5a6c4055a6a629f571f4f05c15cb2104b4f9d0bd1b1f0ab8252da384eeae -f5fd5c663ad7a2c29f65a48a30ed8de196f9eb8ea314c6e86989298146a5 -589f76f12664c8d008228b33144679d16ff564453b5e4e9f813191b6c99e -2680e20a410949ac30691b1428a255b6185b7e3802e8511192e73c376f3d -eb807ad2727fbb4b27538b3213da0746231b1c1b595a958466155835c537 -e0df4a0ef272d4c3f7f2ef011daed38bc58bb0fd7458e48060db98971bd4 -b24bc7bd0de92573a1c7a80a5fa2b34fbe50271dabeb83aaa4235cb7f63d -6a6b399360df8b1235e4e9ab59698930044a98d5e083b5f5a5772309b390 -9e1ff2a252734b32fee3940f0e1ba61f54dd1d3f6ff0d57c9ae75a302d14 -b9dd9034279aaca80b6bd05c74bf3d968305a5046910871223a3ef8c77d8 -25d7e6d3d2809e76064c473d1cd7c05666040b6eba647e34588f49fd70a0 -3c937933a2272c938d2fd3aa8149f215bb48f3bb45090bcb9a6ace393a44 -f1a9bda2ad09a5f566b2e8887880afa45a603a63ffe7c188e3eae926a903 -4f1803368e773f42c7391dff1b9ce8599161515c549aca46aebae7db23ec -8f09db0e0f590aab75e8eb890df354b37cd886bdc230369783a4f22ab51e -0f623738681b0d3f0099c925b93bbb56411205d63f6c05647b3e460ab354 -1bf98c59f7f6c2ea8f29d8fe08df254d8a16aab686baf6856c4fed3ec96b -0328738183dbc1eebb2a3d301b0390ed8bd128bd8e7801c89941485c3c86 -22b5f223cb07dca74f0e8643240044e8c376abbd8c82ff98c6dba9b6d244 -5b6cf4189d63c6acd6e45f07485a0fa55eff370da7e71c26469740a68627 -a3c297d2bf215121fb67815b7b9403aecca10d21e59fabcbe38f5ca66e7b -551b22e28f2d1fd7303d15a42c45bf54b40ef7fc93060ae5164e54f91c55 -20bd303a98d0667a02a900813b260c0343021ac01872fd62cb6abebc7ad3 -a4456805159839ca4a3e35db586221169ded66f852e8974e3815d4d7659f -6a9bb93585aaf264f06cb6da6a26e51683945224158ea69719b8e4e36eb1 -01333aac974db8f84b051724cf245fe7a4c86582b5dbb9a5d9318180e33b -8d92c22c44b0d18f8ca34dfa4ee9693c1a26fedece01635fc5eac1fefa81 -32458254ad46dfdfd2be12a1e7f32f3728f286f1d5d4394424a073696b65 -e3c459aee9310752231fa703faf35e11796c4eeef698f4109ca8c46ee322 -5dc2e3e04fa787188e583321f8410b68b9624ff60679d3f25c13e5ea7506 -a3ce8d0bebb99d9a959ad92d8cf909988d9250b310629903d6bfcad4581a -504b91b2c91889987f36d6fd0be1d0ee5aac00aa0cb48d78a1f7a64a777f -089573ba79452efcc31c8258fb317369feb0d7ccd48cf13da6d1ccb59a4a -48ea0b398e590c1169113fed81639e13e96aa268d99cfdb7aee977fbe85f -f784853a06642b5521ae0a7f610c9739af31ba7a5157ebbbad999e23794a -d2cf25af987dc85dfa29639957cf28e7f2b7671188045130a6e2785f8d8e -30e91f0f68c1cc9f2de902952730003e816e4f5703db7a97b4c566f80547 -42fa77be563ef681a4513b9a68b2b0956551c74545cc9883428dfa72fd5c -4eee93256b26bc86ea34f7427cb0c0cc22c0cc343f739c6c0c46d0923675 -5e04d70587426ef875f8c89ff8492ea23e4e4d763b84a6437a440e69eb70 -65ab6d8cf5f8444a844e6ef3d158b451d121daea2d0e2b423eea24254226 -7eff1b4224c4e80af2a7becac1649e4bbef09f39415e9b1e3750d7ac47a1 -068a4f5ce30840b00574eb4e683e3ec25f6e690feeb0d354568efbc354ba -813ca1400734a67693af127b0f636d58b83e91548f98e3d87da7fd7cdebf -f3ecb4b9272d1c83d4980170378d32f1d98b87c440881af9ec052510982a -0c02ba6743bdc7691a44bae5e044c25304c1a2525cf2c0694494a2e9aa34 -f36af43ab288807ffa4bd418ad51d98c75f2b2f01abfd834d3305682b6b8 -62ef69d05962aac485bb4f560583a5dbb74e967eaf6d299160753ec32249 -bb1d9851d5441cb0c624208e69dc876cd8841a66976b5d7f9c99be68363b -8112d33d971f2c4f2a1feca88ba1a794ddb725c5e2e2c248082231059aef -729bb5fee5006ab8809f63e162fc0743c047c7984a9e6333b433fa143d73 -72d4a74fe37314508e04f54dc7a1445e2d6178ec9c041d0cd4fda5cae830 -4b16feb21f3222261c293a8b058dc708405c1a97ff34eee4ca69ff4e1ee2 -a03380d52297574e3aa50c8afb826fc94a14e8caa9ba89d6e92913be9e07 -bf7ae011e6bd142d8952d9c2304735e875d1ddcf82fa9fc0c6449df2acf0 -d5f6cff6d21ef6b2d29022ed79c4226c97f163284f2311cf34d5b0524a1a -a446645b9d05554f8b49075075f0734b3d1ea31410759c174fcc7305d2c1 -d7128781043cba326251a3375784a506cf32d6a11a4876f85ffa2606fbdf -27dd16d64b2108d808e33c409dd33f6e0c6079e47e7196016f261e824fba -b0e4f91a189747053e648ad2d942ece8f582f052668b63a23a2fae4c75a5 -180db7811aac654270ec6e341126e3561429f1d41fe7ba3f1de9f8bbb8d9 -fc5cebdef869376a2e42dcaa578c0807835e58d75c39f91a83d5c1eb86a1 -b0f7aab991f65eef030f212d38d10b1913bff71717c06c78d9a1be136f21 -4be157ba11ba309326c55c23ae8512646751fb82ae200c06bd2e644bed38 -c7cee826cb587ee8ff378b7fdc00ec316bd4a9c24e2c250cb3d64f8ecbb8 -7f4d81626d7f1e4491908bf17c48c84bb1736693eb4d0fe634484cdd590f -a40ae94d44f348ba683a43004b487f047745fcdfdee2e913328a11a99530 -9bd117e0e5be4fb25d176d59dc2b1842418141190ed9ae1f33e5354cacfd -a5e4bc186119e1461bcd98517e675276ddf0296d3b3cef617dfa36b4759c -944fd721e1bf63d45cea90b5817a40d153a2f779e03487cad3c1375425ac -8cbabf7f754d16cabe45c65f1be4441908e0969d5a5111c931e724537dea -7cd3fbfec9b2f7d3efa747bf586e9218c3106c49276b89fa28f770fa0644 -fe1f3fe3adf07f59c755a5b39a2ac1d6f23c256a293bf3b31b6b9cf4c622 -b188d6e7401c038657c78bfde9ba09f508f1bbe3ed79793772cfc928c4da -519f7dbf3ff7074284437d2de8d7b7c78829642d924abacf353119e9088d -14739935a23667c432806085c3af71ffb7c5fe6b4412b9b1044c1e62ee0a -a5ce7e0322bc65a8c7d874270d84136526e52d0c7f9f93199c6bb7301216 -a19bebcef3c5633f21d012b448d367157ad928e21f8e471e46982bc46a7f -df1bf816a86dc62657c4ebf286134b327ce363ab6a66634eaa2a42e99034 -069fe1302febf06959eab8e7304da4d94a83ac1650a02c38c1c4b7e65c43 -e3a6fb0213e57ac49e58721a4f36996069caedefeb48f1a59303459d5873 -f3bedcdb9d00c1cf31130c27b60928f210e1aa5e1c8e04b86d2049f31265 -9198fa646c53afa9058eb8ceb41bda65f415c79ac92af5790b176de1d300 -f1c06b782d584f458dbd07d32c427d894f84215a8e7819e295ee98d976d5 -644f11920ff2f49cb1075c3bb42b9fe4b561362902f11a75669b7e7c4475 -b65f1ae48834cd67816eb63b58cda2f50bc22eeb0cc965569b476bedded1 -2701668f609393659b266bb0e37bb27afc90bca271366e34754383363592 -0f9a3b508aabfe8deef585b07a992460c592a150b325b1e50e4214a2f483 -e9dfc826c54b488493a96eaa37276f5a9666f0a5388fe388263d2c0cf614 -c6cd01571da4389f01fcdbd0ade1c435d64c5921b5bf7dbebd5268100a03 -1e1abb8cbd83873089a9e08cf80276c7e30d2bb40280278c29fa818eb079 -87623b1cfe13e0b01e27be0a8320b69b5afee820f4705202158b7f3059b3 -655bc28a754d088fde23d43d6a9389da8bc1cf3e8ea1a6f4328c196e655e -42184444d8c0614c7167c91a492c24c8357794c61f5e47cdaf4b38004a5c -8fceaa8151e929328bce1b8f67b22034f3f75e4d105283337c3d460e7d99 -89920c43f5e1449c74ad6ab5ea029cc6e497ea60068451c4ef2132fb87ae -049077a156c868b768df4a4c475a532e2a22d999931c64f8bcc18f51d25f -0f94fbd3e9e6c094f78da062f80c4aa2b86fa572cc469e629deb4ba0c553 -55e8422b562ed2f694d0e8e5540144e30841d7593b255edd4a61dd345d5a -00e411d2c50d64782a3ebedf945fc31c00d2fe4ca800f5aeeaf12ab399db -956362e979bd7ef0787188e43835e5389ac444d13204af6bf1875622f175 -09f32015c28729cfa3b3cca90308eefaf260e3fd9df10f3e76786b8bc0eb -a30e8cd33689aabc55e3ce387cdb89a30573495852a48009cb58a0fd34bd -da911159ccacc94698ffb94c5f45f15ecc9e82365174cefbe746f95eee44 -7a33b4d823487e203478eeb2d8c4bc7b743427778249c56e48fe17d0a501 -7b693509ddfe1f42bdef97aedcc26ceffa9357dd985cdf2c70bbfc987354 -6f0aa7df227ec42f9ca2482f58809e3f9650444568c54d3520bd0a7301ef -48bfebef1fc4332b5ca851fd786c1ece136fe9e575b69393b5aec2611903 -fae6e7a5046e2ff350becb8700f209b1131044afd32fed1bc1297b6a2f29 -6ec3b87f170e92aabacc8867360e4dbce9ea29f0c1df981f6cecc8986767 -0ccfb4c9faeaad7ca9029b8ff0129fec4a040f80ead041b3bc8af7526675 -ed9e13204e64d76440a097d77c535d34165bfe9ffcade530abcc75ae224e -890d5c110004e218bd827a02ac7340e18bf3684c43e664e0a37d5fd4fd1c -4d4489d25a99d542c16e06685652cfa3567da4eb0cb517be1482939da0cd -d0ea3519ad1e51bd9dc7b9077375a8cd3b5de9888697e853bacddbbdd1a3 -0e442e1d6f2d652046821813d0cc0e8f16c97cdd32daf239f5b2b65ef620 -46f6e9821b2e2ec539302747795fa746318514d38bdf0d0e490c00e114d5 -03e7fc9a8fb83b14337a5bb4d640b52630f5450bb3bfcf7cecfbb1ef5192 -ae401265450db197bcfa07315ff95a809bc5fb4249e3a728a817f2580ae3 -50d8d6577f79c883ab4a3119d9ab98219aed0d1e826023a66da814396058 -d95e52d9af8bdbcb0454721f27855b686d13bdb473f650c9865f3e04f08d -b10f5256a3e59bcf16b12a84bb7ef3b370647cdad5929b722a05f5b3669e -14c232bb82fcb9c1dd8155ff4515f4e83c895cafb86754e896f38e5f3beb -5d29f1bd99cb8a09c5e50f412f6d8a773b79021ab2c4831aa663c5defc4d -553616874dd5bd8b75c7a2af7d029aab5a72528fbc4b5ee3d30d523412c9 -60b432434017c4cd68b2062d28f307fc287e11663511d1a6b52143afac0d -ce0f7ba3f326fb707fb8d2c985dd60090e6664f2344e098a7a1a6448026a -2ee651e8141cd7786b6543f512e4c31d25dcaf6652b1eb52706300b771cc -0c49295067befc044ea46341927123ad4b7d094784bda7fa7b568853d0b6 -1e4cc39e1abcc9479f91a2501009ae34ef7d5ff56205cf5288503591cc55 -c48abcc78daa4804549562afc713a4c11152e6e4331619b2e474a25ffb62 -7c46112fa4259f07871f8d6882e9a7ec62d20a86a0c502815d0a8f3f5ce7 -cb4a6a74b6db8e17d54bc919b82c7c729cc05b98855b9d8a0fabd8a9bdfd -4333f395607631f57c0473be0fb290c4f40a7aa6ac49208570ffa1d0f849 -d4871ebcf9ef6f5106301cf54ff8cc9918d6de74d519fccba58bb1c21543 -f3bca9f43c211b2e5c233ff6dff2c9b56d3f656f6070d13dfd0be04653e4 -98c670770e01c07b731ca0e2eb56e608828fedaf1a31087f2d43cb4c0074 -e576769b0830577c86ad5de48ee216df02d7c4e4ec231afd8e76c608fc9d -06cc86f38cf4d839e0a0829902f56cf2f86f08b975a6bdd0642d6b4c78e2 -57cf9a4f52646a952f6a220c36c91db7f44c7f44bddf33328ea8cc01827b -5f2d79e3ee6c514a4f8597a847ef5f32c6400736e6ade28faa7bc6e9c6ba -e4bbff236fa6dd2b0ed23fc77f92649feba149f82488260b0bea2a4fe1f4 -65d96d8c51719e5e10d4c17d1b67e700aac36b1ed55c93b4b2604e72f51e -b30fbf5b64c6fcaaef764639ebd789f82ed354712c7f9fcd1df257e14c0e -8fd59a0eddab684bb1b4176d79b22ad2605bf534e4b8fac2272fbdeaf210 -0424a2c5cc65f8dd5faa13313dd926128ed466046ee94bd3eb41f3ea5505 -5a70603a2ae1981bfae8e77d850fc5a5bf1bacb3df9b7cbce68ce7979fad -a73c2900526b68236c6d37197b0c521c5b1cf5cbbc89238586eceb99818e -aa47ca94ff615233575fe83d0d50d734351e0363030a12300f7b20450946 -17bb209c346ac1d35402b617d6260fce04ce8b3231ab5c05af30b0f3ccb3 -3616d3df334c8d963279537563222dfbb705c3e14616ad01927f952e6364 -4c4b7fa44ac97616c1521facd066aa33b2296dc03682eb6a3b9dd8e5bf62 -53f10667ecb07bbd50553f1b211067f5cf098b64b84d94ba9ad8b146dc9e -8e9be06bc14cfe0945e22fd819856d6996e857c0bb5f292defeb493589f4 -515700753885d61eee1b8c19e6e94fe2302c07933f949d6bf119d207fb04 -dae7bcff7578bf33d77e29611c7cf03b2df12c242827ec4c4e5b5343ca3e -4f7f38ed337583e30dedd78a082f41d60cbad55d59dbba11af1bd296ed6f -e31d2e10d3a8b5ea698e656ff97755a47ddd862d23309e2e6ed3e3e111c0 -2c3a713d782fe301dbaff0a4225f932576622d1cbae40d20f46958298d01 -783851c894f2712bfc4736d3802e548a704878e2d139348671fb96d0ddbb -f56d9349172caef0dfed4b84d867116d91063dcdf9ec401dfe8abb269ee6 -0d646bd12e0752313e2ddc272d9f4aeb9d940987596ab623f9198765cec4 -62f7b6c540c9a70c9a872bd28ea62e056560b61ec51fc68eafe008f20760 -246e06374ae5a6bd2577217700507978811ec29985ab644e474e41e8a105 -295fa67ae05e0739e8c7fbc51104522934942f53e1e1df1ec2a66f0a74b5 -9885cf2c2fad1cab3e2b609f126ac8b7350d5408a7df9ed5c27a10ef6505 -6f0d877cd7bb902977ba93e6e8520d2d018560ec8143876ad0dcb95b173d -af72c0d413bbb5541f14faa57eedb3ac2430e36911d2f486d9ebf9cb6745 -2ccc763e1e46e7a4b8373e06082176a6c66d045e18f90b4b2ad15802f6ef -cf2130cdc627601ecc19887784b6de7fb6a193bc3d057ace29f74199acae -69526ba6f7a2c669593f9d0849f12e37201c32c88384e4548a6718cbb2ab -714ccc917d93b865ac7d7d4dbd13979843f4f5c1f8b937ef12fcdc9aff50 -f09d2625f4367ee70a98772a273d8919952102aa03297e3cbcd876da5abd -2ceb162b8fe1d9a22ff694495528c09a8819fbfb6946ab205d4b2424f6d5 -6fa1c704065cb64fb2aa0fdf291fd5e7daa38667e6d8e889be7f4c453da0 -59c492cd25fcf4a03a6995897145273a66cd6ba999138bc8e2aa7d080f9d -231497ed28a9a27b6b0d4785bfaee46fee71b26d6839f2549a14e7ab7347 -0b6cf368d2d49e74c78d93477828e4582589cb447d795181d3f13dd8ad52 -3c750df8f19b3260c17a6598b406472a7204dd26c5988911ce9884de9a1d -ce33d834becb1dc80efb07f32d3ed6c2a484c5d53746071576c3f67f25ff -1558986fe2dc2265b4fff79c07e3f4c6c0ce8319e04c14728ed722cf214f -65066148bc817753dfdcc0950bf80dc515002e1a92e7d8936e9b3aa9635a -a6d512c68aebc79a62a6bd17a411bba7684e1f06be9bc3d1aca25d50c8bd -1d75597194cf87c9ffe04ff28bea91b5b9521fd356ed9e036466137586ee -f0a8795486438d0d9707cb2854f12963929edac394c562235ca71376d938 -e4e1518668180b857d75318bc22e9f0683749047e7649f9e20b35204b6ee -60c0d47bebf53179a083f0b4cad5b3327a3faf2cf03753e3e46c05773629 -7e9bb305f603369cbb568350b2b5c6d23a35c551e0ab28b082e321ef4ed0 -e2704d35c75b4750af782160c2f2e9aab0e14e541e95b64ebedd66db2c12 -a8935a60177cab634e20a8871a3a72f4b21c3a34d9dac37176a321c2ce3e -e828d140c8445117e7fe4738000c30ffae8e2a48bd618cc8813e38fa0f86 -92ca634d1e56010987483aa0f08980d91528df3d370ac724acb238e141ab -595dcb3da7a769de170edd5763078d1084e2ebefadf8a50a816b50722617 -c9539dbd68d9062b015639708dd900aecf4f15adb36339c05a9aec7403ed -771f9f28c60e52bda3ba6902e06334036c1dfd66d35ed00e3fc0bebf55da -416093b5cf512217c47f905ccc91fad879d63dd1380519a02025ddf15d70 -eaa1bd8cb6be67608fbc5c94796bd09ba35933f64c5e72a26db1ae40ef49 -af5e972fa44660588292b67ac670bf046cb1f5a7a0d73ffd6df862744786 -4a56393b0f1b4cfcfa362c74634713093161b29c94a2526b7138aa92fdde -b37a8c1f30a6b3837d9500b340515f0412e681f5bf36e7869fa157df18e5 -c79df3e6aca924d7b7dd2e0d5b87682d7ea6913b26397ac180fb75fabc1b -8e156ed542b9d8c83079bccd141c187f90d72694de4f6d08520d11cd454b -bd3c2e6d259694fda0c8decc724bdd650163b7f6ce1181590c06de4c0dd8 -536aba318cabf54782c919e07c2ffa1034143175d05deddfcd7dce6c86a9 -ec9bf6a4437da474aac2dbce2c91aedc20043f179d5c9120f3dfb1cf6906 -c27f2ec68cd75035c283e1672ea90d953a23a1515c420b81c3270fa06573 -4d003eca1bb71a2dacdab67e44f47c266c2ea1776648b62bc110671e6eca -4546d3c72c8acd956e10452c32532ed51bf3d0518467fa829efd9c896e8e -1e5c7ff6da0b51e872e403470affc95f25e1d2b9b59ddb0472705e14fdc8 -fc2af16527188508be10d098372cd7eb7d62a85c8d8dd1d0f55ae3ccd0a6 -5dd6bf776dc187bf4de409d5db3fcc5a6d852848a251f4fb4e01dac5e9b9 -587fa8c46ce03689709008b34dfb3dc105def80a1b515abcbe06e73fdf7e -7136e40cc922fe9a9da1726747e84427f288d934747b6c587490734906b8 -a91144ac82a57957cffab561714e1ff5148a39499dfc8cc96bf5d87ced17 -825e8f80cd943d9a73945fb8bc51cf1f9cb39c605491c1bb8f1c4139974a -59471ead310d041b1ca1ecd5e9f92007cd8243cb3fb1ec5256444699a9fc -ed6cb31eaf0912c16fa480a1cb4a8f4a9cb6a4d9a9903d1e2f674286032b -489b8a23ac4719fe435a9fa2d79abdbaba740e69d5ed611421b1aefcd06a -362ddbb7b79aac41e3e90657afc0b87a6e8c57ceef70a628efe19f568634 -50f47b5c6d95870039caa3d07a54e58df064bb5f59dbe9b9a2c7c84d7e0f -32386309560a0efa2cbfa27f861b208b2df4a062ffe2c59c057296aaf5c2 -0f48ffc9ff0692f8cfbd6fc6ed1f3a14537ba40d7267e6b5f69c997a949b -26577a9a99db3f53167355c4967dabd522292ddaca3c537bcf303ce76add -eb99f6664227a94d6a698dd5a5d40008349376067d057e28e55972264502 -e035b1f5e33d7b3aeae016f9be50f2aa09aa138d15d7af3c1ccb805f2d5b -cd4e9b2b5c288b2af4a25abf0a9093749377c9e8232ba1af17962f85064a -23b0a13f11acbb471cc700f9f1b588f72cb63d3d1a95a93502ef74ed212a -c452f1a84619bbdf61a1dc79c0d9ba29c7f19b400f682cf66f7705849314 -f5c8bbf973f2c53bdb060932156bf2c9cd8d36cf6271075500b0e3e6ad49 -958af46a9dc950f4c29f1ab5dc0a85924f7ffef259f778459c80118b1eb1 -ed29208d1145b21b19d62f755de4972c57a09b3decb0a8096ab025fe6b9d -be49ae35394f0ea40d3693980f97f712b27f0e28d8a549acbf1da63518d0 -374941effacf63ac3de0523cfac0dcaeb690de5836741fe58917c7ecffc1 -95e7b560a3e763aa70fc883751bd60ea0a0f893d8e9fe75a66c67e202c24 -84f66708ae74413c0101fe0b5003be20881345d917203b582a247e6c74a8 -1d0479f317aba7b9dbbc0a92e91c51fbe8775a44c57699acc9da84ad60fb -9629929d1edabbd70b4ef9887ce4ec2469f154fada42de54240cf3302364 -7c492ba17e6936a4d85e0751df0945463368a803fb40d8ded22abe118250 -86cfff1878abe5b100bc08b991cda6fdfd579332360f0c3374842edce6ed -e43649d6702f34668a29bf387e647f96d78f33395e8d4b3521cb4fb0956d -12c924c16eee798cde68e319a358cc3524c753177d976d4e14a2e0cb72a4 -80cd87bfb842060b1266568af298bbec58a717c577be73ad808e004348f1 -6aead32a3d57457376ab57197534d6e469ed24474a83618f3ce21df515a1 -22918f4b62c642de0c8a62315ebe02bcfc529c5b8f7c127085c2d819e29a -f44be20fa077ee01a8d427bbe3d97a9d2bafd77f17835279bf135900aee5 -9bc49582b18d468bf93e47ce0bdd627775264ebe9e4172839a444f928580 -8c95895b7e23592b2dcd41ee82e966c26aa2143e3057161511796e980998 -1f2e4ef5868b3bf4576e3546e6407e35cdf14654bcefa7557d09407545a2 -38173080b4771ea52054736677a8d9749a2b22b46b24fbff93c55aa2274b -8c7ddbd751bcaf1df00ccbe1f24a80622aff192fd6db2238db941ec44ae0 -dd73f6b2f80d89bd0aa30c038583deba14913d38a7b61b54522755e251b2 -aeca62033a39ec1143b2b960f9cb87f748428bec3243b8164f07d5ff72eb -f2ef69347bb933241c2401a96ba5ffa3f9ad060c41f4e6bf7280af65293a -bbae49d723dbc4be61d7e13f7a5931a697e7f2c6582dff416341ccf5a24e -9a53686a1e13bbe0bb480c19a4e72a5e477bd29f39dce1a17f63f1e8c696 -d5f8855cefdbf7ce681c7d6ac46798ca9bbdc01f9ad78ce26011ee4b0a55 -786bb41995e509058610650d4858836fcedfe72b42e1d8ba4d607e7ddbbe -3b0222919c85de3cd428fed182f37f0d38e254378c56358e258f8e336126 -9b1f1acd7f387686e8022326a6bbc1511ed3684e2d2fc9b4e53e83e127e7 -84da13550e593bbad1c87493f27b60240852e7fa24392fbf3f478f411047 -3f00a8fdb6dcb8aae629dc7f055d85341d119f7f6951ae612ffa7df82111 -d1ca48306a57a922cf4c3106f0b5e87efba6815f6de4294c7a0394087067 -677889d22a3fd86b0796200300d2716445078027fe0c0b05c86ac80d2095 -ae874324ee6ea3553bcb92fc1522a6d1524f6fa22b71598fbce784a10b5b -61e50307ef4409ffb7b38f27800f2185140ed08fc4ab396050b068025a9d -e4bddcad201e72ed9b41c4ffd4cee743c9c2345b95c5071442defc8ba5fa -9c63c56e209df41d10d93135a8080f7cccacf67e0b0ddb3e0a31df32b83f -290b3c536e9949973cdc80aa5c8a4feee20290a95f68e59f54050192de42 -f27464ee374e4d2451ee8708933b970402c90ca3070843a449d7c3146347 -1efa666a60fd5cbf55a47e4a3c5c318fc1af944d58d32690a2c7eeef09b2 -d94721896e1e3e76e44a8efd524ed5d6f5eb9da093d277441546c6828745 -ad71b6c13f653dd631bc6fc55d0eb4648b7bd9c0eddb13222542f2b6e8d8 -b80bfab4365f4199a41ac690979285d917de79359a183e6fc254b63e6408 -6d33e3c029f472f40742a99f92999f302f79994ffd615f1a848194cb56c7 -12146850f5e400303bf5bcd4e5fdccd1fe2edf5352d525cb15d8327f45a2 -6e3ac276dc8780c65724d28dc6bf9c7c985840070c35e32859168890d599 -a884dc2a90194cc2e9cc6a20c6c0ee11b20adf3aff01db48eb8dba7b0c81 -7fc10cf5a66e8171a2823a4cd22f0e80c82011ae56dd895ae2d3ebe84ff3 -d521c31453e0909cb9b1cf0b030eb6b7059ec38038cae12d0e1cc4b5b3bf -e6c821faac9b8792441e2612aa1ee9318b71f9966d7d3a64abe349be68b1 -744de7b212f6be73a0e1eb2fa30850acc3d9562f989cb2d4fbfbcd5d3ef7 -ba55717da1cabf197b06ee4d8650e968518b6103fbe68fcd5aab70bdd21d -66f09f96208db67c1b345672486657295a39a7fd689b2c9216c6b46a29dd -1283bdba295dfa839a45b86c14f553ff903a6f7a962f035ce90c241f7cde -13bab01d8b94d89abdf5288288a5b32879f0532148c188d42233613b7a1a -7f68e98e63b44af842b924167da2ab0cab8c470a1696a92a19e190a8e84b -1d307b824506e72e68377107166c9c6b6dc0eed258e71e2c6c7d3e63d921 -39690865d3f347c95070cd9691a025825421be84bd571802c85e2c83ba53 -841223435a9ced5dead103b470a4c6ae9efcc8b53331c61d0e1e6d3246cd -aa1b0da347685121196a07e97d21b10ad34e7031d95c1bafa37b4141bf33 -a6be401129dcd64086885f4b5f1b25bce75a4cc8be60af35479509e64044 -d49c8a0c286e4158a5f346ef5fe93a6d4b0a9372233c7434a7a6f9e7ea21 -30c0b4b9f62e3a74cc5d2916ebdaa51a1ef81fceb6cf221e70002a8a3106 -bfbccc2d1809dde18e9607fcaac008fabb72e8c50244507f4013c5a268a3 -6135ead9cc25362c37aa9511589f18d812e6039490f9c599f44e88754ac1 -4f6c1841d570efde27958c7f1b2c68772584e1d12fea252e3a6ec3b051a7 -6faebbf6f5101978e24a9ca927c02065e8e49150a55c64dd30757e8a33d5 -2a788437a9181efb47414dbc22fdeda203d4122137bd045611f68314e12d -1d6a5ec270c8919562c03e3af7b0e0deceeddbdaf3eab8fb5632e44dc1e8 -d46e2396b0236a46659164e33709415e7b347f7f7b87a9224a189ddf5178 -2cf66c9d385470a51efc88696176f6d3ac3b7b95fa074c981194e22981f5 -1d925f980393b7102f1f836b12855149ef1a20d2949371ddba037b53a389 -7617c257bbdfcd74bc51c2b40f8addfe1b5f8bc45aa4d953c0d1d5f4091c -6af796af6513c820499969593bfd22f8c6dcde1d2ee2c0ceebb5bd6a1ce4 -5fa61094e932b380cee381f4485e39b4b1797f2a7d8d90bcbf89b9cb1006 -2d50fff083743bf318157caac1c0179c87c03a2857fc002979e7cc97feda -966b09ceb761d3f55cf07637256c6aa8b8e5cb6aa9739452a330afbe7082 -975ee39fad5e8106e8ee05771157e92d99003533d922ccc37add065b6236 -7613d039741f99edc77c230fe8d1baba720a185186662376b947bbe1a686 -4b42c61ebe1abd40d890751ab8945c629de3b6d2a49809dc693f9e397097 -cf1e568c258081242460af2de0ca44b7ba2734573967b3bdec0e5e64598c -cbf41e630d821491504f414d9b54a3100dd5105a141cf61bd3ec41b67368 -c8cd366c543754ee800ffee3d19c9cd0d408cc772da10e4d8134964b0a61 -232e2dfbeacd0fdee12792504bb327a2e1fc44127f8577ca51d380a760b3 -740e6be46455cbf3917b90f0dfeadaa25d5d9f66cda43ebf9f75e0191a06 -25ba29666bbe8678822a453d4e876bad4a6b0d4b6cf98feb60339c9eba2a -dce4ef7faba428422c503d0210dcf8d884ca9f5094aab9f3b1a2238b569f -444748902907cb0d9d7ca33fccdd0cd29bc68e44f7bca5092be6272bc949 -baae5af92c302bb21f91b6ea8463265680f7c16f45d8ff35392a10eab87e -296f3af4478032b5b021db8510deb617941130d45c46fb3647d94b162fe2 -2738766fb6d76a06ab6803818b27c5ff4205ba668f95b5ec5ce4ce6da545 -c13ff56f417a4e0b3b8554a1e2a985a167e168adc8c4db28a601a80ab451 -91bf32acfd8d25c39c2f17fb3bca1296d3d160f25b43b4d6b94f20ffe012 -b779339b12860dfc897b366e3d400e756f4f9f4d2c86fb9d94c11ebd1450 -eaf720056e2c39529331bdcb104d113b42c94af2c6a5035750b7ae7fdcba -b6116d74bc07a11d4357ecf73d99221dad5cba4a7136425c2a3ac0e092fd -606a4ab722195e3b7fdfb5a5e3ccbb85fc701c42bec43b54e964dff3fa04 -193043eead7681cedae9cce6919949ea60ef5630c4b9263c8f98b4bc74a1 -63ccf3d0a0bc1deff39b800ac90bd734dda7ecdc73169ad77e129887db80 -7a253f8807a422eda8a16c9ee9bb8fc0942634bfe035dac9f7e36d09844e -39477c043399db4d07b3617da9d6eee76d0fde9201da98b906050748b68d -8c944ace3c96e90a3c2b63eae27b9152cb7274fa336866d71b65a57f1bc2 -bb1f482a67f3993dcb3ff24abb0223f9a026c81b2b33127a1dad8929dec7 -5d46bdd790eb1addd771c5c3965a2f514d3a128117a44560cc10a729bade -4e6c86de7c09a39602235c803902e34f5c176b18e127d71a011dd9a3a61e -ebfaa4a4e2a5651be6f4067e5e09bb4f3514d67c2129e4d3ea9568661138 -1e45af07bd84f883c70577a986416747f3bd8d1bf86d3d7b07e8a350899d -3c2dae237bd5ece45faba7a0ba30fcda7b7eec9fbeaa5a94620686d1e403 -1cd2512e8d89451c7bd8eb432c8862023d66f3f9fcec0d47598e2df59525 -d673a5ff493d458748cd6341f161a0a3e8996ca5b496508578fe4f653924 -2ae28bf4b7397c02b726fd5f9d8b898938bb668a546be6e42865f4f030d9 -5faa289eb24f7b8e249b224a95a2245605d67417a489626df7417855b8d3 -1c0043cadd2b461d32e1b39ccf409757c37b68f84e752bde6b5bbb847bf1 -57ea3434802def983d6ce5ceb3e9fbc4911b5484e99bb94dc3f383e50672 -0e85a91ed378e352838cf02921ee0ea94be01b5a60f9b1f58fcc1b4f527e -43725de9b9dadc3ef462fa279bd7138095d4cff2a0563039f71e383430dc -f628dc9611b2e3db08fb2da1d5383dc1a3c784e1e64541fde1d9d7f42505 -de96d3d0a401099fc2879af0293b0eeb143b78cc221f670c0479bc150047 -0cacb9a282e334e428b527acdfbfc56e6aec8d4d60745c1dc000011b6248 -d9ab4a17dca7cc74e17d33c0641710b02cb1edb0addc6be214b17e9f845b -2d9c8bf03c19e131e00f91f2a393b5f2ae7c3d4ae9021c4d7891d84d5067 -377ce92836e42eacd7e540824f7ac95360ce116d41d17a50748748971c82 -27f089a22ee0d21940de854f737547b73c7517addd9bdaab425a6c2908f6 -87dd990d6cba4d84308bdd4c4435a6480ecfa1a14daabd4d8e2398178e48 -de28b84f7ce4b61d2e6e64fe043c29a941f6de7621ee6f6d8b506221df05 -db238b8fe4323cb5f259d4d3d9c94d4ae1ca37d6c34345489c0284171346 -e9830e2e3c6c167238a7ffe0989d3eac870cd44102cae139469b9d909b5a -9c34792f693ac94ecd35d2277080e30a2d24b50391b6f2a3d3b6c81f7ed1 -a7b218903e7fed7a63269e27d793a2e0b40320ebf447c71f36d40dee002d -7257f43c8add31edf2c571123e46fdb413e007cc89e99b6f98d77ab38bff -cf140f787e45ffb2c7cc4ddbb59a4e32dfc36e2875f204ac851d757c1236 -12deb31324ea4c201d27fdab46e9f3988ad2bcfb8e9cfa8c487831a9b0c6 -60b20fb66b4c77f52359ac96f3b3d189aa0571c1c53db06ddb10f08882db -0b1e93e9478d4c75626c5fbdbc6044c4d82684b310ab2af144d12bf36f1a -c0bf6249d1da9ab319453594cb19d0e93c4e047fb49229c0cce76d0cece4 -2e76fabd2425382afe707db032cf617b046a59a2fc1bb3838d98fd5c8053 -ecb918bc14762e4ca45027623988f434ff4cb08bc9bff5d7de21940e3e03 -1ee042d9c30662aa76f96213fb5a92047af60f320e4660eadd1ec19d0086 -072f2202af5f219725f81882f10d1e065a8035a9946d0ca0e48a5e7dcf61 -0283b834eda01e7d94b3453830daade2aa6c947989b290c02ade0d7b2620 -813ad177ed82813b6a985d5c0a2d42419bda763d409da085936e33c817ae -68e5467eddc30be172de855a0f7f5c527555b3f4d942401b450f08273b1e -c5b5352fdb8562a71f276284cf7c27537e628f94bcbffe8d669ea2645752 -60830f1e65e83a2204cec393f6d92d4f61f317471b4b93039d298ca2cc94 -eeada0140823a2bcd1573e732e7b4bde7368f2ecca5961ad547f554ae989 -98d87b7e5d07a85c382bcea1693a697224f41eb8b406bc6a0c3eddfe8b5c -f25b11c3e4bd91ea7d6274cd6b3ee7b8f18cc3fd502a324c645568dce9e0 -d43caa61f7306fd5488fcfc439d85f8160ebf0ac90fc541f9c74d35d7833 -09309807a639477bb038200738342e50136dc64baa7cc1b879c61f7e1b90 -e1f2bd4f6e54c4dc97b8e4adeb102979203a31fe26a7f58c609915a95abc -4acc263179423f8ab16b04272d5592fc536f29a45cbcdbe15890f119ca9f -c7a52eef41dfa5c4fed087eef8e698ba738e300bd58f2a1a10da1198c1f9 -b60e2032f8384a86aa84027df21cb87977528e3bb9bea1e3a6879c56402e -a29063afc6ac0194f4944433f9a5872cf0a2a741382d7f3c0ca7817d5d7c -4b8bf53af0f18b1eb54480519cebb61d983157e039b13025e7980eb36f54 -3451bbb84e470ffd0f98eba80c74f238729dd6278294388a2e06de68a719 -47b6d478c85f124d14aaa835620e49b7f5a4f21347302c0f0864f7ebaeec -d0831c36187cbe9c848736764a31056d2cef27c07cca00033dcddca9a2f3 -b9ebf28e67257b69cd38bc23c711b6a2f6e4dda9bf5a19da275e6a8d683c -723bfbb95a90a344a6f421f0b67ae84c74652288b0597e4c86c28f73808a -77455f2948e8df634c2d14f221626b019033f9230c9167982cca9ae6dc37 -aecbcb49fd9fc1dbf2d11bba7187888721bc42a7f47c23e07d2fc5a7a91c -0dfe255a7f9d17e69af1618502a6b90b1dd748c7eaca1e1ebe8b861b04ff -e5f628f47eb4e7e65311037d7a5713d7cc3552dc85f452ba74c4f12aecd0 -d72892c940c3325640d62fe3bbbc71361dce6d54766e1fb99dedcb2d19d2 -fa6fa21f9116e03952ebbef659816a62db51a9b5b3916ff818518774ccd6 -79d44100d7236f211f36fa80a4cbafb3db76ba1e7e7f12082b0140eed2cb -5e793e24501715c6c170ad4f856a4bf16bb10210025156e635264d3cf18b -1fc1e8cd2fcfdc2ab1a24af9087975bfcf6fb703fb36e288e58d0d2ffc98 -bb4318001d931ad6161dcdf8984e6690e0f6bb07af81bf07445f8f57b355 -6b960d24e7cd152708489e4d953ab6a155a757e002ead97585e6c5333d7e -5aaab2731f047f3490432e0ebf3d0d628eefa8c1f665b9c86aabb0706639 -5bc372e16378f0d9b439c98e7bf87be73e934995d58e4e70d3ae9a5b54c8 -87a19f2826a772c39d41805c642354d9bec75b065f148f7c1e435dabbeaf -e4a5744e3f2894a928121ab069bffa3218a106a9dbb83971353a7c7a5616 -d9da66fbb908173f9b07aadcbd4d112cc353e7b70476046ce5a92e86eaff -4eec40acc840005f51f55c9f5874216851e9cf3fa431d95d3032e779e356 -4bdce33966a3a798b170a06c4cc9f73700224c858c36bbf2d0326c337ce9 -46f69c19a84187fa50afc5b36010f9a7612e3a25e846d49bb907af9505e7 -d8c78748d7dcb501bbb3d6603e829deee3784f2f3ca583d3738d6d2ecfb8 -eaa887103606211a3c1b5cd74a3e0e96fb57da91baebaecd3669661e7b1d -579ba41928a40a7028acff6cd409e601d23ff66ff2c8acb12e535360d727 -60d2e988d801930e0e9443d60dcb9f378fa75d58d73e6a3b6e5b26407c82 -67d50ad97787f8a9b91765e41552283cb67e43e59bf71cf08b9755c8ce47 -0cf374832c72d1e9702b55bcfc8b5a4e966d5072fb2a72a2108574c58601 -03082ac8c4bba3e7eeb34d6b13181365a0fbd4e0aa25ffded22008d76f67 -d44c3e29741961dbe7cbaae1622a9d2c8bca23056d2a609581d5b5e3d697 -08d7e369b48b08fa69660e0ce3157c24f8d6e59bf2f564ce495d0fca4741 -c3a58ec9f924986399480ee547ad1853288e994940bd1d0a2d2519797bf2 -8f345e1bb9cbf6997dae764e69c64534e7f9dd98f86b5710ff8b500e1c4d -f509da50c64e213ebdf91978553a5d90908eb554f09b8fc2748c9c405903 -e7bfbf0ea7e84254fb6735f09bf865244238e5fed85336c995bc3a3b9948 -947a6eb95db4cd1b64c0fccf82d247a2202e9e7eef5a550557625a0192bc -8bcc9e461e52833f6b8729ccd957d5c4b6e07016e864fc02b792c7400ace -d0a8f43c755f87bba6e5c6e1022416e5454cb34a19865d951f7aea527760 -53658cbf306ead832244f3062c39a0a121a1157a8e47008163c5bfc88197 -be16e9a1ba26a035a16dd38cc28dffb666dd4ba7356c66b7bced9e26e905 -4ce25f6d36607d8f5dda1e21ac96a815bb2989f01130ba1aca9aade554fe -effdfef5d6b0d2a01aad92f599f6a12e121010ae6acc6f150f19e7305271 -97da761b07530ca19b84b119e5edca1fad18462143b8913d6b3f6864b713 -7a93bb9e1bc29c09d660704e8d8292c61072ebfe35c354a2342b2458a353 -31d043874380d439388e46688a53bcfe01bc190ef1a6b5dec9d40aafe822 -261b28bf3e2d76f3dc4302506ce3387b4aa2a51cd4ba1faa2ed1fd7df664 -6772fe9f83d253451eeb0448b444b8ca80cc7cb653c2d1eaa0de6f2b1c72 -47e6d24ae72e620e200aff83a557a1aa7a0ce0a9cfbbeae03c31d8cbf1d8 -20b53b688ed2ffbd83418d743ee31e3d62216ac7be6c12bc1917548cf670 -d69fd2e78d9f7786ada0ea30a6f6d9fbd1f1406337151ffa1d3d40afbe03 -728fd1aa2fa8a4f075796b9de9586b71218b4356fb52daa01d3c18cb75ae -d4d33fc809dcb6e3dcf7aee408a0cef21353d76ed480bf522fdfe86e0e0a -b7d097defcb793057f0ce98ea4989a9b6787b14029a4bf10315a2557149a -fe9c91e7d825f7518b343fb556f0177a8f6ca08fbda9913d52997511590e -b9942c9813b4cf4d4aae4919401f2fc11fef0620eb5c40532cdb22d5fad6 -919a3a710de6c40d54993b5386636499c866938e33bc703a99c73adc228d -95cac73ff4f4a275c04d0d787b62c6a184dacc4024d23f593e7721be232e -9882fb738160e52ab905f0ce2c76ae6ff2c8bbe118a1acdb3b464178cf01 -94bc6a50df1090e9221be11e49f254b06c3236a31569b947ad041d1c6b55 -bfdec3c18c791ace0fe2a59504eef64a4eec4b5c8dd38b092745e0d5ad29 -276bf02c419c546627672a5764a4904635bff86fd0781d36fbdf13485229 -71f355de2b0ad250052f50ad70f61afc870ac7a816561d3232b73360d4ab -2727b2fd045f254c782bb3f1f49d94c6d625047071b7e32da5c6d21a86de -9283fd632074430772bfbd85e0c9ccab1dec16bbc049c3e223bec1b65c8a -9e98cf58b30a74f74f1a842dc91e30c023498e280ac55edd58f4cc731d81 -e443d9b9efdf5fea63c9f357320e01b8740eedaeef2495cd02eb2f338b3e -674fb074cc497d7b1937b188da857c2c230e9a931cbc00c85a7a36fa80b4 -56588e1bbabbe4ef429a6aef9bd4eb89c5752421bd049aa13f4dcf9b51ce -2503e90bc118fac78a25d187353d6f5d496cd6130b337666f49619cea985 -dfbeb7e49c67c1e0f0f8e9ec8ba14624ed0982dcbb69415e4b3c8ddba140 -397eb1fc1ddd36c94c374f018873ba41109e45afa51f0e691157d5958c06 -26fbc0903ae25e47ee372389cf65472a3e4d9769550bdc42c0b72f9a297c -d5d3c16ec67e06036e740ab664abc9f10b9499269b73ad3678daf4474329 -c2c7252c1f0df1e3b5e8f198dfef8325cb1e7e8057897a3d7fb5bb5858e0 -cfc0c115bbd7362d8e8ee41862af6eeda681cabbb06f72ebd2ae0b0be45b -a9e1be83f1da30687a655e5d148fcc17d9f53b760810a565f6d2f4cd5da3 -5434116edef756adb4d3df544a1de593be988f2bb8d36c34deaac7d9dc15 -cba49764f1e03aa09fe21fcd7c74e3d6487ebe219569e019f10dd163046b -c1a3cb2bcbaa8558197cb2c18709a998b4efa8ab8c9a71d2ccf942c17662 -1b88dee6b424165d6ce10ac48375e760983818e0085276b1674dd41042e1 -a01a8de111c903f74834199b3230bd475d92c6226ef74eb1daaec3475a6a -fcb47644a17c7e390ee3b16bef1c1ca6c55eddc44fbefbdde525921b3047 -0d76817bd8ac724739a8e743eb09cf78e88adad527d4f115b8a32ed4898f -45bab3eb802b8168aec061e3ecdb026c056fb9efe7e2df48bd516ccb12ce -00de08ed8be4ee0c41f40f4c8f64483e0ade90a78d6d4fe9203fe0b97c60 -3b2f8882bc15a212453c691c52d00fae8a3a26934ff8acf68d4352eef75a -0b10d938e55b7333dda2db0296a69e9775bf82b1aa6d684fd9080fc1c11f -ab4369c7a95a9504063db900a6e345bf6dd99be041230b2e60cc86b8c345 -1d84a9c2cb4ab6d74d63dd43dc26eb6b384f5222796d4083dcc3e1651548 -d9469f09a33b213a33ac52a6a2e23802d8f8a75c01a607940daab0051410 -73a88130bc192f303616adb113c0051b65e12086cb319c0a5323fa7def40 -402f5f87a3b2c2cf0e92789985f6775ac2743e1ffe2d0668291059740d45 -43bae7a2897e5e658592bf5a72966097742e0702deecb0cb12499eab701d -34ba37a08346217a415e44297a181bbf3744f0a49230ad6f030e11462be9 -afc2ae14e0587bc02311b48b8e2122c28cdf14414f3680fa52dbbb63b17f -6ebe4a1204f3c5d6150cbf89a8023890383153838d4dde77d4c8b1b78823 -8918c564d3babfe58eeb154307dd1997f5ab7105426e35c279008b2677e4 -695c60f956b348799c04b734338018fc27f7de7ad9d73468fdbc5283bd14 -c066ddad9a3562f16baae15d72d7bfcb409e1c874e9db1a8cde233b282b9 -6e76e9c08d85ddfbd3cce7e64104d0b0e95291bd91f405ff82f41601ee20 -8471e613fbbee67f269e4e954c36d1d18ca9880b7cc2b08fc990978efdc5 -1d157deefedaa765c1e26ee125d4a2514a41a3b95e9151a824532d7d6486 -35ad622718fe71219a697e94c2e64f26424cbb767acdef5cda70e179cd29 -b7e318d1c6d3ad26fd5fdcbf2fc221301cc1f10f5ed86b40a1a6bcc01c90 -eafd65183e75609610637b99fea57885efe76437df02a2ffc21223d039b5 -74955d9a54ff41980eddaa8768c5ad883a0c9150877392b990d63c6805db -7b8d6ab1358cbedaedb6feadb0ee4fb8f9c1ca03a3e755a74227a8930bb7 -2ea0a00b48fc626fa14d7d48624aedc31c556f44e982f3ccbde7ee735f73 -629ab1b65bcbcf0a3586a920477e8c960219802fcb1bc3a179032b324f8d -c424899b38275886cb5bc771f26a0880767d49cc23426a40a4b6ff8fe48f -d747565fc537565f6d7fd08706accc60f5fbcb45bc785f45ee9b0812366f -ae71b23ec43f3549c8224d78baf18719f05108d5741e681457ead8abc050 -462481771a8dc6cfeb98956e163981a98c59ab44d90e9c3a946c453b5071 -db0c769f7fb5144c7ab0c9ef1a6db1addcde1d4ae1daee1b4035af256a04 -df53926c7a2dcdb94caaf12f986e20929ba4e396f3aa7c93a7abaef1294f -5f13a0dd3c3aaa8fb38da3e15daa32163b7437af683b4f5e64cb14aebbde -8c69ed2e8cdbfb213fc8129af29ca2c06c8f85a5038d688d1fa5d1b54ebe -4dea81a49ce24131f8e6702e7aa4e2cba078d5dd373f894ccb275f49c690 -1dc772e1d2f5fb3fe15dbfffac62c87110162074eb72ae4e5e446bf7e650 -a554178d0d64d3c07f330f0d99e99f2239cb1597f2e5f443854cdb0f5fab -b28fe62f22e7f3419d017980f325351bb04f8f3c3dc57fee03cc029bd29b -202308d5a800ed2d500d41ace8e54e2557bf25b627883beb8118d800eb94 -f4253f855168f7fc8a2d29c5fcb76bb90a6c4e345722b8991a854047f46e -4e97336be85470b6be2b9ba573dbc4967ddcdbfc3b6fc35b0c7f3f2f570c -55dc3fee6d80bc6f46cc7e4d86a0b86f6fa61d062e213d9e442db63fbf11 -d03165b44572096995ed342893bb672f6bb55ff8fed944667995f0f89a48 -a904c47420f32afd14129c6e2bedffce1f07ea69d550b6909bb5beb4aa08 -b0b44f35e018ba5206fdb4df0228462c1fdbb95a429e53eb27bb1b0490db -f07202c3608d0f4ce08570e3d6aa3d4581c569b57bd8c1ea0e4ed3fc5497 -e316ecec06e6be582d9170d426f6d22d8c7287b8219945c124941ca8812b -e97efd9105eb6999edc0665016633b3b48820df736125b7c76c9f3a67d93 -8a2a0a6b743fd42aebc46a0249be459f16811ac9eba7b63bad7c2e88f175 -0eff8da5faaab5659824f9d19b3225aad2ac17c52c523414d3031d08a926 -30abf474fe02a32b44d3b7d9fe0c19aec16ca6d018b71d9d395ffaea0788 -0d4501d7cdf0f7077a2d63303d09083080d67f1f714a1b271dab9fc9866e -4b0571a171eec8a4e351ba2d02438cd108a33b1106acaad0ccdb051061ea -7f40543748115f29debfb4be4b42cae8762d62114ec6f8ef68c478a8e05d -ecfa18b0368428efec9eafb2353f95e3d71e1636b9d9f94a77e692843255 -698576dce13b2b858d2d15ee47cdba3ed08d64b77ab46dd29bba6aac2106 -ab847de378cccdaf35c64e50840248915f4fc110992c493cb1b9cd0b483f -0f1abf5e9b018210b477fea28234ffbe5e0bbe01338e0842a89f1e00a0ca -7cdde0b2d7c324d5e17d8d3415ccad703507497ac95360ce660b656e5f66 -72a2f50761f3d02ccdc1d5692d7797699b8e2147cfd4817c81a432ff6a5f -39cc54927fa146cbed56a55f85f123c0a94b7553a8819b329d9dd122c502 -94e3f6314d5117db89ae7597c4691b6c542979a1ca3d26a8e23d3eb698c7 -1841651e08ec771cfb974d6613f2143872c739b62796bd0a45172530793c -28d93a65b59f79c245248d2c09428657a35b0c0e367bf7a4a4f0425b3f4b -485d9f402e164328a4b963f456829a39035c00283d2e4fcb71a42da6d42a -d46cb751287de34e6519c60bb3f1a6ba91f7bfa21dca96ee712af5681701 -18ece8a0535d9ba1dd4bd835e004a2f38c5ba43c9b30d17045e5649fbbac -188922e442182d4bdafaefb39e00106a5a7765f3d67850471e3629e526af -8691f935b57bd38465665204a214fef1006ea37dc0781073ced5fc042781 -93650393c3cadfddedcc5550ed483bb6355f54600e9758e647f9c9711f1b -e7df05d0e50a698615307c18f6d4886f50188011ba499d03831185915f3f -77c4b9ce708d78423b110776aaaf90396be0381616d1e9b0c1dcf68b6396 -82399da2a7323bf42ae5347599ef4ae9e5c135522c5ecb87e201853eb899 -db60d24acad17d6b7c2c7ea4dc221f3cb6d6caacd1ac0822ea3242ad9b4d -d15116c3874e3012fad26074a23b3cc7e25d67ef349811dbc6b87b53377f -0cf972040a037ecb91e3406a9bac68c9cab9be9a6bb28e93e3275b177cd5 -0b66935cbe8dd3d6a8365625db936b2cfc87d4d6e7322df3dbe6ccda2421 -a5e5372566f626a5e9d8bc66959e443286f8eb4bcfdeb6c49a799f1efa69 -63260d0ea2d51260baba9207fb246da927fc4c89e9c4dd5848fd4ef6f81a -cd836f5f06ff0fe135cafd7ab512af55a57727dd05a5fe1f7c3c7bbe8ea7 -e6680fcb3bbbee1cf2e2c0bba20185f00e2dc3afd42f22de472cdb3eaa5a -ddf8c6fb3682eea5548c51ddca25ca615221127b4438ea535ab3089c9ed9 -b971f35245cf831d9461a5da9d57bc4e5606d26535a7414cef6aee2a7b95 -bf2276044818ee0f3b0a16532934b8b745d8137b42ec2b28fae7d55fc02c -9ccfa4e0055f8a4be96e1e235c01b8b6ad509b832a3e90161e0a449934e7 -4be973c939b31cbc19dad4c58e9be89d242f0ce200548cdd4fa2081ab3f8 -e01f358d5db24b7a50eb2096d833378921f561f132cd7988708ee10cffb6 -2256201801c667e176b1dfaecde9756d725bef093457805e16f550e8a7de -87ecd46e5b09646b73ee74f890a36867636911e4cda2c46a40e7d57cf297 -9696046614c85b1a47ba55c60544ebd3ad7d750d003bda56dd7eed8c4702 -f8b319aaeef9d3cdc59b3e63ee93c6e1e857af273eb90909ecf36ef4c276 -895c78aa762e5376c5c542f854fba864ebce56e4b0207091139f053c2c08 -3b7ddcd0a9909b52100002bc3f8c47bcb19e7a9cb58b1ac03fee95e81195 -072d3aa7c8079632725f63425a3550a947834d29ac9a26d0774e90248e18 -996731fd9aa53ab62b40ce557d98e874b763d9d629a173f0c7babfc00ae7 -82daef5f00cf3608ebeef403dbbc19e16a1d160b889f4a10359d9eacc19d -7b5f126b31720dce7fc35ec861dfa56ea23fa18423ff4e8fe6e53fc6ba16 -b95a2b5dec00f614e4f835281ee0b4bf549e7e882689e0b445dd46fc40c9 -090e5575fa2c34b02a51ad0bccf6a7bb83ca3b929285e5e9fd054b72c47b -733a66c5abda526b18b2e49d0746e067e63b948a45eab2f4221c5b62ae21 -a5d9d7cd8aa9eeb49588891d22c56b14b55ceb6488f02b73ab3b7f6c5555 -b75452594658255e4cd58ac4815f2e1bc3888c6777f62aac2f0a57d416c3 -765c991f0f9a33d888aeb2d527b482c042ee23783a04a73ad13dfc590a52 -f3116f8296cacc7ab29b7d87e7864561a5d0a12bde2d36ee697064f41d1b -ca6ef2f801caab5295d19bf4c02b10c19f73b44635ba48a0806b967d7dfc -ce9a4850171a78532cb30020c0d66b3b1e7c75eaa7894904c181a022e8bc -9b2b8ef1202f3c7d36bcab4742d4a4761bb55b64da0d99685d319f5da8fa -132be6c0483f50e2657ae8af1e28f969440d6ed43eb00e95fd9e1cd490a4 -8646f6d008598751f7a41b43fbec7770fe591012b6b0c4ae18775ccc7db5 -de0ded2dd53e82c89648d46f0d0cc5d3ac5aa104239608d512a4353b9547 -04fe6eb7e73d718323cf9d748b8ec5da01ec9358267de12cc22b05ef0312 -e4b6ac5dbb6d06d7f2d911f20d527f504d62547aef136834b3695df8044c -383b6145e824d3931a602f081d9d656f84987a1ef121772f1f5b37a116bb -d2e77d4ccda01411545d24e15ce595db4cd62ee876b8754df0b85b44e011 -b82d76ce45795e6c2c58be8690b734a8880a074f303a70da4a1b086a6de6 -56c02cc7a4c25258eff18cb0fd868214bb46f972e26509f868d065b3cb14 -1c316898cf22293391bd7051ac3a6927aada952a8fd0658ce63357c07f34 -acbf8c99a5537da0023e901f0eb5547e1b466b7d982c8c539798b76ee2a2 -252437a81a37c3b63f625172d682eeed0b795860b2755f020ef52a138353 -003c61be2052cdd7d73b2cdcd26b127660a7b22fc51a6a2f6034f37e3e46 -c1d7f83f8b28c7c965993abba1d358362833580d9c63fa85d4cb949f97de -579fb6807b95a58b78f596db50055947dd0d0e597d9687083e9bc0266e86 -90b884b27f4094d8fb82ffdbaac4d580340a9ef8aa242be87e54b601af19 -87a48d267c04e371ae77163ebd0de3f5297b1060442ecdeac38334844e38 -0f294d4be73935fd8a38de7fba6d082c3d9156d7e88f2cfff0459377cbb6 -041f37a7e05010753b98e0b67d5827aa312129bb3c3bd883c12323756406 -d555720da8a0bb30edcfa760c01ecc2ba3b15fecccf5a10e9f358822e0ff -b64178fce2ea6a1105bfb72df0e4bc499b207ae26b8ea960de48e7ee7010 -b4e671dff795e4cdc5b43e81b1604d224f0616ae311f1208859c502c1a10 -940e7b9cd11be728bd3a0c8005ae23aea32c1b642812198a6f1aed32cb75 -97152b1340dd35ada1b81051e393d38f3740fa9523df6a83b8ca7dbceb33 -6e299b54cd998d4dfef804733c76156585e42b7284cbcc4047ba6b290efc -aa60953e98cd2b4bc2893857fa6a339f820142a52ccab0df09a2709df550 -f22e5921cbca408e7998cc1cccb8adf6d8f8b71e6685ae59d290fa33f5cd -664d73e434237424060f634262f04e9a71a977556e93b692ddc3aad26d92 -97dde71e4def64932151ad572af6e681082e9944ddbec6e7a8bdfd534233 -9ca3106ca1ccc80eab14f1655978b137fad8f399df7cbfa2d7d3d9675e0e -9afec37369a8ede2c93145ab3f42a375926946680c215fa16bf7416fc892 -bacd806cd424b9f85b47802c4336918f7486af2a03bf0d39b10169d35494 -419cb1ab7b8f407897f70c18303e91563b497d70b7181ede6aa0c3efe089 -ca6135b34dd1019b298e3677f8da61f864a67023c31eaa716c40cf3d397f -9a1209564c9ec759c37028079661d2a56374203c78b023ec61340bce5d96 -e477a4f77e5c0db7c0d1257b4bbbc6f889b17e6eaab045b8adef6f931e4d -0795583d60a6b7002cf61639c6f930671f3b8ac05a1c4e002f4bfc50d8b2 -3029fc4dce1b602cc3a5533336271bccc226559ffb127e3a562f92f89824 -552b9a70466d5a3c74ae515a222b109d490f26e8fc2d9d72bc8af6d1dcc7 -80463c7af81993bac2ce4aece9d95ab736b1dc73e32d1237bc8ec2b52513 -36dbabb4ecc7ceb5d18b02043281eb9a3bfdf19bc4853c9b1722ef1cdcf4 -fcec534923db2e2653dc48545a9850c0ac2e4594abc9f7d18a0bcf2fadfb -bf085d465a4d10528312f5d790eb9511ca01061c0d94136b99a043bcf278 -c18223b1e0f1cc062b32b79e28dec2dc59a0aaa4b5f3506923c83e6a87fa -08a1d941bb644c994491cf7f3b0e2ccf6c8a8ba89376f76dfdb592374f93 -528e78e31e0b18719346b9f1486f652638e3120687774030444674cb0778 -96385c41f6566819652d825dd58f9a4308ff79b45d7828dcbfebc406e40a -c46e866cb0e3e97d6ce7fcac19a9d0fe39bbde66c5f0cf775eb3b1e6d7e1 -1f67e7edb3d5c4facc85c916bf13322b56a0414ca27d145cb740fa2c37cd -8c142d9301f1ac3704cf6a8e93973a07fde5a331cf0cbb370c7ba555de61 -18a6cea0ecb2c0e37152390cc57e2e4fb3791ddbc383ee26b6f4006d0d68 -4880888011020f856a9de47f45440f127cf27ccaea7d40a3869d39ec7dec -ebc06382d294717644b6118354e15544fd4c6d88df9245c9a83b30e6ce09 -e2498dd1df488a019b179cb859889e6ad2838f749e3b038b280ebc8d5c3a -b03e8f15751214691edf0f86281e612d7ec0773c8a5d2b433266402df62f -fcc06879ca196aaf1fc73a5f01ac46b44d6cbe7743ae9a862c20445ae2be -1544f413d010280cc2941900bf3c42ec088cb21b44a915bb810e7666b545 -5324465c5943eedcef0c09128a995f431382e2062f5e39f4338c8eba1bca -e553cb60bb8f3e5038ac8073398c49f06dc734b18afa7921ea0d455e6e73 -db8ad9f77fb5ba6c28af6b4f18cbe46cf842c82d6c960be1520a5fd929df -ac7e00ede976fb2be0a07f659079a421fca693de89ce9b8fcb42b0176d9d -f3ddd58f921e13e216933d27b49d175b423751c451be7618eaab054d3b8c -23e8dd6fd60182d61e9b5c86b3b764a29a62f913ee7524d8cb33737d7224 -d95dc4bb8c2ad6397604a0ffecc8865adcb540e5da1cd769077838515118 -ebc9f0b988545c1881dd2e7a8fd73e11bd7ae9085fb4d45526b23a346b0f -e4281ee3d588106db5f7c386c488d8f2f4dd02d4c08e74c1034f987a44e5 -d39fd07538de57a42987ce290fb2f6557e8b5cbcaec168f5780927226415 -1e11e3667d33b36a793aa53e9e2d1102c9eb30cb3ba0ebac953e0227fe4a -3d3c0eb57e4390c3d35db0c41946e45be2830a1ae33fa25cf2c7c9cb4550 -ce9ff6c6e3d628fc7284daa6241604c90dde6339b7f7e7df3733416cdac8 -e5291357e4983d74d3582a490438a7fdb0af97001a31990b1de68e6adb48 -917daa387e647f9f13312db57310c7dedc2a2ea80800b4f4bbaa99c6b7b2 -7ac8345cb659489307e2565ebfd17774642c9ae5d3c18068dc35170c7d58 -4cf4173f1baf98137fa249c81f3347e1dadd6b1ba0f50c3b64c1eab183a0 -937b0f7278eff101e5267fa6480da7d602844416490c2c2c7eb0d44ac8f4 -75cfd611db5ec268db07c0b3608825c3e12834a2b2efaf5e2723c5199c42 -6011cf22e64e4c0d31d563f321097935ea0c6fcbf5acd3748d90079f6ab8 -687288dc55df29fe7958f566b27b73e2ea30747247f7a2b2add0602c7d64 -d23f52e7c96748e6a54ee8c4629b2aab8882169653f0ba7f05236bf14364 -244720f3259cbed73a318b29e4a9305deb65a2c9dec8a9d0f9a9f6fae541 -83e0f4b9a9a567057a1794945168dc23cec25d1c02ea9242c9fb6d8fc11e -e8874bd80a5226373ae87cea91853d0625c777ceb1f5a6f3debcf2f75a61 -460c7b4067f568ecd01f62901ade8bf8fbc5db9c6720420496f0cb48a002 -99870773c2e7b12e83987a5d0290d9bbf589ac889bf7d4334a5147187a7f -71008f216ce917ca4cfba5347078f354897fd87ac48af6a6c62711d2eb3a -5882bf3b32c0f1bfda976f850c9dcb97170e78c229a27fd5e292d161ece9 -a8c47a223cbdc28e24f79f6429c72b5752a08f917feda941582c36d9acb5 -748c86072858d053170fdbf708971a0bd5a8d8034ec769cb72ea88eb5cd7 -49f35be6ee5e9b5df6021926cae9dac3f5ec2b33680b12e95fd4ecbf28eb -a0503c10c6f2be6c7c47e9d66a0fae6038441c50e6447892f4aaf0a25ccd -952c2e8b201bb479099f16fc4903993ac18d4667c84c124685ae7648a826 -6bc1701cc600964fdcc01258a72104a0e5e9996b34c2691a66fa20f48d7c -2522333dfdabf3785f37dd9b021e8ee29fa10f76f43d5f935996cbf9d98d -92d0a84ce65613f7c4a5052f4c408bf10679fc28a4a9ff848d9e0c4976bb -dfdfb78bb934cd72434db596cb49e199f386a0bda69449ce2e11e3a4f53d -be134c6d7fe452a0927cf6a9a15b2406f8bd354adcde0ce136378baa565f -b9c51a03b1fbe1e166a1f92af26bd9f072250aaa6596a236ba2d5a200c90 -a760ca050421abc78223b2e8b2eea958ab23084fa1947574e846e48aeb12 -26cebb8b5a92089e9ea771557599e2fff44d75bcf600e76ae7289ba98cf3 -98208c5104562834f568ebd62801b988b0a9fdf132b6564566103b3d2d8e -6a099b7fbad8a13b8cd7f6729bb6651fc1019e66c4bd6ff27410bd5cdae7 -4010bd68b066bffdb4fd5e3dd9cf7e1a1353f7a4c5157e3ad508f4ca0259 -9761b7cdd6a81b3560b8765be3b0432fe4c25dcb4001b00c7fa62874f681 -ed22127dc3974605a05be8d8fcf9701f859ffce4dc598091891ab7596ac3 -4cd851ecfd2dbbaa2f99dac376f7bb40703fd0700d7499a7c24726bdc9bb -3b88c6a82e52686c1ee945d8825092bc81848a08722ac5a1d24353f95ec8 -18f3fa487d9600318091b0ae9874b42bb3cb683a2518b18cc1bd86c6e5e8 -3d37c14ef4fe0c77b03a3314995b1e7c1066b98c4375bd1fc5fadee1b024 -7ece4f95a0f59978d543910deb2e5761632c74c508269c4e4b9e315bda02 -975dc771fc30c8164b9df9172a4e571d8ca578cd2aaeaa0dd083e74cdc2e -d938b984b96d76a64b8c5fd12e63220bbac41e5bcd5ccb6b84bdbf6a02d5 -934ac50c654c0853209a6758bcdf560e53566d78987484bb6672ebe93f22 -dcba14e3acc132a2d9ae837adde04d8b16 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -0000000000000000000000000000000000000000000000000000000000000000 -cleartomark -%%BeginResource: procset Altsys_header 4 0 -userdict begin /AltsysDict 245 dict def end -AltsysDict begin -/bdf{bind def}bind def -/xdf{exch def}bdf -/defed{where{pop true}{false}ifelse}bdf -/ndf{1 index where{pop pop pop}{dup xcheck{bind}if def}ifelse}bdf -/d{setdash}bdf -/h{closepath}bdf -/H{}bdf -/J{setlinecap}bdf -/j{setlinejoin}bdf -/M{setmiterlimit}bdf -/n{newpath}bdf -/N{newpath}bdf -/q{gsave}bdf -/Q{grestore}bdf -/w{setlinewidth}bdf -/sepdef{ - dup where not - { -AltsysSepDict - } - if - 3 1 roll exch put -}bdf -/st{settransfer}bdf -/colorimage defed /_rci xdf -/_NXLevel2 defed { - _NXLevel2 not { -/colorimage where { -userdict eq { -/_rci false def -} if -} if - } if -} if -/md defed{ - md type /dicttype eq { -/colorimage where { -md eq { -/_rci false def -}if -}if -/settransfer where { -md eq { -/st systemdict /settransfer get def -}if -}if - }if -}if -/setstrokeadjust defed -{ - true setstrokeadjust - /C{curveto}bdf - /L{lineto}bdf - /m{moveto}bdf -} -{ - /dr{transform .25 sub round .25 add -exch .25 sub round .25 add exch itransform}bdf - /C{dr curveto}bdf - /L{dr lineto}bdf - /m{dr moveto}bdf - /setstrokeadjust{pop}bdf -}ifelse -/rectstroke defed /xt xdf -xt {/yt save def} if -/privrectpath { - 4 -2 roll m - dtransform round exch round exch idtransform - 2 copy 0 lt exch 0 lt xor - {dup 0 exch rlineto exch 0 rlineto neg 0 exch rlineto} - {exch dup 0 rlineto exch 0 exch rlineto neg 0 rlineto} - ifelse - closepath -}bdf -/rectclip{newpath privrectpath clip newpath}def -/rectfill{gsave newpath privrectpath fill grestore}def -/rectstroke{gsave newpath privrectpath stroke grestore}def -xt {yt restore} if -/_fonthacksave false def -/currentpacking defed -{ - /_bfh {/_fonthacksave currentpacking def false setpacking} bdf - /_efh {_fonthacksave setpacking} bdf -} -{ - /_bfh {} bdf - /_efh {} bdf -}ifelse -/packedarray{array astore readonly}ndf -/` -{ - false setoverprint - - - /-save0- save def - 5 index concat - pop - storerect left bottom width height rectclip - pop - - /dict_count countdictstack def - /op_count count 1 sub def - userdict begin - - /showpage {} def - - 0 setgray 0 setlinecap 1 setlinewidth - 0 setlinejoin 10 setmiterlimit [] 0 setdash newpath - -} bdf -/currentpacking defed{true setpacking}if -/min{2 copy gt{exch}if pop}bdf -/max{2 copy lt{exch}if pop}bdf -/xformfont { currentfont exch makefont setfont } bdf -/fhnumcolors 1 - statusdict begin -/processcolors defed -{ -pop processcolors -} -{ -/deviceinfo defed { -deviceinfo /Colors known { -pop deviceinfo /Colors get -} if -} if -} ifelse - end -def -/printerRes - gsave - matrix defaultmatrix setmatrix - 72 72 dtransform - abs exch abs - max - grestore - def -/graycalcs -[ - {Angle Frequency} - {GrayAngle GrayFrequency} - {0 Width Height matrix defaultmatrix idtransform -dup mul exch dup mul add sqrt 72 exch div} - {0 GrayWidth GrayHeight matrix defaultmatrix idtransform -dup mul exch dup mul add sqrt 72 exch div} -] def -/calcgraysteps { - forcemaxsteps - { -maxsteps - } - { -/currenthalftone defed -{currenthalftone /dicttype eq}{false}ifelse -{ -currenthalftone begin -HalftoneType 4 le -{graycalcs HalftoneType 1 sub get exec} -{ -HalftoneType 5 eq -{ -Default begin -{graycalcs HalftoneType 1 sub get exec} -end -} -{0 60} -ifelse -} -ifelse -end -} -{ -currentscreen pop exch -} -ifelse - -printerRes 300 max exch div exch -2 copy -sin mul round dup mul -3 1 roll -cos mul round dup mul -add 1 add -dup maxsteps gt {pop maxsteps} if - } - ifelse -} bdf -/nextrelease defed { - /languagelevel defed not { -/framebuffer defed { -0 40 string framebuffer 9 1 roll 8 {pop} repeat -dup 516 eq exch 520 eq or -{ -/fhnumcolors 3 def -/currentscreen {60 0 {pop pop 1}}bdf -/calcgraysteps {maxsteps} bdf -}if -}if - }if -}if -fhnumcolors 1 ne { - /calcgraysteps {maxsteps} bdf -} if -/currentpagedevice defed { - - - currentpagedevice /PreRenderingEnhance known - { -currentpagedevice /PreRenderingEnhance get -{ -/calcgraysteps -{ -forcemaxsteps -{maxsteps} -{256 maxsteps min} -ifelse -} def -} if - } if -} if -/gradfrequency 144 def -printerRes 1000 lt { - /gradfrequency 72 def -} if -/adjnumsteps { - - dup dtransform abs exch abs max - - printerRes div - - gradfrequency mul - round - 5 max - min -}bdf -/goodsep { - spots exch get 4 get dup sepname eq exch (_vc_Registration) eq or -}bdf -/BeginGradation defed -{/bb{BeginGradation}bdf} -{/bb{}bdf} -ifelse -/EndGradation defed -{/eb{EndGradation}bdf} -{/eb{}bdf} -ifelse -/bottom -0 def -/delta -0 def -/frac -0 def -/height -0 def -/left -0 def -/numsteps1 -0 def -/radius -0 def -/right -0 def -/top -0 def -/width -0 def -/xt -0 def -/yt -0 def -/df currentflat def -/tempstr 1 string def -/clipflatness currentflat def -/inverted? - 0 currenttransfer exec .5 ge def -/tc1 [0 0 0 1] def -/tc2 [0 0 0 1] def -/storerect{/top xdf /right xdf /bottom xdf /left xdf -/width right left sub def /height top bottom sub def}bdf -/concatprocs{ - systemdict /packedarray known - {dup type /packedarraytype eq 2 index type /packedarraytype eq or}{false}ifelse - { -/proc2 exch cvlit def /proc1 exch cvlit def -proc1 aload pop proc2 aload pop -proc1 length proc2 length add packedarray cvx - } - { -/proc2 exch cvlit def /proc1 exch cvlit def -/newproc proc1 length proc2 length add array def -newproc 0 proc1 putinterval newproc proc1 length proc2 putinterval -newproc cvx - }ifelse -}bdf -/i{dup 0 eq - {pop df dup} - {dup} ifelse - /clipflatness xdf setflat -}bdf -version cvr 38.0 le -{/setrgbcolor{ -currenttransfer exec 3 1 roll -currenttransfer exec 3 1 roll -currenttransfer exec 3 1 roll -setrgbcolor}bdf}if -/vms {/vmsv save def} bdf -/vmr {vmsv restore} bdf -/vmrs{vmsv restore /vmsv save def}bdf -/eomode{ - {/filler /eofill load def /clipper /eoclip load def} - {/filler /fill load def /clipper /clip load def} - ifelse -}bdf -/normtaper{}bdf -/logtaper{9 mul 1 add log}bdf -/CD{ - /NF exch def - { -exch dup -/FID ne 1 index/UniqueID ne and -{exch NF 3 1 roll put} -{pop pop} -ifelse - }forall - NF -}bdf -/MN{ - 1 index length - /Len exch def - dup length Len add - string dup - Len - 4 -1 roll - putinterval - dup - 0 - 4 -1 roll - putinterval -}bdf -/RC{4 -1 roll /ourvec xdf 256 string cvs(|______)anchorsearch - {1 index MN cvn/NewN exch def cvn - findfont dup maxlength dict CD dup/FontName NewN put dup - /Encoding ourvec put NewN exch definefont pop}{pop}ifelse}bdf -/RF{ - dup - FontDirectory exch - known - {pop 3 -1 roll pop} - {RC} - ifelse -}bdf -/FF{dup 256 string cvs(|______)exch MN cvn dup FontDirectory exch known - {exch pop findfont 3 -1 roll pop} - {pop dup findfont dup maxlength dict CD dup dup - /Encoding exch /Encoding get 256 array copy 7 -1 roll - {3 -1 roll dup 4 -2 roll put}forall put definefont} - ifelse}bdf -/RFJ{ - dup - FontDirectory exch - known - {pop 3 -1 roll pop - FontDirectory /Ryumin-Light-83pv-RKSJ-H known - {pop pop /Ryumin-Light-83pv-RKSJ-H dup}if - } - {RC} - ifelse -}bdf -/FFJ{dup 256 string cvs(|______)exch MN cvn dup FontDirectory exch known - {exch pop findfont 3 -1 roll pop} - {pop -dup FontDirectory exch known not - {FontDirectory /Ryumin-Light-83pv-RKSJ-H known -{pop /Ryumin-Light-83pv-RKSJ-H}if - }if - dup findfont dup maxlength dict CD dup dup - /Encoding exch /Encoding get 256 array copy 7 -1 roll - {3 -1 roll dup 4 -2 roll put}forall put definefont} - ifelse}bdf -/fps{ - currentflat - exch - dup 0 le{pop 1}if - { -dup setflat 3 index stopped -{1.3 mul dup 3 index gt{pop setflat pop pop stop}if} -{exit} -ifelse - }loop - pop setflat pop pop -}bdf -/fp{100 currentflat fps}bdf -/clipper{clip}bdf -/W{/clipper load 100 clipflatness dup setflat fps}bdf -userdict begin /BDFontDict 29 dict def end -BDFontDict begin -/bu{}def -/bn{}def -/setTxMode{av 70 ge{pop}if pop}def -/gm{m}def -/show{pop}def -/gr{pop}def -/fnt{pop pop pop}def -/fs{pop}def -/fz{pop}def -/lin{pop pop}def -/:M {pop pop} def -/sf {pop} def -/S {pop} def -/@b {pop pop pop pop pop pop pop pop} def -/_bdsave /save load def -/_bdrestore /restore load def -/save { dup /fontsave eq {null} {_bdsave} ifelse } def -/restore { dup null eq { pop } { _bdrestore } ifelse } def -/fontsave null def -end -/MacVec 256 array def -MacVec 0 /Helvetica findfont -/Encoding get 0 128 getinterval putinterval -MacVec 127 /DEL put MacVec 16#27 /quotesingle put MacVec 16#60 /grave put -/NUL/SOH/STX/ETX/EOT/ENQ/ACK/BEL/BS/HT/LF/VT/FF/CR/SO/SI -/DLE/DC1/DC2/DC3/DC4/NAK/SYN/ETB/CAN/EM/SUB/ESC/FS/GS/RS/US -MacVec 0 32 getinterval astore pop -/Adieresis/Aring/Ccedilla/Eacute/Ntilde/Odieresis/Udieresis/aacute -/agrave/acircumflex/adieresis/atilde/aring/ccedilla/eacute/egrave -/ecircumflex/edieresis/iacute/igrave/icircumflex/idieresis/ntilde/oacute -/ograve/ocircumflex/odieresis/otilde/uacute/ugrave/ucircumflex/udieresis -/dagger/degree/cent/sterling/section/bullet/paragraph/germandbls -/registered/copyright/trademark/acute/dieresis/notequal/AE/Oslash -/infinity/plusminus/lessequal/greaterequal/yen/mu/partialdiff/summation -/product/pi/integral/ordfeminine/ordmasculine/Omega/ae/oslash -/questiondown/exclamdown/logicalnot/radical/florin/approxequal/Delta/guillemotleft -/guillemotright/ellipsis/nbspace/Agrave/Atilde/Otilde/OE/oe -/endash/emdash/quotedblleft/quotedblright/quoteleft/quoteright/divide/lozenge -/ydieresis/Ydieresis/fraction/currency/guilsinglleft/guilsinglright/fi/fl -/daggerdbl/periodcentered/quotesinglbase/quotedblbase -/perthousand/Acircumflex/Ecircumflex/Aacute -/Edieresis/Egrave/Iacute/Icircumflex/Idieresis/Igrave/Oacute/Ocircumflex -/apple/Ograve/Uacute/Ucircumflex/Ugrave/dotlessi/circumflex/tilde -/macron/breve/dotaccent/ring/cedilla/hungarumlaut/ogonek/caron -MacVec 128 128 getinterval astore pop -end %. AltsysDict -%%EndResource -%%EndProlog -%%BeginSetup -AltsysDict begin -_bfh -%%IncludeResource: font Symbol -_efh -0 dict dup begin -end -/f0 /Symbol FF def -_bfh -%%IncludeResource: font ZapfHumanist601BT-Bold -_efh -0 dict dup begin -end -/f1 /ZapfHumanist601BT-Bold FF def -end %. AltsysDict -%%EndSetup -AltsysDict begin -/onlyk4{false}ndf -/ccmyk{dup 5 -1 roll sub 0 max exch}ndf -/cmyk2gray{ - 4 -1 roll 0.3 mul 4 -1 roll 0.59 mul 4 -1 roll 0.11 mul - add add add 1 min neg 1 add -}bdf -/setcmykcolor{1 exch sub ccmyk ccmyk ccmyk pop setrgbcolor}ndf -/maxcolor { - max max max -} ndf -/maxspot { - pop -} ndf -/setcmykcoloroverprint{4{dup -1 eq{pop 0}if 4 1 roll}repeat setcmykcolor}ndf -/findcmykcustomcolor{5 packedarray}ndf -/setcustomcolor{exch aload pop pop 4{4 index mul 4 1 roll}repeat setcmykcolor pop}ndf -/setseparationgray{setgray}ndf -/setoverprint{pop}ndf -/currentoverprint false ndf -/cmykbufs2gray{ - 0 1 2 index length 1 sub - { -4 index 1 index get 0.3 mul -4 index 2 index get 0.59 mul -4 index 3 index get 0.11 mul -4 index 4 index get -add add add cvi 255 min -255 exch sub -2 index 3 1 roll put - }for - 4 1 roll pop pop pop -}bdf -/colorimage{ - pop pop - [ -5 -1 roll/exec cvx -6 -1 roll/exec cvx -7 -1 roll/exec cvx -8 -1 roll/exec cvx -/cmykbufs2gray cvx - ]cvx - image -} -%. version 47.1 on Linotronic of Postscript defines colorimage incorrectly (rgb model only) -version cvr 47.1 le -statusdict /product get (Lino) anchorsearch{pop pop true}{pop false}ifelse -and{userdict begin bdf end}{ndf}ifelse -fhnumcolors 1 ne {/yt save def} if -/customcolorimage{ - aload pop - (_vc_Registration) eq - { -pop pop pop pop separationimage - } - { -/ik xdf /iy xdf /im xdf /ic xdf -ic im iy ik cmyk2gray /xt xdf -currenttransfer -{dup 1.0 exch sub xt mul add}concatprocs -st -image - } - ifelse -}ndf -fhnumcolors 1 ne {yt restore} if -fhnumcolors 3 ne {/yt save def} if -/customcolorimage{ - aload pop - (_vc_Registration) eq - { -pop pop pop pop separationimage - } - { -/ik xdf /iy xdf /im xdf /ic xdf -1.0 dup ic ik add min sub -1.0 dup im ik add min sub -1.0 dup iy ik add min sub -/ic xdf /iy xdf /im xdf -currentcolortransfer -4 1 roll -{dup 1.0 exch sub ic mul add}concatprocs 4 1 roll -{dup 1.0 exch sub iy mul add}concatprocs 4 1 roll -{dup 1.0 exch sub im mul add}concatprocs 4 1 roll -setcolortransfer -{/dummy xdf dummy}concatprocs{dummy}{dummy}true 3 colorimage - } - ifelse -}ndf -fhnumcolors 3 ne {yt restore} if -fhnumcolors 4 ne {/yt save def} if -/customcolorimage{ - aload pop - (_vc_Registration) eq - { -pop pop pop pop separationimage - } - { -/ik xdf /iy xdf /im xdf /ic xdf -currentcolortransfer -{1.0 exch sub ik mul ik sub 1 add}concatprocs 4 1 roll -{1.0 exch sub iy mul iy sub 1 add}concatprocs 4 1 roll -{1.0 exch sub im mul im sub 1 add}concatprocs 4 1 roll -{1.0 exch sub ic mul ic sub 1 add}concatprocs 4 1 roll -setcolortransfer -{/dummy xdf dummy}concatprocs{dummy}{dummy}{dummy} -true 4 colorimage - } - ifelse -}ndf -fhnumcolors 4 ne {yt restore} if -/separationimage{image}ndf -/newcmykcustomcolor{6 packedarray}ndf -/inkoverprint false ndf -/setinkoverprint{pop}ndf -/setspotcolor { - spots exch get - dup 4 get (_vc_Registration) eq - {pop 1 exch sub setseparationgray} - {0 5 getinterval exch setcustomcolor} - ifelse -}ndf -/currentcolortransfer{currenttransfer dup dup dup}ndf -/setcolortransfer{st pop pop pop}ndf -/fas{}ndf -/sas{}ndf -/fhsetspreadsize{pop}ndf -/filler{fill}bdf -/F{gsave {filler}fp grestore}bdf -/f{closepath F}bdf -/S{gsave {stroke}fp grestore}bdf -/s{closepath S}bdf -/bc4 [0 0 0 0] def -/_lfp4 { - /iosv inkoverprint def - /cosv currentoverprint def - /yt xdf - /xt xdf - /ang xdf - storerect - /taperfcn xdf - /k2 xdf /y2 xdf /m2 xdf /c2 xdf - /k1 xdf /y1 xdf /m1 xdf /c1 xdf - c1 c2 sub abs - m1 m2 sub abs - y1 y2 sub abs - k1 k2 sub abs - maxcolor - calcgraysteps mul abs round - height abs adjnumsteps - dup 2 lt {pop 1} if - 1 sub /numsteps1 xdf - currentflat mark - currentflat clipflatness - /delta top bottom sub numsteps1 1 add div def - /right right left sub def - /botsv top delta sub def - { -{ -W -xt yt translate -ang rotate -xt neg yt neg translate -dup setflat -/bottom botsv def -0 1 numsteps1 -{ -numsteps1 dup 0 eq {pop 0.5 } { div } ifelse -taperfcn /frac xdf -bc4 0 c2 c1 sub frac mul c1 add put -bc4 1 m2 m1 sub frac mul m1 add put -bc4 2 y2 y1 sub frac mul y1 add put -bc4 3 k2 k1 sub frac mul k1 add put -bc4 vc -1 index setflat -{ -mark {newpath left bottom right delta rectfill}stopped -{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} -{cleartomark exit}ifelse -}loop -/bottom bottom delta sub def -}for -} -gsave stopped grestore -{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} -{exit}ifelse - }loop - cleartomark setflat - iosv setinkoverprint - cosv setoverprint -}bdf -/bcs [0 0] def -/_lfs4 { - /iosv inkoverprint def - /cosv currentoverprint def - /yt xdf - /xt xdf - /ang xdf - storerect - /taperfcn xdf - /tint2 xdf - /tint1 xdf - bcs exch 1 exch put - tint1 tint2 sub abs - bcs 1 get maxspot - calcgraysteps mul abs round - height abs adjnumsteps - dup 2 lt {pop 2} if - 1 sub /numsteps1 xdf - currentflat mark - currentflat clipflatness - /delta top bottom sub numsteps1 1 add div def - /right right left sub def - /botsv top delta sub def - { -{ -W -xt yt translate -ang rotate -xt neg yt neg translate -dup setflat -/bottom botsv def -0 1 numsteps1 -{ -numsteps1 div taperfcn /frac xdf -bcs 0 -1.0 tint2 tint1 sub frac mul tint1 add sub -put bcs vc -1 index setflat -{ -mark {newpath left bottom right delta rectfill}stopped -{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} -{cleartomark exit}ifelse -}loop -/bottom bottom delta sub def -}for -} -gsave stopped grestore -{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} -{exit}ifelse - }loop - cleartomark setflat - iosv setinkoverprint - cosv setoverprint -}bdf -/_rfs4 { - /iosv inkoverprint def - /cosv currentoverprint def - /tint2 xdf - /tint1 xdf - bcs exch 1 exch put - /radius xdf - /yt xdf - /xt xdf - tint1 tint2 sub abs - bcs 1 get maxspot - calcgraysteps mul abs round - radius abs adjnumsteps - dup 2 lt {pop 2} if - 1 sub /numsteps1 xdf - radius numsteps1 div 2 div /halfstep xdf - currentflat mark - currentflat clipflatness - { -{ -dup setflat -W -0 1 numsteps1 -{ -dup /radindex xdf -numsteps1 div /frac xdf -bcs 0 -tint2 tint1 sub frac mul tint1 add -put bcs vc -1 index setflat -{ -newpath mark xt yt radius 1 frac sub mul halfstep add 0 360 -{ arc -radindex numsteps1 ne -{ -xt yt -radindex 1 add numsteps1 -div 1 exch sub -radius mul halfstep add -dup xt add yt moveto -360 0 arcn -} if -fill -}stopped -{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} -{cleartomark exit}ifelse -}loop -}for -} -gsave stopped grestore -{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} -{exit}ifelse - }loop - cleartomark setflat - iosv setinkoverprint - cosv setoverprint -}bdf -/_rfp4 { - /iosv inkoverprint def - /cosv currentoverprint def - /k2 xdf /y2 xdf /m2 xdf /c2 xdf - /k1 xdf /y1 xdf /m1 xdf /c1 xdf - /radius xdf - /yt xdf - /xt xdf - c1 c2 sub abs - m1 m2 sub abs - y1 y2 sub abs - k1 k2 sub abs - maxcolor - calcgraysteps mul abs round - radius abs adjnumsteps - dup 2 lt {pop 1} if - 1 sub /numsteps1 xdf - radius numsteps1 dup 0 eq {pop} {div} ifelse - 2 div /halfstep xdf - currentflat mark - currentflat clipflatness - { -{ -dup setflat -W -0 1 numsteps1 -{ -dup /radindex xdf -numsteps1 dup 0 eq {pop 0.5 } { div } ifelse -/frac xdf -bc4 0 c2 c1 sub frac mul c1 add put -bc4 1 m2 m1 sub frac mul m1 add put -bc4 2 y2 y1 sub frac mul y1 add put -bc4 3 k2 k1 sub frac mul k1 add put -bc4 vc -1 index setflat -{ -newpath mark xt yt radius 1 frac sub mul halfstep add 0 360 -{ arc -radindex numsteps1 ne -{ -xt yt -radindex 1 add -numsteps1 dup 0 eq {pop} {div} ifelse -1 exch sub -radius mul halfstep add -dup xt add yt moveto -360 0 arcn -} if -fill -}stopped -{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if} -{cleartomark exit}ifelse -}loop -}for -} -gsave stopped grestore -{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if} -{exit}ifelse - }loop - cleartomark setflat - iosv setinkoverprint - cosv setoverprint -}bdf -/lfp4{_lfp4}ndf -/lfs4{_lfs4}ndf -/rfs4{_rfs4}ndf -/rfp4{_rfp4}ndf -/cvc [0 0 0 1] def -/vc{ - AltsysDict /cvc 2 index put - aload length 4 eq - {setcmykcolor} - {setspotcolor} - ifelse -}bdf -/origmtx matrix currentmatrix def -/ImMatrix matrix currentmatrix def -0 setseparationgray -/imgr {1692 1570.1102 2287.2756 2412 } def -/bleed 0 def -/clpr {1692 1570.1102 2287.2756 2412 } def -/xs 1 def -/ys 1 def -/botx 0 def -/overlap 0 def -/wdist 18 def -0 2 mul fhsetspreadsize -0 0 ne {/df 0 def /clipflatness 0 def} if -/maxsteps 256 def -/forcemaxsteps false def -vms --1845 -1956 translate -/currentpacking defed{false setpacking}if -/spots[ -1 0 0 0 (Process Cyan) false newcmykcustomcolor -0 1 0 0 (Process Magenta) false newcmykcustomcolor -0 0 1 0 (Process Yellow) false newcmykcustomcolor -0 0 0 1 (Process Black) false newcmykcustomcolor -]def -/textopf false def -/curtextmtx{}def -/otw .25 def -/msf{dup/curtextmtx xdf makefont setfont}bdf -/makesetfont/msf load def -/curtextheight{.707104 .707104 curtextmtx dtransform - dup mul exch dup mul add sqrt}bdf -/ta2{ -tempstr 2 index gsave exec grestore -cwidth cheight rmoveto -4 index eq{5 index 5 index rmoveto}if -2 index 2 index rmoveto -}bdf -/ta{exch systemdict/cshow known -{{/cheight xdf/cwidth xdf tempstr 0 2 index put ta2}exch cshow} -{{tempstr 0 2 index put tempstr stringwidth/cheight xdf/cwidth xdf ta2}forall} -ifelse 6{pop}repeat}bdf -/sts{/textopf currentoverprint def vc setoverprint -/ts{awidthshow}def exec textopf setoverprint}bdf -/stol{/xt currentlinewidth def - setlinewidth vc newpath - /ts{{false charpath stroke}ta}def exec - xt setlinewidth}bdf - -/strk{/textopf currentoverprint def vc setoverprint - /ts{{false charpath stroke}ta}def exec - textopf setoverprint - }bdf -n -[] 0 d -3.863708 M -1 w -0 j -0 J -false setoverprint -0 i -false eomode -[0 0 0 1] vc -vms -%white border -- disabled -%1845.2293 2127.8588 m -%2045.9437 2127.8588 L -%2045.9437 1956.1412 L -%1845.2293 1956.1412 L -%1845.2293 2127.8588 L -%0.1417 w -%2 J -%2 M -%[0 0 0 0] vc -%s -n -1950.8 2097.2 m -1958.8 2092.5 1967.3 2089 1975.5 2084.9 C -1976.7 2083.5 1976.1 2081.5 1976.7 2079.9 C -1979.6 2081.1 1981.6 2086.8 1985.3 2084 C -1993.4 2079.3 2001.8 2075.8 2010 2071.7 C -2010.5 2071.5 2010.5 2071.1 2010.8 2070.8 C -2011.2 2064.3 2010.9 2057.5 2011 2050.8 C -2015.8 2046.9 2022.2 2046.2 2026.6 2041.7 C -2026.5 2032.5 2026.8 2022.9 2026.4 2014.1 C -2020.4 2008.3 2015 2002.4 2008.8 1997.1 C -2003.8 1996.8 2000.7 2001.2 1996.1 2002.1 C -1995.2 1996.4 1996.9 1990.5 1995.6 1984.8 C -1989.9 1979 1984.5 1973.9 1978.8 1967.8 C -1977.7 1968.6 1976 1967.6 1974.5 1968.3 C -1967.4 1972.5 1960.1 1976.1 1952.7 1979.3 C -1946.8 1976.3 1943.4 1970.7 1938.5 1966.1 C -1933.9 1966.5 1929.4 1968.8 1925.1 1970.7 C -1917.2 1978.2 1906 1977.9 1897.2 1983.4 C -1893.2 1985.6 1889.4 1988.6 1885 1990.1 C -1884.6 1990.6 1883.9 1991 1883.8 1991.6 C -1883.7 2000.4 1884 2009.9 1883.6 2018.9 C -1887.7 2024 1893.2 2028.8 1898 2033.8 C -1899.1 2035.5 1900.9 2036.8 1902.5 2037.9 C -1903.9 2037.3 1905.2 2036.6 1906.4 2035.5 C -1906.3 2039.7 1906.5 2044.6 1906.1 2048.9 C -1906.3 2049.6 1906.7 2050.2 1907.1 2050.8 C -1913.4 2056 1918.5 2062.7 1924.8 2068.1 C -1926.6 2067.9 1928 2066.9 1929.4 2066 C -1930.2 2071 1927.7 2077.1 1930.6 2081.6 C -1936.6 2086.9 1941.5 2092.9 1947.9 2097.9 C -1949 2098.1 1949.9 2097.5 1950.8 2097.2 C -[0 0 0 0.18] vc -f -0.4 w -S -n -1975.2 2084.7 m -1976.6 2083.4 1975.7 2081.1 1976 2079.4 C -1979.3 2079.5 1980.9 2086.2 1984.8 2084 C -1992.9 2078.9 2001.7 2075.6 2010 2071.2 C -2011 2064.6 2010.2 2057.3 2010.8 2050.6 C -2015.4 2046.9 2021.1 2045.9 2025.9 2042.4 C -2026.5 2033.2 2026.8 2022.9 2025.6 2013.9 C -2020.5 2008.1 2014.5 2003.1 2009.3 1997.6 C -2004.1 1996.7 2000.7 2001.6 1995.9 2002.6 C -1995.2 1996.7 1996.3 1990.2 1994.9 1984.6 C -1989.8 1978.7 1983.6 1973.7 1978.4 1968 C -1977.3 1969.3 1976 1967.6 1974.8 1968.5 C -1967.7 1972.7 1960.4 1976.3 1952.9 1979.6 C -1946.5 1976.9 1943.1 1970.5 1937.8 1966.1 C -1928.3 1968.2 1920.6 1974.8 1911.6 1978.4 C -1901.9 1979.7 1893.9 1986.6 1885 1990.6 C -1884.3 1991 1884.3 1991.7 1884 1992.3 C -1884.5 2001 1884.2 2011 1884.3 2019.9 C -1890.9 2025.3 1895.9 2031.9 1902.3 2037.4 C -1904.2 2037.9 1905.6 2034.2 1906.8 2035.7 C -1907.4 2040.9 1905.7 2046.1 1907.3 2050.8 C -1913.6 2056.2 1919.2 2062.6 1925.1 2067.9 C -1926.9 2067.8 1928 2066.3 1929.6 2065.7 C -1929.9 2070.5 1929.2 2076 1930.1 2080.8 C -1936.5 2086.1 1941.6 2092.8 1948.4 2097.6 C -1957.3 2093.3 1966.2 2088.8 1975.2 2084.7 C -[0 0 0 0] vc -f -S -n -1954.8 2093.8 m -1961.6 2090.5 1968.2 2087 1975 2084 C -1975 2082.8 1975.6 2080.9 1974.8 2080.6 C -1974.3 2075.2 1974.6 2069.6 1974.5 2064 C -1977.5 2059.7 1984.5 2060 1988.9 2056.4 C -1989.5 2055.5 1990.5 2055.3 1990.8 2054.4 C -1991.1 2045.7 1991.4 2036.1 1990.6 2027.8 C -1990.7 2026.6 1992 2027.3 1992.8 2027.1 C -1997 2032.4 2002.6 2037.8 2007.6 2042.2 C -2008.7 2042.3 2007.8 2040.6 2007.4 2040 C -2002.3 2035.6 1997.5 2030 1992.8 2025.2 C -1991.6 2024.7 1990.8 2024.9 1990.1 2025.4 C -1989.4 2024.9 1988.1 2025.2 1987.2 2024.4 C -1987.1 2025.8 1988.3 2026.5 1989.4 2026.8 C -1989.4 2026.6 1989.3 2026.2 1989.6 2026.1 C -1989.9 2026.2 1989.9 2026.6 1989.9 2026.8 C -1989.8 2026.6 1990 2026.5 1990.1 2026.4 C -1990.2 2027 1991.1 2028.3 1990.1 2028 C -1989.9 2037.9 1990.5 2044.1 1989.6 2054.2 C -1985.9 2058 1979.7 2057.4 1976 2061.2 C -1974.5 2061.6 1975.2 2059.9 1974.5 2059.5 C -1973.9 2058 1975.6 2057.8 1975 2056.6 C -1974.5 2057.1 1974.6 2055.3 1973.6 2055.9 C -1971.9 2059.3 1974.7 2062.1 1973.1 2065.5 C -1973.1 2071.2 1972.9 2077 1973.3 2082.5 C -1967.7 2085.6 1962 2088 1956.3 2090.7 C -1953.9 2092.4 1951 2093 1948.6 2094.8 C -1943.7 2089.9 1937.9 2084.3 1933 2079.6 C -1931.3 2076.1 1933.2 2071.3 1932.3 2067.2 C -1931.3 2062.9 1933.3 2060.6 1932 2057.6 C -1932.7 2056.5 1930.9 2053.3 1933.2 2051.8 C -1936.8 2050.1 1940.1 2046.9 1944 2046.8 C -1946.3 2049.7 1949.3 2051.9 1952 2054.4 C -1954.5 2054.2 1956.4 2052.3 1958.7 2051.3 C -1960.8 2050 1963.2 2049 1965.6 2048.4 C -1968.3 2050.8 1970.7 2054.3 1973.6 2055.4 C -1973 2052.2 1969.7 2050.4 1967.6 2048.2 C -1967.1 2046.7 1968.8 2046.6 1969.5 2045.8 C -1972.8 2043.3 1980.6 2043.4 1979.3 2038.4 C -1979.4 2038.6 1979.2 2038.7 1979.1 2038.8 C -1978.7 2038.6 1978.9 2038.1 1978.8 2037.6 C -1978.9 2037.9 1978.7 2038 1978.6 2038.1 C -1978.2 2032.7 1978.4 2027.1 1978.4 2021.6 C -1979.3 2021.1 1980 2020.2 1981.5 2020.1 C -1983.5 2020.5 1984 2021.8 1985.1 2023.5 C -1985.7 2024 1987.4 2023.7 1986 2022.8 C -1984.7 2021.7 1983.3 2020.8 1983.9 2018.7 C -1987.2 2015.9 1993 2015.4 1994.9 2011.5 C -1992.2 2004.9 1999.3 2005.2 2002.1 2002.4 C -2005.9 2002.7 2004.8 1997.4 2009.1 1999 C -2011 1999.3 2010 2002.9 2012.7 2002.4 C -2010.2 2000.7 2009.4 1996.1 2005.5 1998.5 C -2002.1 2000.3 1999 2002.5 1995.4 2003.8 C -1995.2 2003.6 1994.9 2003.3 1994.7 2003.1 C -1994.3 1997 1995.6 1991.1 1994.4 1985.3 C -1994.3 1986 1993.8 1985 1994 1985.6 C -1993.8 1995.4 1994.4 2001.6 1993.5 2011.7 C -1989.7 2015.5 1983.6 2014.9 1979.8 2018.7 C -1978.3 2019.1 1979.1 2017.4 1978.4 2017 C -1977.8 2015.5 1979.4 2015.3 1978.8 2014.1 C -1978.4 2014.6 1978.5 2012.8 1977.4 2013.4 C -1975.8 2016.8 1978.5 2019.6 1976.9 2023 C -1977 2028.7 1976.7 2034.5 1977.2 2040 C -1971.6 2043.1 1965.8 2045.6 1960.1 2048.2 C -1957.7 2049.9 1954.8 2050.5 1952.4 2052.3 C -1947.6 2047.4 1941.8 2041.8 1936.8 2037.2 C -1935.2 2033.6 1937.1 2028.8 1936.1 2024.7 C -1935.1 2020.4 1937.1 2018.1 1935.9 2015.1 C -1936.5 2014.1 1934.7 2010.8 1937.1 2009.3 C -1944.4 2004.8 1952 2000.9 1959.9 1997.8 C -1963.9 1997 1963.9 2001.9 1966.8 2003.3 C -1970.3 2006.9 1973.7 2009.9 1976.9 2012.9 C -1977.9 2013 1977.1 2011.4 1976.7 2010.8 C -1971.6 2006.3 1966.8 2000.7 1962 1995.9 C -1960 1995.2 1960.1 1996.6 1958.2 1995.6 C -1957 1997 1955.1 1998.8 1953.2 1998 C -1951.7 1994.5 1954.1 1993.4 1952.9 1991.1 C -1952.1 1990.5 1953.3 1990.2 1953.2 1989.6 C -1954.2 1986.8 1950.9 1981.4 1954.4 1981.2 C -1954.7 1981.6 1954.7 1981.7 1955.1 1982 C -1961.9 1979.1 1967.6 1975 1974.3 1971.6 C -1974.7 1969.8 1976.7 1969.5 1978.4 1969.7 C -1980.3 1970 1979.3 1973.6 1982 1973.1 C -1975.8 1962.2 1968 1975.8 1960.8 1976.7 C -1956.9 1977.4 1953.3 1982.4 1949.1 1978.8 C -1946 1975.8 1941.2 1971 1939.5 1969.2 C -1938.5 1968.6 1938.9 1967.4 1937.8 1966.8 C -1928.7 1969.4 1920.6 1974.5 1912.4 1979.1 C -1904 1980 1896.6 1985 1889.3 1989.4 C -1887.9 1990.4 1885.1 1990.3 1885 1992.5 C -1885.4 2000.6 1885.2 2012.9 1885.2 2019.9 C -1886.1 2022 1889.7 2019.5 1888.4 2022.8 C -1889 2023.3 1889.8 2024.4 1890.3 2024 C -1891.2 2023.5 1891.8 2028.2 1893.4 2026.6 C -1894.2 2026.3 1893.9 2027.3 1894.4 2027.6 C -1893.4 2027.6 1894.7 2028.3 1894.1 2028.5 C -1894.4 2029.6 1896 2030 1896 2029.2 C -1896.2 2029 1896.3 2029 1896.5 2029.2 C -1896.8 2029.8 1897.3 2030 1897 2030.7 C -1896.5 2030.7 1896.9 2031.5 1897.2 2031.6 C -1898.3 2034 1899.5 2030.6 1899.6 2033.3 C -1898.5 2033 1899.6 2034.4 1900.1 2034.8 C -1901.3 2035.8 1903.2 2034.6 1902.5 2036.7 C -1904.4 2036.9 1906.1 2032.2 1907.6 2035.5 C -1907.5 2040.1 1907.7 2044.9 1907.3 2049.4 C -1908 2050.2 1908.3 2051.4 1909.5 2051.6 C -1910.1 2051.1 1911.6 2051.1 1911.4 2052.3 C -1909.7 2052.8 1912.4 2054 1912.6 2054.7 C -1913.4 2055.2 1913 2053.7 1913.6 2054.4 C -1913.6 2054.5 1913.6 2055.3 1913.6 2054.7 C -1913.7 2054.4 1913.9 2054.4 1914 2054.7 C -1914 2054.9 1914.1 2055.3 1913.8 2055.4 C -1913.7 2056 1915.2 2057.6 1916 2057.6 C -1915.9 2057.3 1916.1 2057.2 1916.2 2057.1 C -1917 2056.8 1916.7 2057.7 1917.2 2058 C -1917 2058.3 1916.7 2058.3 1916.4 2058.3 C -1917.1 2059 1917.3 2060.1 1918.4 2060.4 C -1918.1 2059.2 1919.1 2060.6 1919.1 2059.5 C -1919 2060.6 1920.6 2060.1 1919.8 2061.2 C -1919.6 2061.2 1919.3 2061.2 1919.1 2061.2 C -1919.6 2061.9 1921.4 2064.2 1921.5 2062.6 C -1922.4 2062.1 1921.6 2063.9 1922.2 2064.3 C -1922.9 2067.3 1926.1 2064.3 1925.6 2067.2 C -1927.2 2066.8 1928.4 2064.6 1930.1 2065.2 C -1931.8 2067.8 1931 2071.8 1930.8 2074.8 C -1930.6 2076.4 1930.1 2078.6 1930.6 2080.4 C -1936.6 2085.4 1941.8 2091.6 1948.1 2096.9 C -1950.7 2096.7 1952.6 2094.8 1954.8 2093.8 C -[0 0.33 0.33 0.99] vc -f -S -n -1989.4 2080.6 m -1996.1 2077.3 2002.7 2073.8 2009.6 2070.8 C -2009.6 2069.6 2010.2 2067.7 2009.3 2067.4 C -2008.9 2062 2009.1 2056.4 2009.1 2050.8 C -2012.3 2046.6 2019 2046.6 2023.5 2043.2 C -2024 2042.3 2025.1 2042.1 2025.4 2041.2 C -2025.3 2032.7 2025.6 2023.1 2025.2 2014.6 C -2025 2015.3 2024.6 2014.2 2024.7 2014.8 C -2024.5 2024.7 2025.1 2030.9 2024.2 2041 C -2020.4 2044.8 2014.3 2044.2 2010.5 2048 C -2009 2048.4 2009.8 2046.7 2009.1 2046.3 C -2008.5 2044.8 2010.2 2044.6 2009.6 2043.4 C -2009.1 2043.9 2009.2 2042.1 2008.1 2042.7 C -2006.5 2046.1 2009.3 2048.9 2007.6 2052.3 C -2007.7 2058 2007.5 2063.8 2007.9 2069.3 C -2002.3 2072.4 1996.5 2074.8 1990.8 2077.5 C -1988.4 2079.2 1985.6 2079.8 1983.2 2081.6 C -1980.5 2079 1977.9 2076.5 1975.5 2074.1 C -1975.5 2075.1 1975.5 2076.2 1975.5 2077.2 C -1977.8 2079.3 1980.3 2081.6 1982.7 2083.7 C -1985.3 2083.5 1987.1 2081.6 1989.4 2080.6 C -f -S -n -1930.1 2079.9 m -1931.1 2075.6 1929.2 2071.1 1930.8 2067.2 C -1930.3 2066.3 1930.1 2064.6 1928.7 2065.5 C -1927.7 2066.4 1926.5 2067 1925.3 2067.4 C -1924.5 2066.9 1925.6 2065.7 1924.4 2066 C -1924.2 2067.2 1923.6 2065.5 1923.2 2065.7 C -1922.3 2063.6 1917.8 2062.1 1919.6 2060.4 C -1919.3 2060.5 1919.2 2060.3 1919.1 2060.2 C -1919.7 2060.9 1918.2 2061 1917.6 2060.2 C -1917 2059.6 1916.1 2058.8 1916.4 2058 C -1915.5 2058 1917.4 2057.1 1915.7 2057.8 C -1914.8 2057.1 1913.4 2056.2 1913.3 2054.9 C -1913.1 2055.4 1911.3 2054.3 1910.9 2053.2 C -1910.7 2052.9 1910.2 2052.5 1910.7 2052.3 C -1911.1 2052.5 1910.9 2052 1910.9 2051.8 C -1910.5 2051.2 1909.9 2052.6 1909.2 2051.8 C -1908.2 2051.4 1907.8 2050.2 1907.1 2049.4 C -1907.5 2044.8 1907.3 2040 1907.3 2035.2 C -1905.3 2033 1902.8 2039.3 1902.3 2035.7 C -1899.6 2036 1898.4 2032.5 1896.3 2030.7 C -1895.7 2030.1 1897.5 2030 1896.3 2029.7 C -1896.3 2030.6 1895 2029.7 1894.4 2029.2 C -1892.9 2028.1 1894.2 2027.4 1893.6 2027.1 C -1892.1 2027.9 1891.7 2025.6 1890.8 2024.9 C -1891.1 2024.6 1889.1 2024.3 1888.4 2023 C -1887.5 2022.6 1888.2 2021.9 1888.1 2021.3 C -1886.7 2022 1885.2 2020.4 1884.8 2019.2 C -1884.8 2010 1884.6 2000.2 1885 1991.8 C -1886.9 1989.6 1889.9 1989.3 1892.2 1987.5 C -1898.3 1982.7 1905.6 1980.1 1912.8 1978.6 C -1921 1974.2 1928.8 1968.9 1937.8 1966.6 C -1939.8 1968.3 1938.8 1968.3 1940.4 1970 C -1945.4 1972.5 1947.6 1981.5 1954.6 1979.3 C -1952.3 1981 1950.4 1978.4 1948.6 1977.9 C -1945.1 1973.9 1941.1 1970.6 1938 1966.6 C -1928.4 1968.5 1920.6 1974.8 1911.9 1978.8 C -1907.1 1979.2 1902.6 1981.7 1898.2 1983.6 C -1893.9 1986 1889.9 1989 1885.5 1990.8 C -1884.9 1991.2 1884.8 1991.8 1884.5 1992.3 C -1884.9 2001.3 1884.7 2011.1 1884.8 2019.6 C -1890.6 2025 1896.5 2031.2 1902.3 2036.9 C -1904.6 2037.6 1905 2033 1907.3 2035.5 C -1907.2 2040.2 1907 2044.8 1907.1 2049.6 C -1913.6 2055.3 1918.4 2061.5 1925.1 2067.4 C -1927.3 2068.2 1929.6 2062.5 1930.6 2066.9 C -1929.7 2070.7 1930.3 2076 1930.1 2080.1 C -1935.6 2085.7 1941.9 2090.7 1947.2 2096.7 C -1942.2 2091.1 1935.5 2085.2 1930.1 2079.9 C -[0.18 0.18 0 0.78] vc -f -S -n -1930.8 2061.9 m -1930.3 2057.8 1931.8 2053.4 1931.1 2050.4 C -1931.3 2050.3 1931.7 2050.5 1931.6 2050.1 C -1933 2051.1 1934.4 2049.5 1935.9 2048.7 C -1937 2046.5 1939.5 2047.1 1941.2 2045.1 C -1939.7 2042.6 1937.3 2041.2 1935.4 2039.3 C -1934 2039.7 1934.5 2038.1 1933.7 2037.6 C -1934 2033.3 1933.1 2027.9 1934.4 2024.4 C -1934.3 2023.8 1933.9 2022.8 1933 2022.8 C -1931.6 2023.1 1930.5 2024.4 1929.2 2024.9 C -1928.4 2024.5 1929.8 2023.5 1928.7 2023.5 C -1927.7 2024.1 1926.2 2022.6 1925.6 2021.6 C -1926.9 2021.6 1924.8 2020.6 1925.6 2020.4 C -1924.7 2021.7 1923.9 2019.6 1923.2 2019.2 C -1923.3 2018.3 1923.8 2018.1 1923.2 2018 C -1922.9 2017.8 1922.9 2017.5 1922.9 2017.2 C -1922.8 2018.3 1921.3 2017.3 1920.3 2018 C -1916.6 2019.7 1913 2022.1 1910 2024.7 C -1910 2032.9 1910 2041.2 1910 2049.4 C -1915.4 2055.2 1920 2058.7 1925.3 2064.8 C -1927.2 2064 1929 2061.4 1930.8 2061.9 C -[0 0 0 0] vc -f -S -n -1907.6 2030.4 m -1907.5 2027.1 1906.4 2021.7 1908.5 2019.9 C -1908.8 2020.1 1908.9 2019 1909.2 2019.6 C -1910 2019.6 1912 2019.2 1913.1 2018.2 C -1913.7 2016.5 1920.2 2015.7 1917.4 2012.7 C -1918.2 2011.2 1917 2013.8 1917.2 2012 C -1916.9 2012.3 1916 2012.4 1915.2 2012 C -1912.5 2010.5 1916.6 2008.8 1913.6 2009.6 C -1912.6 2009.2 1911.1 2009 1910.9 2007.6 C -1911 1999.2 1911.8 1989.8 1911.2 1982.2 C -1910.1 1981.1 1908.8 1982.2 1907.6 1982.2 C -1900.8 1986.5 1893.2 1988.8 1887.2 1994.2 C -1887.2 2002.4 1887.2 2010.7 1887.2 2018.9 C -1892.6 2024.7 1897.2 2028.2 1902.5 2034.3 C -1904.3 2033.3 1906.2 2032.1 1907.6 2030.4 C -f -S -n -1910.7 2025.4 m -1912.7 2022.4 1916.7 2020.8 1919.8 2018.9 C -1920.2 2018.7 1920.6 2018.6 1921 2018.4 C -1925 2020 1927.4 2028.5 1932 2024.2 C -1932.3 2025 1932.5 2023.7 1932.8 2024.4 C -1932.8 2028 1932.8 2031.5 1932.8 2035 C -1931.9 2033.9 1932.5 2036.3 1932.3 2036.9 C -1933.2 2036.4 1932.5 2038.5 1933 2038.4 C -1933.1 2040.5 1935.6 2042.2 1936.6 2043.2 C -1936.2 2042.4 1935.1 2040.8 1933.7 2040.3 C -1932.2 2034.4 1933.8 2029.8 1933 2023.2 C -1931.1 2024.9 1928.4 2026.4 1926.5 2023.5 C -1925.1 2021.6 1923 2019.8 1921.5 2018.2 C -1917.8 2018.9 1915.2 2022.5 1911.6 2023.5 C -1910.8 2023.8 1911.2 2024.7 1910.4 2025.2 C -1910.9 2031.8 1910.6 2039.1 1910.7 2045.6 C -1910.1 2048 1910.7 2045.9 1911.2 2044.8 C -1910.6 2038.5 1911.2 2031.8 1910.7 2025.4 C -[0.07 0.06 0 0.58] vc -f -S -n -1910.7 2048.9 m -1910.3 2047.4 1911.3 2046.5 1911.6 2045.3 C -1912.9 2045.3 1913.9 2047.1 1915.2 2045.8 C -1915.2 2044.9 1916.6 2043.3 1917.2 2042.9 C -1918.7 2042.9 1919.4 2044.4 1920.5 2043.2 C -1921.2 2042.2 1921.4 2040.9 1922.4 2040.3 C -1924.5 2040.3 1925.7 2040.9 1926.8 2039.6 C -1927.1 2037.9 1926.8 2038.1 1927.7 2037.6 C -1929 2037.5 1930.4 2037 1931.6 2037.2 C -1932.3 2038.2 1933.1 2038.7 1932.8 2040.3 C -1935 2041.8 1935.9 2043.8 1938.5 2044.8 C -1938.6 2045 1938.3 2045.5 1938.8 2045.3 C -1939.1 2042.9 1935.4 2044.2 1935.4 2042.2 C -1932.1 2040.8 1932.8 2037.2 1932 2034.8 C -1932.3 2034 1932.7 2035.4 1932.5 2034.8 C -1931.3 2031.8 1935.5 2020.1 1928.9 2025.9 C -1924.6 2024.7 1922.6 2014.5 1917.4 2020.4 C -1915.5 2022.8 1912 2022.6 1910.9 2025.4 C -1911.5 2031.9 1910.9 2038.8 1911.4 2045.3 C -1911.1 2046.5 1910 2047.4 1910.4 2048.9 C -1915.1 2054.4 1920.4 2058.3 1925.1 2063.8 C -1920.8 2058.6 1914.9 2054.3 1910.7 2048.9 C -[0.4 0.4 0 0] vc -f -S -n -1934.7 2031.9 m -1934.6 2030.7 1934.9 2029.5 1934.4 2028.5 C -1934 2029.5 1934.3 2031.2 1934.2 2032.6 C -1933.8 2031.7 1934.9 2031.6 1934.7 2031.9 C -[0.92 0.92 0 0.67] vc -f -S -n -vmrs -1934.7 2019.4 m -1934.1 2015.3 1935.6 2010.9 1934.9 2007.9 C -1935.1 2007.8 1935.6 2008.1 1935.4 2007.6 C -1936.8 2008.6 1938.2 2007 1939.7 2006.2 C -1940.1 2004.3 1942.7 2005 1943.6 2003.8 C -1945.1 2000.3 1954 2000.8 1950 1996.6 C -1952.1 1993.3 1948.2 1989.2 1951.2 1985.6 C -1953 1981.4 1948.4 1982.3 1947.9 1979.8 C -1945.4 1979.6 1945.1 1975.5 1942.4 1975 C -1942.4 1972.3 1938 1973.6 1938.5 1970.4 C -1937.4 1969 1935.6 1970.1 1934.2 1970.2 C -1927.5 1974.5 1919.8 1976.8 1913.8 1982.2 C -1913.8 1990.4 1913.8 1998.7 1913.8 2006.9 C -1919.3 2012.7 1923.8 2016.2 1929.2 2022.3 C -1931.1 2021.6 1932.8 2018.9 1934.7 2019.4 C -[0 0 0 0] vc -f -0.4 w -2 J -2 M -S -n -2024.2 2038.1 m -2024.1 2029.3 2024.4 2021.7 2024.7 2014.4 C -2024.4 2013.6 2020.6 2013.4 2021.3 2011.2 C -2020.5 2010.3 2018.4 2010.6 2018.9 2008.6 C -2019 2008.8 2018.8 2009 2018.7 2009.1 C -2018.2 2006.7 2015.2 2007.9 2015.3 2005.5 C -2014.7 2004.8 2012.4 2005.1 2013.2 2003.6 C -2012.3 2004.2 2012.8 2002.4 2012.7 2002.6 C -2009.4 2003.3 2011.2 1998.6 2008.4 1999.2 C -2007 1999.1 2006.1 1999.4 2005.7 2000.4 C -2006.9 1998.5 2007.7 2000.5 2009.3 2000.2 C -2009.2 2003.7 2012.4 2002.1 2012.9 2005.2 C -2015.9 2005.6 2015.2 2008.6 2017.7 2008.8 C -2018.4 2009.6 2018.3 2011.4 2019.6 2011 C -2021.1 2011.7 2021.4 2014.8 2023.7 2015.1 C -2023.7 2023.5 2023.9 2031.6 2023.5 2040.5 C -2021.8 2041.7 2020.7 2043.6 2018.4 2043.9 C -2020.8 2042.7 2025.5 2041.8 2024.2 2038.1 C -[0 0.87 0.91 0.83] vc -f -S -n -2023.5 2040 m -2023.5 2031.1 2023.5 2023.4 2023.5 2015.1 C -2020.2 2015 2021.8 2010.3 2018.4 2011 C -2018.6 2007.5 2014.7 2009.3 2014.8 2006.4 C -2011.8 2006.3 2012.2 2002.3 2009.8 2002.4 C -2009.7 2001.5 2009.2 2000.1 2008.4 2000.2 C -2008.7 2000.9 2009.7 2001.2 2009.3 2002.4 C -2008.4 2004.2 2007.5 2003.1 2007.9 2005.5 C -2007.9 2010.8 2007.7 2018.7 2008.1 2023.2 C -2009 2024.3 2007.3 2023.4 2007.9 2024 C -2007.7 2024.6 2007.3 2026.3 2008.6 2027.1 C -2009.7 2026.8 2010 2027.6 2010.5 2028 C -2010.5 2028.2 2010.5 2029.1 2010.5 2028.5 C -2011.5 2028 2010.5 2030 2011.5 2030 C -2014.2 2029.7 2012.9 2032.2 2014.8 2032.6 C -2015.1 2033.6 2015.3 2033 2016 2033.3 C -2017 2033.9 2016.6 2035.4 2017.2 2036.2 C -2018.7 2036.4 2019.2 2039 2021.3 2038.4 C -2021.6 2035.4 2019.7 2029.5 2021.1 2027.3 C -2020.9 2023.5 2021.5 2018.5 2020.6 2016 C -2020.9 2013.9 2021.5 2015.4 2022.3 2014.4 C -2022.2 2015.1 2023.3 2014.8 2023.2 2015.6 C -2022.7 2019.8 2023.3 2024.3 2022.8 2028.5 C -2022.3 2028.2 2022.6 2027.6 2022.5 2027.1 C -2022.5 2027.8 2022.5 2029.2 2022.5 2029.2 C -2022.6 2029.2 2022.7 2029.1 2022.8 2029 C -2023.9 2032.8 2022.6 2037 2023 2040.8 C -2022.3 2041.2 2021.6 2041.5 2021.1 2042.2 C -2022 2041.2 2022.9 2041.4 2023.5 2040 C -[0 1 1 0.23] vc -f -S -n -2009.1 1997.8 m -2003.8 1997.7 2000.1 2002.4 1995.4 2003.1 C -1995 1999.5 1995.2 1995 1995.2 1992 C -1995.2 1995.8 1995 1999.7 1995.4 2003.3 C -2000.3 2002.2 2003.8 1997.9 2009.1 1997.8 C -2012.3 2001.2 2015.6 2004.8 2018.7 2008.1 C -2021.6 2011.2 2027.5 2013.9 2025.9 2019.9 C -2026.1 2017.9 2025.6 2016.2 2025.4 2014.4 C -2020.2 2008.4 2014 2003.6 2009.1 1997.8 C -[0.18 0.18 0 0.78] vc -f -S -n -2009.3 1997.8 m -2008.7 1997.4 2007.9 1997.6 2007.2 1997.6 C -2007.9 1997.6 2008.9 1997.4 2009.6 1997.8 C -2014.7 2003.6 2020.8 2008.8 2025.9 2014.8 C -2025.8 2017.7 2026.1 2014.8 2025.6 2014.1 C -2020.4 2008.8 2014.8 2003.3 2009.3 1997.8 C -[0.07 0.06 0 0.58] vc -f -S -n -2009.6 1997.6 m -2009 1997.1 2008.1 1997.4 2007.4 1997.3 C -2008.1 1997.4 2009 1997.1 2009.6 1997.6 C -2014.8 2003.7 2021.1 2008.3 2025.9 2014.4 C -2021.1 2008.3 2014.7 2003.5 2009.6 1997.6 C -[0.4 0.4 0 0] vc -f -S -n -2021.8 2011.5 m -2021.9 2012.2 2022.3 2013.5 2023.7 2013.6 C -2023.4 2012.7 2022.8 2011.8 2021.8 2011.5 C -[0 0.33 0.33 0.99] vc -f -S -n -2021.1 2042 m -2022.1 2041.1 2020.9 2040.2 2020.6 2039.6 C -2018.4 2039.5 2018.1 2036.9 2016.3 2036.4 C -2015.8 2035.5 2015.3 2033.8 2014.8 2033.6 C -2012.4 2033.8 2013 2030.4 2010.5 2030.2 C -2009.6 2028.9 2009.6 2028.3 2008.4 2028 C -2006.9 2026.7 2007.5 2024.3 2006 2023.2 C -2006.6 2023.2 2005.7 2023.3 2005.7 2023 C -2006.4 2022.5 2006.3 2021.1 2006.7 2020.6 C -2006.6 2015 2006.9 2009 2006.4 2003.8 C -2006.9 2002.5 2007.6 2001.1 2006.9 2000.7 C -2004.6 2003.6 2003 2002.9 2000.2 2004.3 C -1999.3 2005.8 1997.9 2006.3 1996.1 2006.7 C -1995.7 2008.9 1996 2011.1 1995.9 2012.9 C -1993.4 2015.1 1990.5 2016.2 1987.7 2017.7 C -1987.1 2019.3 1991.1 2019.4 1990.4 2021.3 C -1990.5 2021.5 1991.9 2022.3 1992 2023 C -1994.8 2024.4 1996.2 2027.5 1998.5 2030 C -2002.4 2033 2005.2 2037.2 2008.8 2041 C -2010.2 2041.3 2011.6 2042 2011 2043.9 C -2011.2 2044.8 2010.1 2045.3 2010.5 2046.3 C -2013.8 2044.8 2017.5 2043.4 2021.1 2042 C -[0 0.5 0.5 0.2] vc -f -S -n -2019.4 2008.8 m -2018.9 2009.2 2019.3 2009.9 2019.6 2010.3 C -2022.2 2011.5 2020.3 2009.1 2019.4 2008.8 C -[0 0.33 0.33 0.99] vc -f -S -n -2018 2007.4 m -2015.7 2006.7 2015.3 2003.6 2012.9 2002.8 C -2013.5 2003.7 2013.5 2005.1 2015.6 2005.2 C -2016.4 2006.1 2015.7 2007.7 2018 2007.4 C -f -S -n -vmrs -1993.5 2008.8 m -1993.4 2000 1993.7 1992.5 1994 1985.1 C -1993.7 1984.3 1989.9 1984.1 1990.6 1982 C -1989.8 1981.1 1987.7 1981.4 1988.2 1979.3 C -1988.3 1979.6 1988.1 1979.7 1988 1979.8 C -1987.5 1977.5 1984.5 1978.6 1984.6 1976.2 C -1983.9 1975.5 1981.7 1975.8 1982.4 1974.3 C -1981.6 1974.9 1982.1 1973.1 1982 1973.3 C -1979 1973.7 1980 1968.8 1976.9 1969.7 C -1975.9 1969.8 1975.3 1970.3 1975 1971.2 C -1976.2 1969.2 1977 1971.2 1978.6 1970.9 C -1978.5 1974.4 1981.7 1972.8 1982.2 1976 C -1985.2 1976.3 1984.5 1979.3 1987 1979.6 C -1987.7 1980.3 1987.5 1982.1 1988.9 1981.7 C -1990.4 1982.4 1990.7 1985.5 1993 1985.8 C -1992.9 1994.3 1993.2 2002.3 1992.8 2011.2 C -1991.1 2012.4 1990 2014.4 1987.7 2014.6 C -1990.1 2013.4 1994.7 2012.6 1993.5 2008.8 C -[0 0.87 0.91 0.83] vc -f -0.4 w -2 J -2 M -S -n -1992.8 2010.8 m -1992.8 2001.8 1992.8 1994.1 1992.8 1985.8 C -1989.5 1985.7 1991.1 1981.1 1987.7 1981.7 C -1987.9 1978.2 1983.9 1980 1984.1 1977.2 C -1981.1 1977 1981.5 1973 1979.1 1973.1 C -1979 1972.2 1978.5 1970.9 1977.6 1970.9 C -1977.9 1971.6 1979 1971.9 1978.6 1973.1 C -1977.6 1974.9 1976.8 1973.9 1977.2 1976.2 C -1977.2 1981.5 1977 1989.4 1977.4 1994 C -1978.3 1995 1976.6 1994.1 1977.2 1994.7 C -1977 1995.3 1976.6 1997 1977.9 1997.8 C -1979 1997.5 1979.3 1998.3 1979.8 1998.8 C -1979.8 1998.9 1979.8 1999.8 1979.8 1999.2 C -1980.8 1998.7 1979.7 2000.7 1980.8 2000.7 C -1983.5 2000.4 1982.1 2003 1984.1 2003.3 C -1984.4 2004.3 1984.5 2003.7 1985.3 2004 C -1986.3 2004.6 1985.9 2006.1 1986.5 2006.9 C -1988 2007.1 1988.4 2009.7 1990.6 2009.1 C -1990.9 2006.1 1989 2000.2 1990.4 1998 C -1990.2 1994.3 1990.8 1989.2 1989.9 1986.8 C -1990.2 1984.7 1990.8 1986.2 1991.6 1985.1 C -1991.5 1985.9 1992.6 1985.5 1992.5 1986.3 C -1992 1990.5 1992.6 1995 1992 1999.2 C -1991.6 1998.9 1991.9 1998.3 1991.8 1997.8 C -1991.8 1998.5 1991.8 2000 1991.8 2000 C -1991.9 1999.9 1992 1999.8 1992 1999.7 C -1993.2 2003.5 1991.9 2007.7 1992.3 2011.5 C -1991.6 2012 1990.9 2012.2 1990.4 2012.9 C -1991.3 2011.9 1992.2 2012.1 1992.8 2010.8 C -[0 1 1 0.23] vc -f -S -n -1978.4 1968.5 m -1977 1969.2 1975.8 1968.2 1974.5 1969 C -1968.3 1973 1961.6 1976 1955.1 1979.1 C -1962 1975.9 1968.8 1972.5 1975.5 1968.8 C -1976.5 1968.8 1977.6 1968.8 1978.6 1968.8 C -1981.7 1972.1 1984.8 1975.7 1988 1978.8 C -1990.9 1981.9 1996.8 1984.6 1995.2 1990.6 C -1995.3 1988.6 1994.9 1986.9 1994.7 1985.1 C -1989.5 1979.1 1983.3 1974.3 1978.4 1968.5 C -[0.18 0.18 0 0.78] vc -f -S -n -1978.4 1968.3 m -1977.9 1968.7 1977.1 1968.5 1976.4 1968.5 C -1977.3 1968.8 1978.1 1967.9 1978.8 1968.5 C -1984 1974.3 1990.1 1979.5 1995.2 1985.6 C -1995.1 1988.4 1995.3 1985.6 1994.9 1984.8 C -1989.5 1979.4 1983.9 1973.8 1978.4 1968.3 C -[0.07 0.06 0 0.58] vc -f -S -n -1978.6 1968 m -1977.9 1968 1977.4 1968.6 1978.4 1968 C -1983.9 1973.9 1990.1 1979.1 1995.2 1985.1 C -1990.2 1979 1983.8 1974.1 1978.6 1968 C -[0.4 0.4 0 0] vc -f -S -n -1991.1 1982.2 m -1991.2 1982.9 1991.6 1984.2 1993 1984.4 C -1992.6 1983.5 1992.1 1982.5 1991.1 1982.2 C -[0 0.33 0.33 0.99] vc -f -S -n -1990.4 2012.7 m -1991.4 2011.8 1990.2 2010.9 1989.9 2010.3 C -1987.7 2010.2 1987.4 2007.6 1985.6 2007.2 C -1985.1 2006.2 1984.6 2004.5 1984.1 2004.3 C -1981.7 2004.5 1982.3 2001.2 1979.8 2000.9 C -1978.8 1999.6 1978.8 1999.1 1977.6 1998.8 C -1976.1 1997.4 1976.7 1995 1975.2 1994 C -1975.8 1994 1975 1994 1975 1993.7 C -1975.7 1993.2 1975.6 1991.8 1976 1991.3 C -1975.9 1985.7 1976.1 1979.7 1975.7 1974.5 C -1976.2 1973.3 1976.9 1971.8 1976.2 1971.4 C -1973.9 1974.3 1972.2 1973.6 1969.5 1975 C -1967.9 1977.5 1963.8 1977.1 1961.8 1980 C -1959 1980 1957.6 1983 1954.8 1982.9 C -1953.8 1984.2 1954.8 1985.7 1955.1 1987.2 C -1956.2 1989.5 1959.7 1990.1 1959.9 1991.8 C -1965.9 1998 1971.8 2005.2 1978.1 2011.7 C -1979.5 2012 1980.9 2012.7 1980.3 2014.6 C -1980.5 2015.6 1979.4 2016 1979.8 2017 C -1983 2015.6 1986.8 2014.1 1990.4 2012.7 C -[0 0.5 0.5 0.2] vc -f -S -n -1988.7 1979.6 m -1988.2 1979.9 1988.6 1980.6 1988.9 1981 C -1991.4 1982.2 1989.6 1979.9 1988.7 1979.6 C -[0 0.33 0.33 0.99] vc -f -S -n -1987.2 1978.1 m -1985 1977.5 1984.6 1974.3 1982.2 1973.6 C -1982.7 1974.5 1982.8 1975.8 1984.8 1976 C -1985.7 1976.9 1985 1978.4 1987.2 1978.1 C -f -S -n -1975.5 2084 m -1975.5 2082 1975.3 2080 1975.7 2078.2 C -1978.8 2079 1980.9 2085.5 1984.8 2083.5 C -1993 2078.7 2001.6 2075 2010 2070.8 C -2010.1 2064 2009.9 2057.2 2010.3 2050.6 C -2014.8 2046.2 2020.9 2045.7 2025.6 2042 C -2026.1 2035.1 2025.8 2028 2025.9 2021.1 C -2025.8 2027.8 2026.1 2034.6 2025.6 2041.2 C -2022.2 2044.9 2017.6 2046.8 2012.9 2048 C -2012.5 2049.5 2010.4 2049.4 2009.8 2051.1 C -2009.9 2057.6 2009.6 2064.2 2010 2070.5 C -2001.2 2075.4 1992 2079.1 1983.2 2084 C -1980.3 2082.3 1977.8 2079.2 1975.2 2077.5 C -1974.9 2079.9 1977.2 2084.6 1973.3 2085.2 C -1964.7 2088.6 1956.8 2093.7 1948.1 2097.2 C -1949 2097.3 1949.6 2096.9 1950.3 2096.7 C -1958.4 2091.9 1967.1 2088.2 1975.5 2084 C -[0.18 0.18 0 0.78] vc -f -S -n -vmrs -1948.6 2094.5 m -1950.2 2093.7 1951.8 2092.9 1953.4 2092.1 C -1951.8 2092.9 1950.2 2093.7 1948.6 2094.5 C -[0 0.87 0.91 0.83] vc -f -0.4 w -2 J -2 M -S -n -1971.6 2082.3 m -1971.6 2081.9 1970.7 2081.1 1970.9 2081.3 C -1970.7 2081.6 1970.6 2081.6 1970.4 2081.3 C -1970.8 2080.1 1968.7 2081.7 1968.3 2080.8 C -1966.6 2080.9 1966.7 2078 1964.2 2078.2 C -1964.8 2075 1960.1 2075.8 1960.1 2072.9 C -1958 2072.3 1957.5 2069.3 1955.3 2069.3 C -1953.9 2070.9 1948.8 2067.8 1950 2072 C -1949 2074 1943.2 2070.6 1944 2074.8 C -1942.2 2076.6 1937.6 2073.9 1938 2078.2 C -1936.7 2078.6 1935 2078.6 1933.7 2078.2 C -1933.5 2080 1936.8 2080.7 1937.3 2082.8 C -1939.9 2083.5 1940.6 2086.4 1942.6 2088 C -1945.2 2089.2 1946 2091.3 1948.4 2093.6 C -1956 2089.5 1963.9 2086.1 1971.6 2082.3 C -[0 0.01 1 0] vc -f -S -n -1958.2 2089.7 m -1956.4 2090 1955.6 2091.3 1953.9 2091.9 C -1955.6 2091.9 1956.5 2089.7 1958.2 2089.7 C -[0 0.87 0.91 0.83] vc -f -S -n -1929.9 2080.4 m -1929.5 2077.3 1929.7 2073.9 1929.6 2070.8 C -1929.8 2074.1 1929.2 2077.8 1930.1 2080.8 C -1935.8 2085.9 1941.4 2091.3 1946.9 2096.9 C -1941.2 2091 1935.7 2086 1929.9 2080.4 C -[0.4 0.4 0 0] vc -f -S -n -1930.1 2080.4 m -1935.8 2086 1941.5 2090.7 1946.9 2096.7 C -1941.5 2090.9 1935.7 2085.8 1930.1 2080.4 C -[0.07 0.06 0 0.58] vc -f -S -n -1940.9 2087.1 m -1941.7 2088 1944.8 2090.6 1943.6 2089.2 C -1942.5 2089 1941.6 2087.7 1940.9 2087.1 C -[0 0.87 0.91 0.83] vc -f -S -n -1972.8 2082.8 m -1973 2075.3 1972.4 2066.9 1973.3 2059.5 C -1972.5 2058.9 1972.8 2057.3 1973.1 2056.4 C -1974.8 2055.2 1973.4 2055.5 1972.4 2055.4 C -1970.1 2053.2 1967.9 2050.9 1965.6 2048.7 C -1960.9 2049.9 1956.9 2052.7 1952.4 2054.7 C -1949.3 2052.5 1946.3 2049.5 1943.6 2046.8 C -1939.9 2047.7 1936.8 2050.1 1933.5 2051.8 C -1930.9 2054.9 1933.5 2056.2 1932.3 2059.7 C -1933.2 2059.7 1932.2 2060.5 1932.5 2060.2 C -1933.2 2062.5 1931.6 2064.6 1932.5 2067.4 C -1932.9 2069.7 1932.7 2072.2 1932.8 2074.6 C -1933.6 2070.6 1932.2 2066.3 1933 2062.6 C -1934.4 2058.2 1929.8 2053.5 1935.2 2051.1 C -1937.7 2049.7 1940.2 2048 1942.8 2046.8 C -1945.9 2049.2 1948.8 2052 1951.7 2054.7 C -1952.7 2054.7 1953.6 2054.6 1954.4 2054.2 C -1958.1 2052.5 1961.7 2049.3 1965.9 2049.2 C -1968.2 2052.8 1975.2 2055 1972.6 2060.9 C -1973.3 2062.4 1972.2 2065.2 1972.6 2067.6 C -1972.7 2072.6 1972.4 2077.7 1972.8 2082.5 C -1968.1 2084.9 1963.5 2087.5 1958.7 2089.5 C -1963.5 2087.4 1968.2 2085 1972.8 2082.8 C -f -S -n -1935.2 2081.1 m -1936.8 2083.4 1938.6 2084.6 1940.4 2086.6 C -1938.8 2084.4 1936.7 2083.4 1935.2 2081.1 C -f -S -n -1983.2 2081.3 m -1984.8 2080.5 1986.3 2079.7 1988 2078.9 C -1986.3 2079.7 1984.8 2080.5 1983.2 2081.3 C -f -S -n -2006.2 2069.1 m -2006.2 2068.7 2005.2 2067.9 2005.5 2068.1 C -2005.3 2068.4 2005.2 2068.4 2005 2068.1 C -2005.4 2066.9 2003.3 2068.5 2002.8 2067.6 C -2001.2 2067.7 2001.2 2064.8 1998.8 2065 C -1999.4 2061.8 1994.7 2062.6 1994.7 2059.7 C -1992.4 2059.5 1992.4 2055.8 1990.1 2056.8 C -1985.9 2059.5 1981.1 2061 1976.9 2063.8 C -1977.2 2067.6 1974.9 2074.2 1978.8 2075.8 C -1979.6 2077.8 1981.7 2078.4 1982.9 2080.4 C -1990.6 2076.3 1998.5 2072.9 2006.2 2069.1 C -[0 0.01 1 0] vc -f -S -n -vmrs -1992.8 2076.5 m -1991 2076.8 1990.2 2078.1 1988.4 2078.7 C -1990.2 2078.7 1991 2076.5 1992.8 2076.5 C -[0 0.87 0.91 0.83] vc -f -0.4 w -2 J -2 M -S -n -1975.5 2073.4 m -1976.1 2069.7 1973.9 2064.6 1977.4 2062.4 C -1973.9 2064.5 1976.1 2069.9 1975.5 2073.6 C -1976 2074.8 1979.3 2077.4 1978.1 2076 C -1977 2075.7 1975.8 2074.5 1975.5 2073.4 C -f -S -n -2007.4 2069.6 m -2007.6 2062.1 2007 2053.7 2007.9 2046.3 C -2007.1 2045.7 2007.3 2044.1 2007.6 2043.2 C -2009.4 2042 2007.9 2042.3 2006.9 2042.2 C -2002.2 2037.4 1996.7 2032.4 1992.5 2027.3 C -1992 2027.3 1991.6 2027.3 1991.1 2027.3 C -1991.4 2035.6 1991.4 2045.6 1991.1 2054.4 C -1990.5 2055.5 1988.4 2056.6 1990.6 2055.4 C -1991.6 2055.4 1991.6 2054.1 1991.6 2053.2 C -1990.8 2044.7 1991.9 2035.4 1991.6 2027.6 C -1991.8 2027.6 1992 2027.6 1992.3 2027.6 C -1997 2032.8 2002.5 2037.7 2007.2 2042.9 C -2007.3 2044.8 2006.7 2047.4 2007.6 2048.4 C -2006.9 2055.1 2007.1 2062.5 2007.4 2069.3 C -2002.7 2071.7 1998.1 2074.3 1993.2 2076.3 C -1998 2074.2 2002.7 2071.8 2007.4 2069.6 C -f -S -n -2006.7 2069.1 m -2006.3 2068.6 2005.9 2067.7 2005.7 2066.9 C -2005.7 2059.7 2005.9 2051.4 2005.5 2045.1 C -2004.9 2045.3 2004.7 2044.5 2004.3 2045.3 C -2005.1 2045.3 2004.2 2045.8 2004.8 2046 C -2004.8 2052.2 2004.8 2059.2 2004.8 2064.5 C -2005.7 2065.7 2005.1 2065.7 2005 2066.7 C -2003.8 2067 2002.7 2067.2 2001.9 2066.4 C -2001.3 2064.6 1998 2063.1 1998 2061.9 C -1996.1 2062.3 1996.6 2058.3 1994.2 2058.8 C -1992.6 2057.7 1992.7 2054.8 1989.9 2056.6 C -1985.6 2059.3 1980.9 2060.8 1976.7 2063.6 C -1976 2066.9 1976 2071.2 1976.7 2074.6 C -1977.6 2070.8 1973.1 2062.1 1980.5 2061.2 C -1984.3 2060.3 1987.5 2058.2 1990.8 2056.4 C -1991.7 2056.8 1992.9 2057.2 1993.5 2059.2 C -1994.3 2058.6 1994.4 2060.6 1994.7 2059.2 C -1995.3 2062.7 1999.2 2061.4 1998.8 2064.8 C -2001.8 2065.4 2002.5 2068.4 2005.2 2067.4 C -2004.9 2067.9 2006 2068 2006.4 2069.1 C -2001.8 2071.1 1997.4 2073.9 1992.8 2075.8 C -1997.5 2073.8 2002 2071.2 2006.7 2069.1 C -[0 0.2 1 0] vc -f -S -n -1988.7 2056.6 m -1985.1 2058.7 1981.1 2060.1 1977.6 2061.9 C -1981.3 2060.5 1985.6 2058.1 1988.7 2056.6 C -[0 0.87 0.91 0.83] vc -f -S -n -1977.9 2059.5 m -1975.7 2064.5 1973.7 2054.7 1975.2 2060.9 C -1976 2060.6 1977.6 2059.7 1977.9 2059.5 C -f -S -n -1989.6 2051.3 m -1990.1 2042.3 1989.8 2036.6 1989.9 2028 C -1989.8 2027 1990.8 2028.3 1990.1 2027.3 C -1988.9 2026.7 1986.7 2026.9 1986.8 2024.7 C -1987.4 2023 1985.9 2024.6 1985.1 2023.7 C -1984.1 2021.4 1982.5 2020.5 1980.3 2020.6 C -1979.9 2020.8 1979.5 2021.1 1979.3 2021.6 C -1979.7 2025.8 1978.4 2033 1979.6 2038.1 C -1983.7 2042.9 1968.8 2044.6 1978.8 2042.7 C -1979.3 2042.3 1979.6 2041.9 1980 2041.5 C -1980 2034.8 1980 2027 1980 2021.6 C -1981.3 2020.5 1981.7 2021.5 1982.9 2021.8 C -1983.6 2024.7 1986.1 2023.8 1986.8 2026.4 C -1987.1 2027.7 1988.6 2027.1 1989.2 2028.3 C -1989.1 2036.7 1989.3 2044.8 1988.9 2053.7 C -1987.2 2054.9 1986.2 2056.8 1983.9 2057.1 C -1986.3 2055.9 1990.9 2055 1989.6 2051.3 C -f -S -n -1971.6 2078.9 m -1971.4 2070.5 1972.1 2062.2 1971.6 2055.9 C -1969.9 2053.7 1967.6 2051.7 1965.6 2049.6 C -1961.4 2050.4 1957.6 2053.6 1953.4 2055.2 C -1949.8 2055.6 1948.2 2051.2 1945.5 2049.6 C -1945.1 2048.8 1944.5 2047.9 1943.6 2047.5 C -1940.1 2047.8 1937.3 2051 1934 2052.3 C -1933.7 2052.6 1933.7 2053 1933.2 2053.2 C -1933.7 2060.8 1933.4 2067.2 1933.5 2074.6 C -1933.8 2068.1 1934 2060.9 1933.2 2054 C -1935.3 2050.9 1939.3 2049.6 1942.4 2047.5 C -1942.8 2047.5 1943.4 2047.4 1943.8 2047.7 C -1947.1 2050.2 1950.3 2057.9 1955.3 2054.4 C -1955.4 2054.4 1955.5 2054.3 1955.6 2054.2 C -1955.9 2057.6 1956.1 2061.8 1955.3 2064.8 C -1955.4 2064.3 1955.1 2063.8 1955.6 2063.6 C -1956 2066.6 1955.3 2068.7 1958.7 2069.8 C -1959.2 2071.7 1961.4 2071.7 1962 2074.1 C -1964.4 2074.2 1964 2077.7 1967.3 2078.4 C -1967 2079.7 1968.1 2079.9 1969 2080.1 C -1971.1 2079.9 1970 2079.2 1970.4 2078 C -1969.5 2077.2 1970.3 2075.9 1969.7 2075.1 C -1970.1 2069.8 1970.1 2063.6 1969.7 2058.8 C -1969.2 2058.5 1970 2058.1 1970.2 2057.8 C -1970.4 2058.3 1971.2 2057.7 1971.4 2058.3 C -1971.5 2065.3 1971.2 2073.6 1971.6 2081.1 C -1974.1 2081.4 1969.8 2084.3 1972.4 2082.5 C -1971.9 2081.4 1971.6 2080.2 1971.6 2078.9 C -[0 0.4 1 0] vc -f -S -n -1952.4 2052 m -1954.1 2051.3 1955.6 2050.4 1957.2 2049.6 C -1955.6 2050.4 1954.1 2051.3 1952.4 2052 C -[0 0.87 0.91 0.83] vc -f -S -n -1975.5 2039.8 m -1975.5 2039.4 1974.5 2038.7 1974.8 2038.8 C -1974.6 2039.1 1974.5 2039.1 1974.3 2038.8 C -1974.6 2037.6 1972.5 2039.3 1972.1 2038.4 C -1970.4 2038.4 1970.5 2035.5 1968 2035.7 C -1968.6 2032.5 1964 2033.3 1964 2030.4 C -1961.9 2029.8 1961.4 2026.8 1959.2 2026.8 C -1957.7 2028.5 1952.6 2025.3 1953.9 2029.5 C -1952.9 2031.5 1947 2028.2 1947.9 2032.4 C -1946 2034.2 1941.5 2031.5 1941.9 2035.7 C -1940.6 2036.1 1938.9 2036.1 1937.6 2035.7 C -1937.3 2037.5 1940.7 2038.2 1941.2 2040.3 C -1943.7 2041.1 1944.4 2043.9 1946.4 2045.6 C -1949.1 2046.7 1949.9 2048.8 1952.2 2051.1 C -1959.9 2047.1 1967.7 2043.6 1975.5 2039.8 C -[0 0.01 1 0] vc -f -S -n -vmrs -1962 2047.2 m -1960.2 2047.5 1959.5 2048.9 1957.7 2049.4 C -1959.5 2049.5 1960.3 2047.2 1962 2047.2 C -[0 0.87 0.91 0.83] vc -f -0.4 w -2 J -2 M -S -n -2012.4 2046.3 m -2010.3 2051.3 2008.3 2041.5 2009.8 2047.7 C -2010.5 2047.4 2012.2 2046.5 2012.4 2046.3 C -f -S -n -1944.8 2044.6 m -1945.5 2045.6 1948.6 2048.1 1947.4 2046.8 C -1946.3 2046.5 1945.5 2045.2 1944.8 2044.6 C -f -S -n -1987.2 2054.9 m -1983.7 2057.3 1979.6 2058 1976 2060.2 C -1974.7 2058.2 1977.2 2055.8 1974.3 2054.9 C -1973.1 2052 1970.4 2050.2 1968 2048 C -1968 2047.7 1968 2047.4 1968.3 2047.2 C -1969.5 2046.1 1983 2040.8 1972.4 2044.8 C -1971.2 2046.6 1967.9 2046 1968 2048.2 C -1970.5 2050.7 1973.8 2052.6 1974.3 2055.6 C -1975.1 2055 1975.7 2056.7 1975.7 2057.1 C -1975.7 2058.2 1974.8 2059.3 1975.5 2060.4 C -1979.3 2058.2 1983.9 2057.7 1987.2 2054.9 C -[0.18 0.18 0 0.78] vc -f -S -n -1967.8 2047.5 m -1968.5 2047 1969.1 2046.5 1969.7 2046 C -1969.1 2046.5 1968.5 2047 1967.8 2047.5 C -[0 0.87 0.91 0.83] vc -f -S -n -1976.7 2040.3 m -1976.9 2032.8 1976.3 2024.4 1977.2 2017 C -1976.4 2016.5 1976.6 2014.8 1976.9 2013.9 C -1978.7 2012.7 1977.2 2013 1976.2 2012.9 C -1971.5 2008.1 1965.9 2003.1 1961.8 1998 C -1960.9 1998 1960.1 1998 1959.2 1998 C -1951.5 2001.1 1944.3 2005.5 1937.1 2009.6 C -1935 2012.9 1937 2013.6 1936.1 2017.2 C -1937.1 2017.2 1936 2018 1936.4 2017.7 C -1937 2020.1 1935.5 2022.1 1936.4 2024.9 C -1936.8 2027.2 1936.5 2029.7 1936.6 2032.1 C -1937.4 2028.2 1936 2023.8 1936.8 2020.1 C -1938.3 2015.7 1933.6 2011 1939 2008.6 C -1945.9 2004.5 1953.1 2000.3 1960.6 1998.3 C -1960.9 1998.3 1961.3 1998.3 1961.6 1998.3 C -1966.2 2003.5 1971.8 2008.4 1976.4 2013.6 C -1976.6 2015.5 1976 2018.1 1976.9 2019.2 C -1976.1 2025.8 1976.4 2033.2 1976.7 2040 C -1971.9 2042.4 1967.4 2045 1962.5 2047 C -1967.3 2044.9 1972 2042.6 1976.7 2040.3 C -f -S -n -1939 2038.6 m -1940.6 2040.9 1942.5 2042.1 1944.3 2044.1 C -1942.7 2041.9 1940.6 2040.9 1939 2038.6 C -f -S -n -2006.2 2065.7 m -2006 2057.3 2006.7 2049 2006.2 2042.7 C -2002.1 2038.4 1997.7 2033.4 1993 2030 C -1992.9 2029.3 1992.5 2028.6 1992 2028.3 C -1992.1 2036.6 1991.9 2046.2 1992.3 2054.9 C -1990.8 2056.2 1989 2056.7 1987.5 2058 C -1988.7 2057.7 1990.7 2054.4 1993 2056.4 C -1993.4 2058.8 1996 2058.2 1996.6 2060.9 C -1999 2061 1998.5 2064.5 2001.9 2065.2 C -2001.5 2066.5 2002.7 2066.7 2003.6 2066.9 C -2005.7 2066.7 2004.6 2066 2005 2064.8 C -2004 2064 2004.8 2062.7 2004.3 2061.9 C -2004.6 2056.6 2004.6 2050.4 2004.3 2045.6 C -2003.7 2045.3 2004.6 2044.9 2004.8 2044.6 C -2005 2045.1 2005.7 2044.5 2006 2045.1 C -2006 2052.1 2005.8 2060.4 2006.2 2067.9 C -2008.7 2068.2 2004.4 2071.1 2006.9 2069.3 C -2006.4 2068.2 2006.2 2067 2006.2 2065.7 C -[0 0.4 1 0] vc -f -S -n -2021.8 2041.7 m -2018.3 2044.1 2014.1 2044.8 2010.5 2047 C -2009.3 2045 2011.7 2042.6 2008.8 2041.7 C -2004.3 2035.1 1997.6 2030.9 1993 2024.4 C -1992.1 2024 1991.5 2024.3 1990.8 2024 C -1993.2 2023.9 1995.3 2027.1 1996.8 2029 C -2000.4 2032.6 2004.9 2036.9 2008.4 2040.8 C -2008.2 2043.1 2011.4 2042.8 2009.8 2045.8 C -2009.8 2046.3 2009.7 2046.9 2010 2047.2 C -2013.8 2045 2018.5 2044.5 2021.8 2041.7 C -[0.18 0.18 0 0.78] vc -f -S -n -2001.6 2034 m -2000.7 2033.1 1999.9 2032.3 1999 2031.4 C -1999.9 2032.3 2000.7 2033.1 2001.6 2034 C -[0 0.87 0.91 0.83] vc -f -S -n -vmrs -1989.4 2024.4 m -1989.5 2025.4 1988.6 2024.3 1988.9 2024.7 C -1990.5 2025.8 1990.7 2024.2 1992.8 2024.9 C -1993.8 2025.9 1995 2027.1 1995.9 2028 C -1994.3 2026 1991.9 2023.4 1989.4 2024.4 C -[0 0.87 0.91 0.83] vc -f -0.4 w -2 J -2 M -S -n -1984.8 2019.9 m -1984.6 2018.6 1986.3 2017.2 1987.7 2016.8 C -1987.2 2017.5 1982.9 2017.9 1984.4 2020.6 C -1984.1 2019.9 1984.9 2020 1984.8 2019.9 C -f -S -n -1981.7 2017 m -1979.6 2022 1977.6 2012.3 1979.1 2018.4 C -1979.8 2018.1 1981.5 2017.2 1981.7 2017 C -f -S -n -1884.3 2019.2 m -1884.7 2010.5 1884.5 2000.6 1884.5 1991.8 C -1886.6 1989.3 1889.9 1988.9 1892.4 1987 C -1890.8 1988.7 1886 1989.1 1884.3 1992.3 C -1884.7 2001 1884.5 2011.3 1884.5 2019.9 C -1891 2025.1 1895.7 2031.5 1902 2036.9 C -1896.1 2031 1890 2024.9 1884.3 2019.2 C -[0.07 0.06 0 0.58] vc -f -S -n -1884 2019.4 m -1884.5 2010.6 1884.2 2000.4 1884.3 1991.8 C -1884.8 1990.4 1887.8 1989 1884.8 1990.8 C -1884.3 1991.3 1884.3 1992 1884 1992.5 C -1884.5 2001.2 1884.2 2011.1 1884.3 2019.9 C -1887.9 2023.1 1891.1 2026.4 1894.4 2030 C -1891.7 2026.1 1887.1 2022.9 1884 2019.4 C -[0.4 0.4 0 0] vc -f -S -n -1885 2011.7 m -1885 2006.9 1885 2001.9 1885 1997.1 C -1885 2001.9 1885 2006.9 1885 2011.7 C -[0 0.87 0.91 0.83] vc -f -S -n -1975.5 2036.4 m -1975.2 2028 1976 2019.7 1975.5 2013.4 C -1971.1 2008.5 1965.6 2003.6 1961.6 1999 C -1958.8 1998 1956 2000 1953.6 2001.2 C -1948.2 2004.7 1941.9 2006.5 1937.1 2010.8 C -1937.5 2018.3 1937.3 2024.7 1937.3 2032.1 C -1937.6 2025.6 1937.9 2018.4 1937.1 2011.5 C -1937.3 2011 1937.6 2010.5 1937.8 2010 C -1944.6 2005.7 1951.9 2002.3 1959.2 1999 C -1960.1 1998.5 1960.1 1999.8 1960.4 2000.4 C -1959.7 2006.9 1959.7 2014.2 1959.4 2021.1 C -1959 2021.1 1959.2 2021.9 1959.2 2022.3 C -1959.2 2021.9 1959 2021.3 1959.4 2021.1 C -1959.8 2024.1 1959.2 2026.2 1962.5 2027.3 C -1963 2029.2 1965.3 2029.2 1965.9 2031.6 C -1968.3 2031.8 1967.8 2035.2 1971.2 2036 C -1970.8 2037.2 1971.9 2037.5 1972.8 2037.6 C -1974.9 2037.4 1973.9 2036.7 1974.3 2035.5 C -1973.3 2034.7 1974.1 2033.4 1973.6 2032.6 C -1973.9 2027.3 1973.9 2021.1 1973.6 2016.3 C -1973 2016 1973.9 2015.6 1974 2015.3 C -1974.3 2015.9 1975 2015.3 1975.2 2015.8 C -1975.3 2022.8 1975.1 2031.2 1975.5 2038.6 C -1977.9 2039 1973.7 2041.8 1976.2 2040 C -1975.7 2039 1975.5 2037.8 1975.5 2036.4 C -[0 0.4 1 0] vc -f -S -n -1991.1 2012.4 m -1987.5 2014.8 1983.4 2015.6 1979.8 2017.7 C -1978.5 2015.7 1981 2013.3 1978.1 2012.4 C -1973.6 2005.8 1966.8 2001.6 1962.3 1995.2 C -1961.4 1994.7 1960.8 1995 1960.1 1994.7 C -1962.5 1994.6 1964.6 1997.8 1966.1 1999.7 C -1969.7 2003.3 1974.2 2007.6 1977.6 2011.5 C -1977.5 2013.8 1980.6 2013.5 1979.1 2016.5 C -1979.1 2017 1979 2017.6 1979.3 2018 C -1983.1 2015.7 1987.8 2015.2 1991.1 2012.4 C -[0.18 0.18 0 0.78] vc -f -S -n -1970.9 2004.8 m -1970 2003.9 1969.2 2003 1968.3 2002.1 C -1969.2 2003 1970 2003.9 1970.9 2004.8 C -[0 0.87 0.91 0.83] vc -f -S -n -1887.9 1994.9 m -1888.5 1992.3 1891.4 1992.2 1893.2 1990.8 C -1898.4 1987.5 1904 1984.8 1909.5 1982.2 C -1909.7 1982.7 1910.3 1982.1 1910.4 1982.7 C -1909.5 1990.5 1910.1 1996.4 1910 2004.5 C -1909.1 2003.4 1909.7 2005.8 1909.5 2006.4 C -1910.4 2006 1909.7 2008 1910.2 2007.9 C -1911.3 2010.6 1912.5 2012.6 1915.7 2013.4 C -1915.8 2013.7 1915.5 2014.4 1916 2014.4 C -1916.3 2015 1915.4 2016 1915.2 2016 C -1916.1 2015.5 1916.5 2014.5 1916 2013.6 C -1913.4 2013.3 1913.1 2010.5 1910.9 2009.8 C -1910.7 2008.8 1910.4 2007.9 1910.2 2006.9 C -1911.1 1998.8 1909.4 1990.7 1910.7 1982.4 C -1910 1982.1 1908.9 1982.1 1908.3 1982.4 C -1901.9 1986.1 1895 1988.7 1888.8 1993 C -1888 1993.4 1888.4 1994.3 1887.6 1994.7 C -1888.1 2001.3 1887.8 2008.6 1887.9 2015.1 C -1887.3 2017.5 1887.9 2015.4 1888.4 2014.4 C -1887.8 2008 1888.4 2001.3 1887.9 1994.9 C -[0.07 0.06 0 0.58] vc -f -S -n -vmrs -1887.9 2018.4 m -1887.5 2016.9 1888.5 2016 1888.8 2014.8 C -1890.1 2014.8 1891.1 2016.6 1892.4 2015.3 C -1892.4 2014.4 1893.8 2012.9 1894.4 2012.4 C -1895.9 2012.4 1896.6 2013.9 1897.7 2012.7 C -1898.4 2011.7 1898.6 2010.4 1899.6 2009.8 C -1901.7 2009.9 1902.9 2010.4 1904 2009.1 C -1904.3 2007.4 1904 2007.6 1904.9 2007.2 C -1906.2 2007 1907.6 2006.5 1908.8 2006.7 C -1910.6 2008.2 1909.8 2011.5 1912.6 2012 C -1912.4 2013 1913.8 2012.7 1914 2013.2 C -1911.5 2011.1 1909.1 2007.9 1909.2 2004.3 C -1909.5 2003.5 1909.9 2004.9 1909.7 2004.3 C -1909.9 1996.2 1909.3 1990.5 1910.2 1982.7 C -1909.5 1982.6 1909.5 1982.6 1908.8 1982.7 C -1903.1 1985.7 1897 1987.9 1891.7 1992 C -1890.5 1993 1888.2 1992.9 1888.1 1994.9 C -1888.7 2001.4 1888.1 2008.4 1888.6 2014.8 C -1888.3 2016 1887.2 2016.9 1887.6 2018.4 C -1892.3 2023.9 1897.6 2027.9 1902.3 2033.3 C -1898 2028.2 1892.1 2023.8 1887.9 2018.4 C -[0.4 0.4 0 0] vc -f -0.4 w -2 J -2 M -S -n -1910.9 1995.2 m -1910.4 1999.8 1911 2003.3 1910.9 2008.1 C -1910.9 2003.8 1910.9 1999.2 1910.9 1995.2 C -[0.18 0.18 0 0.78] vc -f -S -n -1911.2 2004.3 m -1911.2 2001.9 1911.2 1999.7 1911.2 1997.3 C -1911.2 1999.7 1911.2 2001.9 1911.2 2004.3 C -[0 0.87 0.91 0.83] vc -f -S -n -1958.7 1995.2 m -1959 1995.6 1956.2 1995 1956.5 1996.8 C -1955.8 1997.6 1954.2 1998.5 1953.6 1997.3 C -1953.6 1990.8 1954.9 1989.6 1953.4 1983.9 C -1953.4 1983.3 1953.3 1982.1 1954.4 1982 C -1955.5 1982.6 1956.5 1981.3 1957.5 1981 C -1956.3 1981.8 1954.7 1982.6 1953.9 1981.5 C -1951.4 1983 1954.7 1988.8 1952.9 1990.6 C -1953.8 1990.6 1953.2 1992.7 1953.4 1993.7 C -1953.8 1994.5 1952.3 1996.1 1953.2 1997.8 C -1956.3 1999.4 1957.5 1994 1959.9 1995.6 C -1962 1994.4 1963.7 1997.7 1965.2 1998.8 C -1963.5 1996.7 1961.2 1994.1 1958.7 1995.2 C -f -S -n -1945 2000.7 m -1945.4 1998.7 1945.4 1997.9 1945 1995.9 C -1944.5 1995.3 1944.2 1992.6 1945.7 1993.2 C -1946 1992.2 1948.7 1992.5 1948.4 1990.6 C -1947.5 1990.3 1948.1 1988.7 1947.9 1988.2 C -1948.9 1987.8 1950.5 1986.8 1950.5 1984.6 C -1951.5 1980.9 1946.7 1983 1947.2 1979.8 C -1944.5 1979.9 1945.2 1976.6 1943.1 1976.7 C -1941.8 1975.7 1942.1 1972.7 1939.2 1973.8 C -1938.2 1974.6 1939.3 1971.6 1938.3 1970.9 C -1938.8 1969.2 1933.4 1970.3 1937.3 1970 C -1939.4 1971.2 1937.2 1973 1937.6 1974.3 C -1937.2 1976.3 1937.1 1981.2 1937.8 1984.1 C -1938.8 1982.3 1937.9 1976.6 1938.5 1973.1 C -1938.9 1975 1938.5 1976.4 1939.7 1977.2 C -1939.5 1983.5 1938.9 1991.3 1940.2 1997.3 C -1939.4 1999.1 1938.6 1997.1 1937.8 1997.1 C -1937.4 1996.7 1937.6 1996.1 1937.6 1995.6 C -1936.5 1998.5 1940.1 1998.4 1940.9 2000.7 C -1942.1 2000.4 1943.2 2001.3 1943.1 2002.4 C -1943.6 2003.1 1941.1 2004.6 1942.8 2003.8 C -1943.9 2002.5 1942.6 2000.6 1945 2000.7 C -[0.65 0.65 0 0.42] vc -f -S -n -1914.5 2006.4 m -1914.1 2004.9 1915.2 2004 1915.5 2002.8 C -1916.7 2002.8 1917.8 2004.6 1919.1 2003.3 C -1919 2002.4 1920.4 2000.9 1921 2000.4 C -1922.5 2000.4 1923.2 2001.9 1924.4 2000.7 C -1925 1999.7 1925.3 1998.4 1926.3 1997.8 C -1928.4 1997.9 1929.5 1998.4 1930.6 1997.1 C -1930.9 1995.4 1930.7 1995.6 1931.6 1995.2 C -1932.8 1995 1934.3 1994.5 1935.4 1994.7 C -1936.1 1995.8 1936.9 1996.2 1936.6 1997.8 C -1938.9 1999.4 1939.7 2001.3 1942.4 2002.4 C -1942.4 2002.5 1942.2 2003 1942.6 2002.8 C -1942.9 2000.4 1939.2 2001.8 1939.2 1999.7 C -1936.2 1998.6 1937 1995.3 1935.9 1993.5 C -1937.1 1986.5 1935.2 1977.9 1937.6 1971.2 C -1937.6 1970.3 1936.6 1971 1936.4 1970.4 C -1930.2 1973.4 1924 1976 1918.4 1980 C -1917.2 1981 1914.9 1980.9 1914.8 1982.9 C -1915.3 1989.4 1914.7 1996.4 1915.2 2002.8 C -1914.9 2004 1913.9 2004.9 1914.3 2006.4 C -1919 2011.9 1924.2 2015.9 1928.9 2021.3 C -1924.6 2016.2 1918.7 2011.8 1914.5 2006.4 C -[0.4 0.4 0 0] vc -f -S -n -1914.5 1982.9 m -1915.1 1980.3 1918 1980.2 1919.8 1978.8 C -1925 1975.5 1930.6 1972.8 1936.1 1970.2 C -1939.4 1970.6 1936.1 1974.2 1936.6 1976.4 C -1936.5 1981.9 1936.8 1987.5 1936.4 1992.8 C -1935.9 1992.8 1936.2 1993.5 1936.1 1994 C -1937.1 1993.6 1936.2 1995.9 1936.8 1995.9 C -1937 1998 1939.5 1999.7 1940.4 2000.7 C -1940.1 1998.6 1935 1997.2 1937.6 1993.7 C -1938.3 1985.7 1935.9 1976.8 1937.8 1970.7 C -1936.9 1969.8 1935.4 1970.3 1934.4 1970.7 C -1928.3 1974.4 1921.4 1976.7 1915.5 1981 C -1914.6 1981.4 1915.1 1982.3 1914.3 1982.7 C -1914.7 1989.3 1914.5 1996.6 1914.5 2003.1 C -1913.9 2005.5 1914.5 2003.4 1915 2002.4 C -1914.5 1996 1915.1 1989.3 1914.5 1982.9 C -[0.07 0.06 0 0.58] vc -f -S -n -1939.2 1994.9 m -1939.3 1995 1939.4 1995.1 1939.5 1995.2 C -1939.1 1989 1939.3 1981.6 1939 1976.7 C -1938.6 1976.3 1938.6 1974.6 1938.5 1973.3 C -1938.7 1976.1 1938.1 1979.4 1939 1981.7 C -1937.3 1986 1937.7 1991.6 1938 1996.4 C -1937.3 1994.3 1939.6 1996.2 1939.2 1994.9 C -[0.18 0.18 0 0.78] vc -f -S -n -1938.3 1988.4 m -1938.5 1990.5 1937.9 1994.1 1938.8 1994.7 C -1937.9 1992.6 1939 1990.6 1938.3 1988.4 C -[0 0.87 0.91 0.83] vc -f -S -n -1938.8 1985.8 m -1938.5 1985.9 1938.4 1985.7 1938.3 1985.6 C -1938.4 1986.2 1938 1989.5 1938.8 1987.2 C -1938.8 1986.8 1938.8 1986.3 1938.8 1985.8 C -f -S -n -vmrs -1972.8 2062.1 m -1971.9 2061 1972.5 2059.4 1972.4 2058 C -1972.2 2063.8 1971.9 2073.7 1972.4 2081.3 C -1972.5 2074.9 1971.9 2067.9 1972.8 2062.1 C -[0 1 1 0.36] vc -f -0.4 w -2 J -2 M -S -n -1940.2 2071.7 m -1941.3 2072 1943.1 2072.3 1944 2071.5 C -1943.6 2069.9 1945.2 2069.1 1946 2068.8 C -1950 2071.1 1948.7 2065.9 1951.7 2066.2 C -1953.5 2063.9 1956.9 2069.4 1955.6 2063.8 C -1955.5 2064.2 1955.7 2064.8 1955.3 2065 C -1954.3 2063.7 1956.2 2063.6 1955.6 2062.1 C -1954.5 2060 1958.3 2050.3 1952.2 2055.6 C -1949.1 2053.8 1946 2051 1943.8 2048 C -1940.3 2048 1937.5 2051.3 1934.2 2052.5 C -1933.1 2054.6 1934.4 2057.3 1934 2060 C -1934 2065.1 1934 2069.7 1934 2074.6 C -1934.4 2069 1934.1 2061.5 1934.2 2054.9 C -1934.6 2054.5 1935.3 2054.7 1935.9 2054.7 C -1937 2055.3 1935.9 2056.1 1935.9 2056.8 C -1936.5 2063 1935.6 2070.5 1935.9 2074.6 C -1936.7 2074.4 1937.3 2075.2 1938 2074.6 C -1937.9 2073.6 1939.1 2072.1 1940.2 2071.7 C -[0 0.2 1 0] vc -f -S -n -1933.2 2074.1 m -1933.2 2071.5 1933.2 2069 1933.2 2066.4 C -1933.2 2069 1933.2 2071.5 1933.2 2074.1 C -[0 1 1 0.36] vc -f -S -n -2007.4 2048.9 m -2006.5 2047.8 2007.1 2046.2 2006.9 2044.8 C -2006.7 2050.6 2006.5 2060.5 2006.9 2068.1 C -2007.1 2061.7 2006.5 2054.7 2007.4 2048.9 C -f -S -n -1927.2 2062.4 m -1925.8 2060.1 1928.1 2058.2 1927 2056.4 C -1927.3 2055.5 1926.5 2053.5 1926.8 2051.8 C -1926.8 2052.8 1926 2052.5 1925.3 2052.5 C -1924.1 2052.8 1925 2050.5 1924.4 2050.1 C -1925.3 2050.2 1925.4 2048.8 1926.3 2049.4 C -1926.5 2052.3 1928.4 2047.2 1928.4 2051.1 C -1928.9 2050.5 1929 2051.4 1928.9 2051.8 C -1928.9 2052 1928.9 2052.3 1928.9 2052.5 C -1929.4 2051.4 1928.9 2049 1930.1 2048.2 C -1928.9 2047.1 1930.5 2047.1 1930.4 2046.5 C -1931.9 2046.2 1933.1 2046.1 1934.7 2046.5 C -1934.6 2046.9 1935.2 2047.9 1934.4 2048.4 C -1936.9 2048.1 1933.6 2043.8 1935.9 2043.9 C -1935.7 2043.9 1934.8 2041.3 1933.2 2041.7 C -1932.5 2041.6 1932.4 2039.6 1932.3 2041 C -1930.8 2042.6 1929 2040.6 1927.7 2042 C -1927.5 2041.4 1927.1 2040.9 1927.2 2040.3 C -1927.8 2040.6 1927.4 2039.1 1928.2 2038.6 C -1929.4 2038 1930.5 2038.8 1931.3 2037.9 C -1931.7 2039 1932.5 2038.6 1931.8 2037.6 C -1930.9 2037 1928.7 2037.8 1928.2 2037.9 C -1926.7 2037.8 1928 2039 1927 2038.8 C -1927.4 2040.4 1925.6 2040.8 1925.1 2041 C -1924.3 2040.4 1923.2 2040.5 1922.2 2040.5 C -1921.4 2041.7 1921 2043.9 1919.3 2043.9 C -1918.8 2043.4 1917.2 2043.3 1916.4 2043.4 C -1915.9 2044.4 1915.7 2046 1914.3 2046.5 C -1913.1 2046.6 1912 2044.5 1911.4 2046.3 C -1912.8 2046.5 1913.8 2047.4 1915.7 2047 C -1916.9 2047.7 1915.6 2048.8 1916 2049.4 C -1915.4 2049.3 1913.9 2050.3 1913.3 2051.1 C -1913.9 2054.1 1916 2050.2 1916.7 2053 C -1916.9 2053.8 1915.5 2054.1 1916.7 2054.4 C -1917 2054.7 1920.2 2054.3 1919.3 2056.6 C -1918.8 2056.1 1920.2 2058.6 1920.3 2057.6 C -1921.2 2057.9 1922.1 2057.5 1922.4 2059 C -1922.3 2059.1 1922.2 2059.3 1922 2059.2 C -1922.1 2059.7 1922.4 2060.3 1922.9 2060.7 C -1923.2 2060.1 1923.8 2060.4 1924.6 2060.7 C -1925.9 2062.6 1923.2 2062 1925.6 2063.6 C -1926.1 2063.1 1927.3 2062.5 1927.2 2062.4 C -[0.21 0.21 0 0] vc -f -S -n -1933.2 2063.3 m -1933.2 2060.7 1933.2 2058.2 1933.2 2055.6 C -1933.2 2058.2 1933.2 2060.7 1933.2 2063.3 C -[0 1 1 0.36] vc -f -S -n -1965.2 2049.2 m -1967.1 2050.1 1969.9 2053.7 1972.1 2056.4 C -1970.5 2054 1967.6 2051.3 1965.2 2049.2 C -f -S -n -1991.8 2034.8 m -1991.7 2041.5 1992 2048.5 1991.6 2055.2 C -1990.5 2056.4 1991.9 2054.9 1991.8 2054.4 C -1991.8 2047.9 1991.8 2041.3 1991.8 2034.8 C -f -S -n -1988.9 2053.2 m -1988.9 2044.3 1988.9 2036.6 1988.9 2028.3 C -1985.7 2028.2 1987.2 2023.5 1983.9 2024.2 C -1983.9 2022.4 1982 2021.6 1981 2021.3 C -1980.6 2021.1 1980.6 2021.7 1980.3 2021.6 C -1980.3 2027 1980.3 2034.8 1980.3 2041.5 C -1979.3 2043.2 1977.6 2043 1976.2 2043.6 C -1977.1 2043.8 1978.5 2043.2 1978.8 2044.1 C -1978.5 2045.3 1979.9 2045.3 1980.3 2045.8 C -1980.5 2046.8 1980.7 2046.2 1981.5 2046.5 C -1982.4 2047.1 1982 2048.6 1982.7 2049.4 C -1984.2 2049.6 1984.6 2052.2 1986.8 2051.6 C -1987.1 2048.6 1985.1 2042.7 1986.5 2040.5 C -1986.3 2036.7 1986.9 2031.7 1986 2029.2 C -1986.3 2027.1 1986.9 2028.6 1987.7 2027.6 C -1987.7 2028.3 1988.7 2028 1988.7 2028.8 C -1988.1 2033 1988.7 2037.5 1988.2 2041.7 C -1987.8 2041.4 1988 2040.8 1988 2040.3 C -1988 2041 1988 2042.4 1988 2042.4 C -1988 2042.4 1988.1 2042.3 1988.2 2042.2 C -1989.3 2046 1988 2050.2 1988.4 2054 C -1987.8 2054.4 1987.1 2054.7 1986.5 2055.4 C -1987.4 2054.4 1988.4 2054.6 1988.9 2053.2 C -[0 1 1 0.23] vc -f -S -n -1950.8 2054.4 m -1949.7 2053.4 1948.7 2052.3 1947.6 2051.3 C -1948.7 2052.3 1949.7 2053.4 1950.8 2054.4 C -[0 1 1 0.36] vc -f -S -n -vmrs -2006.7 2043.2 m -2004.5 2040.8 2002.4 2038.4 2000.2 2036 C -2002.4 2038.4 2004.5 2040.8 2006.7 2043.2 C -[0 1 1 0.36] vc -f -0.4 w -2 J -2 M -S -n -1976.7 2019.6 m -1975.8 2018.6 1976.4 2016.9 1976.2 2015.6 C -1976 2021.3 1975.8 2031.2 1976.2 2038.8 C -1976.4 2032.4 1975.8 2025.5 1976.7 2019.6 C -f -S -n -1988.4 2053.5 m -1988.6 2049.2 1988.1 2042.8 1988 2040 C -1988.4 2040.4 1988.1 2041 1988.2 2041.5 C -1988.3 2037.2 1988 2032.7 1988.4 2028.5 C -1987.6 2027.1 1987.2 2028.6 1986.8 2028 C -1985.9 2028.5 1986.5 2029.7 1986.3 2030.4 C -1986.9 2029.8 1986.6 2031 1987 2031.2 C -1987.4 2039.6 1985 2043 1987.2 2050.4 C -1987.2 2051.6 1985.9 2052.3 1984.6 2051.3 C -1981.9 2049.7 1982.9 2047 1980.3 2046.5 C -1980.3 2045.2 1978.1 2046.2 1978.6 2043.9 C -1975.6 2043.3 1979.3 2045.6 1979.6 2046.5 C -1980.8 2046.6 1981.5 2048.5 1982.2 2049.9 C -1983.7 2050.8 1984.8 2052.8 1986.5 2053 C -1986.7 2053.5 1987.5 2054.1 1987 2054.7 C -1987.4 2053.9 1988.3 2054.3 1988.4 2053.5 C -[0 1 1 0.23] vc -f -S -n -1988 2038.1 m -1988 2036.7 1988 2035.4 1988 2034 C -1988 2035.4 1988 2036.7 1988 2038.1 C -[0 1 1 0.36] vc -f -S -n -1999.7 2035.7 m -1997.6 2033.5 1995.4 2031.2 1993.2 2029 C -1995.4 2031.2 1997.6 2033.5 1999.7 2035.7 C -f -S -n -1944 2029.2 m -1945.2 2029.5 1946.9 2029.8 1947.9 2029 C -1947.4 2027.4 1949 2026.7 1949.8 2026.4 C -1953.9 2028.6 1952.6 2023.4 1955.6 2023.7 C -1957.4 2021.4 1960.7 2027 1959.4 2021.3 C -1959.3 2021.7 1959.6 2022.3 1959.2 2022.5 C -1958.1 2021.2 1960.1 2021.1 1959.4 2019.6 C -1959.1 2012.7 1959.9 2005.1 1959.6 1999.2 C -1955.3 2000.1 1951.3 2003.1 1947.2 2005 C -1943.9 2006 1941.2 2008.7 1938 2010 C -1936.9 2012.1 1938.2 2014.8 1937.8 2017.5 C -1937.8 2022.6 1937.8 2027.3 1937.8 2032.1 C -1938.2 2026.5 1938 2019 1938 2012.4 C -1938.5 2012 1939.2 2012.3 1939.7 2012.2 C -1940.8 2012.8 1939.7 2013.6 1939.7 2014.4 C -1940.4 2020.5 1939.4 2028 1939.7 2032.1 C -1940.6 2031.9 1941.2 2032.7 1941.9 2032.1 C -1941.7 2031.2 1943 2029.7 1944 2029.2 C -[0 0.2 1 0] vc -f -S -n -1937.1 2031.6 m -1937.1 2029.1 1937.1 2026.5 1937.1 2024 C -1937.1 2026.5 1937.1 2029.1 1937.1 2031.6 C -[0 1 1 0.36] vc -f -S -n -1991.8 2028 m -1992.5 2027.8 1993.2 2029.9 1994 2030.2 C -1992.9 2029.6 1993.1 2028.1 1991.8 2028 C -[0 1 1 0.23] vc -f -S -n -1991.8 2027.8 m -1992.4 2027.6 1992.6 2028.3 1993 2028.5 C -1992.6 2028.2 1992.2 2027.6 1991.6 2027.8 C -1991.6 2028.5 1991.6 2029.1 1991.6 2029.7 C -1991.6 2029.1 1991.4 2028.3 1991.8 2027.8 C -[0 1 1 0.36] vc -f -S -n -1985.8 2025.4 m -1985.3 2025.2 1984.8 2024.7 1984.1 2024.9 C -1983.3 2025.3 1983.6 2027.3 1983.9 2027.6 C -1985 2028 1986.9 2026.9 1985.8 2025.4 C -[0 1 1 0.23] vc -f -S -n -vmrs -1993.5 2024.4 m -1992.4 2023.7 1991.3 2022.9 1990.1 2023.2 C -1990.7 2023.7 1989.8 2023.8 1989.4 2023.7 C -1989.1 2023.7 1988.6 2023.9 1988.4 2023.5 C -1988.5 2023.2 1988.3 2022.7 1988.7 2022.5 C -1989 2022.6 1988.9 2023 1988.9 2023.2 C -1989.1 2022.8 1990.4 2022.3 1990.6 2021.3 C -1990.4 2021.8 1990 2021.3 1990.1 2021.1 C -1990.1 2020.9 1990.1 2020.1 1990.1 2020.6 C -1989.9 2021.1 1989.5 2020.6 1989.6 2020.4 C -1989.6 2019.8 1988.7 2019.6 1988.2 2019.2 C -1987.5 2018.7 1987.7 2020.2 1987 2019.4 C -1987.5 2020.4 1986 2021.1 1987.5 2021.8 C -1986.8 2023.1 1986.6 2021.1 1986 2021.1 C -1986.1 2020.1 1985.9 2019 1986.3 2018.2 C -1986.7 2018.4 1986.5 2019 1986.5 2019.4 C -1986.5 2018.7 1986.4 2017.8 1987.2 2017.7 C -1986.5 2017.2 1985.5 2019.3 1985.3 2020.4 C -1986.2 2022 1987.3 2023.5 1989.2 2024.2 C -1990.8 2024.3 1991.6 2022.9 1993.2 2024.4 C -1993.8 2025.4 1995 2026.6 1995.9 2027.1 C -1995 2026.5 1994.1 2025.5 1993.5 2024.4 C -[0 1 1 0.36] vc -f -0.4 w -2 J -2 M -[0 0.5 0.5 0.2] vc -S -n -2023 2040.3 m -2023.2 2036 2022.7 2029.6 2022.5 2026.8 C -2022.9 2027.2 2022.7 2027.8 2022.8 2028.3 C -2022.8 2024 2022.6 2019.5 2023 2015.3 C -2022.2 2013.9 2021.7 2015.4 2021.3 2014.8 C -2020.4 2015.3 2021 2016.5 2020.8 2017.2 C -2021.4 2016.6 2021.1 2017.8 2021.6 2018 C -2022 2026.4 2019.6 2029.8 2021.8 2037.2 C -2021.7 2038.4 2020.5 2039.1 2019.2 2038.1 C -2016.5 2036.5 2017.5 2033.8 2014.8 2033.3 C -2014.9 2032 2012.6 2033 2013.2 2030.7 C -2011.9 2030.8 2011.2 2030.1 2010.8 2029.2 C -2010.8 2029.1 2010.8 2028.2 2010.8 2028.8 C -2010 2028.8 2010.4 2026.5 2008.6 2027.3 C -2007.9 2026.6 2007.3 2025.9 2007.9 2027.1 C -2009.7 2028 2010 2030.1 2012.2 2030.9 C -2012.9 2032.1 2013.7 2033.6 2015.1 2033.6 C -2015.7 2035.1 2016.9 2036.7 2018.4 2038.4 C -2019.8 2039.3 2022 2039.4 2021.6 2041.5 C -2021.9 2040.7 2022.9 2041.1 2023 2040.3 C -[0 1 1 0.23] vc -f -S -n -2022.5 2024.9 m -2022.5 2023.5 2022.5 2022.2 2022.5 2020.8 C -2022.5 2022.2 2022.5 2023.5 2022.5 2024.9 C -[0 1 1 0.36] vc -f -S -n -1983.2 2022.8 m -1982.4 2022.5 1982.1 2021.6 1981.2 2022.3 C -1981.1 2022.9 1980.5 2024 1981 2024.2 C -1981.8 2024.6 1982.9 2024.4 1983.2 2022.8 C -[0 1 1 0.23] vc -f -S -n -1931.1 2019.9 m -1929.6 2017.7 1932 2015.7 1930.8 2013.9 C -1931.1 2013 1930.3 2011 1930.6 2009.3 C -1930.6 2010.3 1929.8 2010 1929.2 2010 C -1928 2010.3 1928.8 2008.1 1928.2 2007.6 C -1929.1 2007.8 1929.3 2006.3 1930.1 2006.9 C -1930.3 2009.8 1932.2 2004.8 1932.3 2008.6 C -1932.7 2008 1932.8 2009 1932.8 2009.3 C -1932.8 2009.6 1932.8 2009.8 1932.8 2010 C -1933.2 2009 1932.7 2006.6 1934 2005.7 C -1932.7 2004.6 1934.3 2004.6 1934.2 2004 C -1935.8 2003.7 1937 2003.6 1938.5 2004 C -1938.5 2004.5 1939.1 2005.4 1938.3 2006 C -1940.7 2005.7 1937.4 2001.3 1939.7 2001.4 C -1939.5 2001.4 1938.6 1998.8 1937.1 1999.2 C -1936.3 1999.1 1936.2 1997.1 1936.1 1998.5 C -1934.7 2000.1 1932.9 1998.2 1931.6 1999.5 C -1931.3 1998.9 1930.9 1998.5 1931.1 1997.8 C -1931.6 1998.2 1931.3 1996.6 1932 1996.1 C -1933.2 1995.5 1934.3 1996.4 1935.2 1995.4 C -1935.5 1996.5 1936.3 1996.1 1935.6 1995.2 C -1934.7 1994.5 1932.5 1995.3 1932 1995.4 C -1930.5 1995.3 1931.9 1996.5 1930.8 1996.4 C -1931.2 1997.9 1929.5 1998.3 1928.9 1998.5 C -1928.1 1997.9 1927.1 1998 1926 1998 C -1925.3 1999.2 1924.8 2001.4 1923.2 2001.4 C -1922.6 2000.9 1921 2000.9 1920.3 2000.9 C -1919.7 2001.9 1919.6 2003.5 1918.1 2004 C -1916.9 2004.1 1915.8 2002 1915.2 2003.8 C -1916.7 2004 1917.6 2004.9 1919.6 2004.5 C -1920.7 2005.2 1919.4 2006.3 1919.8 2006.9 C -1919.2 2006.9 1917.7 2007.8 1917.2 2008.6 C -1917.8 2011.6 1919.8 2007.8 1920.5 2010.5 C -1920.8 2011.3 1919.3 2011.6 1920.5 2012 C -1920.8 2012.3 1924 2011.8 1923.2 2014.1 C -1922.6 2013.6 1924.1 2016.1 1924.1 2015.1 C -1925.1 2015.4 1925.9 2015 1926.3 2016.5 C -1926.2 2016.6 1926 2016.8 1925.8 2016.8 C -1925.9 2017.2 1926.2 2017.8 1926.8 2018.2 C -1927.1 2017.6 1927.7 2018 1928.4 2018.2 C -1929.7 2020.1 1927.1 2019.5 1929.4 2021.1 C -1929.9 2020.7 1931.1 2020 1931.1 2019.9 C -[0.21 0.21 0 0] vc -f -S -n -1937.1 2020.8 m -1937.1 2018.3 1937.1 2015.7 1937.1 2013.2 C -1937.1 2015.7 1937.1 2018.3 1937.1 2020.8 C -[0 1 1 0.36] vc -f -S -n -2020.4 2012.2 m -2019.8 2012 2019.3 2011.5 2018.7 2011.7 C -2017.9 2012.1 2018.1 2014.1 2018.4 2014.4 C -2019.6 2014.8 2021.4 2013.7 2020.4 2012.2 C -[0 1 1 0.23] vc -f -S -n -1976 2013.9 m -1973.8 2011.5 1971.6 2009.1 1969.5 2006.7 C -1971.6 2009.1 1973.8 2011.5 1976 2013.9 C -[0 1 1 0.36] vc -f -S -n -1995.4 2012.7 m -1996.1 2010.3 1993.8 2006.2 1997.3 2005.7 C -1998.9 2005.4 2000 2003.7 2001.4 2003.1 C -2003.9 2003.1 2005.3 2001.3 2006.9 1999.7 C -2004.5 2003.5 2000 2002.2 1997.6 2005.7 C -1996.5 2005.9 1994.8 2006.1 1995.2 2007.6 C -1995.7 2009.4 1995.2 2011.6 1994.7 2012.9 C -1992 2015.8 1987.8 2015.7 1985.3 2018.7 C -1988.3 2016.3 1992.3 2015.3 1995.4 2012.7 C -[0.18 0.18 0 0.78] vc -f -S -n -1995.6 2012.4 m -1995.6 2011.2 1995.6 2010 1995.6 2008.8 C -1995.6 2010 1995.6 2011.2 1995.6 2012.4 C -[0 1 1 0.36] vc -f -S -n -vmrs -2017.7 2009.6 m -2016.9 2009.3 2016.7 2008.4 2015.8 2009.1 C -2014.2 2010.6 2016 2010.6 2016.5 2011.5 C -2017.2 2010.9 2018.1 2010.8 2017.7 2009.6 C -[0 1 1 0.23] vc -f -0.4 w -2 J -2 M -S -n -2014.4 2006.4 m -2013.5 2006.8 2012.1 2005.6 2012 2006.7 C -2013 2007.3 2011.9 2009.2 2012.9 2008.4 C -2014.2 2008.3 2014.6 2007.8 2014.4 2006.4 C -f -S -n -1969 2006.4 m -1966.5 2003.8 1964 2001.2 1961.6 1998.5 C -1964 2001.2 1966.5 2003.8 1969 2006.4 C -[0 1 1 0.36] vc -f -S -n -2012 2005.2 m -2012.2 2004.2 2011.4 2003.3 2010.3 2003.3 C -2009 2003.6 2010 2004.7 2009.6 2004.8 C -2009.3 2005.7 2011.4 2006.7 2012 2005.2 C -[0 1 1 0.23] vc -f -S -n -1962.8 1995.2 m -1961.7 1994.4 1960.6 1993.7 1959.4 1994 C -1959.5 1994.9 1957.5 1994.1 1956.8 1994.7 C -1955.9 1995.5 1956.7 1997 1955.1 1997.3 C -1956.9 1996.7 1956.8 1994 1959.2 1994.7 C -1961.1 1991 1968.9 2003.2 1962.8 1995.2 C -[0 1 1 0.36] vc -f -S -n -1954.6 1995.6 m -1955.9 1994.7 1955.1 1989.8 1955.3 1988 C -1954.5 1988.3 1954.9 1986.6 1954.4 1986 C -1955.7 1989.2 1953.9 1991.1 1954.8 1994.2 C -1954.5 1995.9 1953.5 1995.3 1953.9 1997.3 C -1955.3 1998.3 1953.2 1995.5 1954.6 1995.6 C -f -S -n -1992.3 2011 m -1992.5 2006.7 1992 2000.3 1991.8 1997.6 C -1992.2 1997.9 1992 1998.5 1992 1999 C -1992.1 1994.7 1991.9 1990.2 1992.3 1986 C -1991.4 1984.6 1991 1986.1 1990.6 1985.6 C -1989.7 1986 1990.3 1987.2 1990.1 1988 C -1990.7 1987.4 1990.4 1988.5 1990.8 1988.7 C -1991.3 1997.1 1988.9 2000.6 1991.1 2007.9 C -1991 2009.1 1989.8 2009.9 1988.4 2008.8 C -1985.7 2007.2 1986.8 2004.5 1984.1 2004 C -1984.2 2002.7 1981.9 2003.7 1982.4 2001.4 C -1981.2 2001.5 1980.5 2000.8 1980 2000 C -1980 1999.8 1980 1998.9 1980 1999.5 C -1979.3 1999.5 1979.7 1997.2 1977.9 1998 C -1977.2 1997.3 1976.6 1996.7 1977.2 1997.8 C -1979 1998.7 1979.3 2000.8 1981.5 2001.6 C -1982.2 2002.8 1983 2004.3 1984.4 2004.3 C -1985 2005.8 1986.2 2007.5 1987.7 2009.1 C -1989 2010 1991.3 2010.2 1990.8 2012.2 C -1991.2 2011.4 1992.2 2011.8 1992.3 2011 C -[0 1 1 0.23] vc -f -S -n -1991.8 1995.6 m -1991.8 1994.3 1991.8 1992.9 1991.8 1991.6 C -1991.8 1992.9 1991.8 1994.3 1991.8 1995.6 C -[0 1 1 0.36] vc -f -S -n -1959.2 1994.2 m -1958.8 1993.3 1960.7 1993.9 1961.1 1993.7 C -1961.5 1993.9 1961.2 1994.4 1961.8 1994.2 C -1960.9 1994 1960.8 1992.9 1959.9 1992.5 C -1959.6 1993.5 1958.3 1993.5 1958.2 1994.2 C -1958.1 1994.1 1958 1994 1958 1994 C -1957.2 1994.9 1958 1993.4 1956.8 1993 C -1955.6 1992.5 1956 1991 1956.3 1989.9 C -1956.5 1989.8 1956.6 1990 1956.8 1990.1 C -1957.1 1989 1956 1989.1 1955.8 1988.2 C -1955.1 1990.4 1956.2 1995 1954.8 1995.9 C -1954.1 1995.5 1954.5 1996.5 1954.4 1997.1 C -1955 1996.8 1954.8 1997.4 1955.6 1996.8 C -1956 1996 1956.3 1993.2 1958.7 1994.2 C -1958.9 1994.2 1959.7 1994.2 1959.2 1994.2 C -[0 1 1 0.23] vc -f -S -n -1958.2 1994 m -1958.4 1993.5 1959.7 1993.1 1959.9 1992 C -1959.7 1992.5 1959.3 1992 1959.4 1991.8 C -1959.4 1991.6 1959.4 1990.8 1959.4 1991.3 C -1959.2 1991.8 1958.8 1991.3 1958.9 1991.1 C -1958.9 1990.5 1958 1990.3 1957.5 1989.9 C -1956.8 1989.5 1956.9 1991 1956.3 1990.1 C -1956.7 1991 1955.4 1992.1 1956.5 1992.3 C -1956.8 1993.5 1958.3 1992.9 1957.2 1994 C -1957.8 1994.3 1958.1 1992.4 1958.2 1994 C -[0 0.5 0.5 0.2] vc -f -S -n -vmrs -1954.4 1982.7 m -1956.1 1982.7 1954.1 1982.5 1953.9 1982.9 C -1953.9 1983.7 1953.7 1984.7 1954.1 1985.3 C -1954.4 1984.2 1953.6 1983.6 1954.4 1982.7 C -[0 1 1 0.36] vc -f -0.4 w -2 J -2 M -S -n -1989.6 1982.9 m -1989.1 1982.7 1988.6 1982.3 1988 1982.4 C -1987.2 1982.8 1987.4 1984.8 1987.7 1985.1 C -1988.9 1985.6 1990.7 1984.4 1989.6 1982.9 C -[0 1 1 0.23] vc -f -S -n -1987 1980.3 m -1986.2 1980 1986 1979.1 1985.1 1979.8 C -1983.5 1981.4 1985.3 1981.4 1985.8 1982.2 C -1986.5 1981.7 1987.4 1981.5 1987 1980.3 C -f -S -n -1983.6 1977.2 m -1982.7 1977.5 1981.4 1976.3 1981.2 1977.4 C -1982.3 1978 1981.2 1979.9 1982.2 1979.1 C -1983.5 1979 1983.9 1978.5 1983.6 1977.2 C -f -S -n -1981.2 1976 m -1981.5 1974.9 1980.6 1974 1979.6 1974 C -1978.3 1974.3 1979.3 1975.4 1978.8 1975.5 C -1978.6 1976.4 1980.7 1977.4 1981.2 1976 C -f -S -n -1972.1 2082.3 m -1971.8 2081.8 1971.3 2080.9 1971.2 2080.1 C -1971.1 2072.9 1971.3 2064.6 1970.9 2058.3 C -1970.3 2058.5 1970.1 2057.7 1969.7 2058.5 C -1970.6 2058.5 1969.7 2059 1970.2 2059.2 C -1970.2 2065.4 1970.2 2072.4 1970.2 2077.7 C -1971.1 2078.9 1970.6 2078.9 1970.4 2079.9 C -1969.2 2080.2 1968.2 2080.4 1967.3 2079.6 C -1966.8 2077.8 1963.4 2076.3 1963.5 2075.1 C -1961.5 2075.5 1962 2071.5 1959.6 2072 C -1959.2 2070 1956.5 2069.3 1955.8 2067.6 C -1956 2068.4 1955.3 2069.7 1956.5 2069.8 C -1958.6 2068.9 1958.1 2073.5 1960.1 2072.4 C -1960.7 2075.9 1964.7 2074.6 1964.2 2078 C -1967.2 2078.6 1967.9 2081.6 1970.7 2080.6 C -1970.3 2081.1 1971.5 2081.2 1971.9 2082.3 C -1967.2 2084.3 1962.9 2087.1 1958.2 2089 C -1962.9 2087 1967.4 2084.4 1972.1 2082.3 C -[0 0.2 1 0] vc -f -S -n -1971.9 2080.1 m -1971.9 2075.1 1971.9 2070 1971.9 2065 C -1971.9 2070 1971.9 2075.1 1971.9 2080.1 C -[0 1 1 0.23] vc -f -S -n -2010.8 2050.6 m -2013.2 2049 2010.5 2050.1 2010.5 2051.3 C -2010.5 2057.7 2010.5 2064.1 2010.5 2070.5 C -2008.7 2072.4 2006 2073.3 2003.6 2074.4 C -2016.4 2073.7 2008 2058.4 2010.8 2050.6 C -[0.4 0.4 0 0] vc -f -S -n -2006.4 2066.9 m -2006.4 2061.9 2006.4 2056.8 2006.4 2051.8 C -2006.4 2056.8 2006.4 2061.9 2006.4 2066.9 C -[0 1 1 0.23] vc -f -S -n -1971.9 2060.7 m -1972.2 2060.3 1971.4 2068.2 1972.4 2061.9 C -1971.8 2061.6 1972.4 2060.9 1971.9 2060.7 C -f -S -n -vmrs -1986.5 2055.2 m -1987.5 2054.3 1986.3 2053.4 1986 2052.8 C -1983.8 2052.7 1983.6 2050.1 1981.7 2049.6 C -1981.2 2048.7 1980.8 2047 1980.3 2046.8 C -1978.5 2047 1978 2044.6 1976.7 2043.9 C -1974 2044.4 1972 2046.6 1969.2 2047 C -1969 2047.2 1968.8 2047.5 1968.5 2047.7 C -1970.6 2049.6 1973.1 2051.3 1974.3 2054.2 C -1975.7 2054.5 1977 2055.2 1976.4 2057.1 C -1976.7 2058 1975.5 2058.5 1976 2059.5 C -1979.2 2058 1983 2056.6 1986.5 2055.2 C -[0 0.5 0.5 0.2] vc -f -0.4 w -2 J -2 M -S -n -1970.2 2054.2 m -1971.5 2055.3 1972.5 2056.8 1972.1 2058.3 C -1972.8 2056.5 1971.6 2055.6 1970.2 2054.2 C -[0 1 1 0.23] vc -f -S -n -1992 2052.5 m -1992 2053.4 1992.2 2054.4 1991.8 2055.2 C -1992.2 2054.4 1992 2053.4 1992 2052.5 C -f -S -n -1957.2 2053 m -1958.1 2052.6 1959 2052.2 1959.9 2051.8 C -1959 2052.2 1958.1 2052.6 1957.2 2053 C -f -S -n -2006.4 2047.5 m -2006.8 2047.1 2006 2055 2006.9 2048.7 C -2006.4 2048.4 2007 2047.7 2006.4 2047.5 C -f -S -n -2004.8 2041 m -2006.1 2042.1 2007.1 2043.6 2006.7 2045.1 C -2007.3 2043.3 2006.2 2042.4 2004.8 2041 C -f -S -n -1976 2039.8 m -1975.6 2039.3 1975.2 2038.4 1975 2037.6 C -1974.9 2030.4 1975.2 2022.1 1974.8 2015.8 C -1974.2 2016 1974 2015.3 1973.6 2016 C -1974.4 2016 1973.5 2016.5 1974 2016.8 C -1974 2022.9 1974 2030 1974 2035.2 C -1974.9 2036.4 1974.4 2036.4 1974.3 2037.4 C -1973.1 2037.7 1972 2037.9 1971.2 2037.2 C -1970.6 2035.3 1967.3 2033.9 1967.3 2032.6 C -1965.3 2033 1965.9 2029.1 1963.5 2029.5 C -1963 2027.6 1960.4 2026.8 1959.6 2025.2 C -1959.8 2025.9 1959.2 2027.2 1960.4 2027.3 C -1962.5 2026.4 1961.9 2031 1964 2030 C -1964.6 2033.4 1968.5 2032.1 1968 2035.5 C -1971 2036.1 1971.8 2039.1 1974.5 2038.1 C -1974.2 2038.7 1975.3 2038.7 1975.7 2039.8 C -1971 2041.8 1966.7 2044.6 1962 2046.5 C -1966.8 2044.5 1971.3 2041.9 1976 2039.8 C -[0 0.2 1 0] vc -f -S -n -1975.7 2037.6 m -1975.7 2032.6 1975.7 2027.6 1975.7 2022.5 C -1975.7 2027.6 1975.7 2032.6 1975.7 2037.6 C -[0 1 1 0.23] vc -f -S -n -1992 2035.5 m -1992 2034.2 1992 2032.9 1992 2031.6 C -1992 2032.9 1992 2034.2 1992 2035.5 C -f -S -n -2015.3 2036 m -2015.4 2034.1 2013.3 2034 2012.9 2033.3 C -2011.5 2031 2009.3 2029.4 2007.4 2028 C -2006.9 2027.1 2006.6 2023.8 2005 2024.9 C -2004 2024.9 2002.9 2024.9 2001.9 2024.9 C -2001.4 2026.5 2001 2028.4 2003.8 2028.3 C -2006.6 2030.4 2008.9 2033.7 2011.2 2036.2 C -2011.8 2036.4 2012.9 2035.8 2012.9 2036.7 C -2013 2035.5 2015.3 2037.4 2015.3 2036 C -[0 0 0 0] vc -f -S -n -vmrs -2009.1 2030.4 m -2009.1 2029 2007.5 2029.4 2006.9 2028.3 C -2007.2 2027.1 2006.5 2025.5 2005.7 2024.7 C -2004.6 2025.1 2003.1 2024.9 2001.9 2024.9 C -2001.8 2026.2 2000.9 2027 2002.4 2028 C -2004.5 2027.3 2004.9 2029.4 2006.9 2029 C -2007 2030.2 2007.6 2030.7 2008.4 2031.4 C -2008.8 2031.5 2009.1 2031.1 2009.1 2030.4 C -[0 0 0 0.18] vc -f -0.4 w -2 J -2 M -S -n -2003.8 2029.5 m -2003 2029.4 2001.9 2029.1 2002.4 2030.4 C -2003.1 2031.3 2005.2 2030.3 2003.8 2029.5 C -[0 1 1 0.23] vc -f -S -n -1999.2 2025.2 m -1999.1 2025.6 1998 2025.7 1998.8 2026.6 C -2000.9 2028.5 1999.5 2023.4 1999.2 2025.2 C -f -S -n -2007.6 2024.2 m -2007.6 2022.9 2008.4 2024.2 2007.6 2022.8 C -2007.6 2017.5 2007.8 2009.1 2007.4 2003.8 C -2007.9 2003.7 2008.7 2002.8 2009.1 2002.1 C -2009.6 2000.8 2008.3 2000.8 2007.9 2000.2 C -2004.9 2000 2008.9 2001.3 2007.2 2002.1 C -2006.7 2007.7 2007 2015.1 2006.9 2021.1 C -2006.7 2022.1 2005.4 2022.8 2006.2 2023.5 C -2006.6 2023.1 2008 2025.9 2007.6 2024.2 C -f -S -n -1989.9 2023.5 m -1989.5 2022.6 1991.4 2023.2 1991.8 2023 C -1992.2 2023.2 1991.9 2023.7 1992.5 2023.5 C -1991.6 2023.2 1991.6 2022.2 1990.6 2021.8 C -1990.4 2022.8 1989 2022.8 1988.9 2023.5 C -1988.5 2023 1988.7 2022.6 1988.7 2023.5 C -1989.1 2023.5 1990.2 2023.5 1989.9 2023.5 C -f -[0 0.5 0.5 0.2] vc -S -n -2003.3 2023.5 m -2003.1 2023.3 2003.1 2023.2 2003.3 2023 C -2003.7 2023.1 2003.9 2022.9 2003.8 2022.5 C -2003.4 2022.2 2001.2 2022.3 2002.4 2023 C -2002.6 2022.9 2002.7 2023.1 2002.8 2023.2 C -2000.7 2023.7 2003.9 2023.4 2003.3 2023.5 C -[0 1 1 0.23] vc -f -S -n -1986.8 2019.4 m -1987.8 2019.8 1987.5 2018.6 1987.2 2018 C -1986.2 2017.8 1987.3 2020.5 1986.3 2019.2 C -1986.3 2017.7 1986.3 2020.6 1986.3 2021.3 C -1988.5 2023.1 1985.6 2020.3 1986.8 2019.4 C -f -S -n -1975.7 2018.2 m -1976.1 2017.8 1975.2 2025.7 1976.2 2019.4 C -1975.7 2019.2 1976.3 2018.4 1975.7 2018.2 C -f -S -n -1974 2011.7 m -1975.4 2012.8 1976.4 2014.3 1976 2015.8 C -1976.6 2014 1975.5 2013.1 1974 2011.7 C -f -S -n -1984.6 2006.7 m -1984.7 2004.8 1982.6 2004.8 1982.2 2004 C -1980.8 2001.7 1978.6 2000.1 1976.7 1998.8 C -1976.1 1997.8 1975.8 1994.5 1974.3 1995.6 C -1973.3 1995.6 1972.2 1995.6 1971.2 1995.6 C -1970.7 1997.2 1970.3 1999.1 1973.1 1999 C -1975.8 2001.2 1978.2 2004.4 1980.5 2006.9 C -1981.1 2007.1 1982.1 2006.5 1982.2 2007.4 C -1982.3 2006.2 1984.5 2008.1 1984.6 2006.7 C -[0 0 0 0] vc -f -S -n -vmrs -1978.4 2001.2 m -1978.4 1999.7 1976.8 2000.1 1976.2 1999 C -1976.5 1997.8 1975.8 1996.2 1975 1995.4 C -1973.9 1995.8 1972.4 1995.6 1971.2 1995.6 C -1971 1997 1970.2 1997.7 1971.6 1998.8 C -1973.8 1998 1974.2 2000.1 1976.2 1999.7 C -1976.3 2000.9 1976.9 2001.4 1977.6 2002.1 C -1978.1 2002.2 1978.4 2001.8 1978.4 2001.2 C -[0 0 0 0.18] vc -f -0.4 w -2 J -2 M -S -n -1973.1 2000.2 m -1972.3 2000.1 1971.2 1999.8 1971.6 2001.2 C -1972.4 2002 1974.5 2001 1973.1 2000.2 C -[0 1 1 0.23] vc -f -S -n -1960.8 1998.5 m -1961.6 1998.2 1962.6 2000.3 1963.2 2000.9 C -1962.3 2000.1 1962.2 1998.7 1960.8 1998.5 C -f -S -n -1968.5 1995.9 m -1968.4 1996.4 1967.3 1996.4 1968 1997.3 C -1970.1 1999.2 1968.8 1994.1 1968.5 1995.9 C -f -S -n -1976.9 1994.9 m -1976.9 1993.7 1977.6 1994.9 1976.9 1993.5 C -1976.9 1988.2 1977.1 1979.8 1976.7 1974.5 C -1977.2 1974.5 1978 1973.5 1978.4 1972.8 C -1978.8 1971.5 1977.6 1971.5 1977.2 1970.9 C -1974.2 1970.7 1978.2 1972 1976.4 1972.8 C -1976 1978.4 1976.3 1985.8 1976.2 1991.8 C -1976 1992.8 1974.6 1993.5 1975.5 1994.2 C -1975.9 1993.8 1977.3 1996.6 1976.9 1994.9 C -f -S -n -1972.6 1994.2 m -1972.4 1994 1972.4 1993.9 1972.6 1993.7 C -1973 1993.8 1973.1 1993.7 1973.1 1993.2 C -1972.7 1992.9 1970.5 1993.1 1971.6 1993.7 C -1971.9 1993.7 1972 1993.8 1972.1 1994 C -1970 1994.4 1973.1 1994.1 1972.6 1994.2 C -f -S -n -1948.1 2093.8 m -1947 2092.7 1945.9 2091.6 1944.8 2090.4 C -1945.9 2091.6 1947 2092.7 1948.1 2093.8 C -[0 0.4 1 0] vc -f -S -n -1953.4 2091.4 m -1954.8 2090.7 1956.3 2090 1957.7 2089.2 C -1956.3 2090 1954.8 2090.7 1953.4 2091.4 C -[0 0.2 1 0] vc -f -S -n -1954.1 2091.4 m -1956.6 2089.6 1957.2 2089.6 1954.1 2091.4 C -[0 0.4 1 0] vc -f -S -n -1962.3 2087.3 m -1963.7 2086.6 1965.2 2085.9 1966.6 2085.2 C -1965.2 2085.9 1963.7 2086.6 1962.3 2087.3 C -f -S -n -vmrs -1967.1 2084.9 m -1968.3 2084.4 1969.7 2083.8 1970.9 2083.2 C -1969.7 2083.8 1968.3 2084.4 1967.1 2084.9 C -[0 0.4 1 0] vc -f -0.4 w -2 J -2 M -S -n -1982.7 2080.6 m -1981.5 2079.5 1980.5 2078.4 1979.3 2077.2 C -1980.5 2078.4 1981.5 2079.5 1982.7 2080.6 C -f -S -n -1988 2078.2 m -1989.4 2077.5 1990.8 2076.8 1992.3 2076 C -1990.8 2076.8 1989.4 2077.5 1988 2078.2 C -[0 0.2 1 0] vc -f -S -n -1988.7 2078.2 m -1991.1 2076.4 1991.8 2076.4 1988.7 2078.2 C -[0 0.4 1 0] vc -f -S -n -1976.2 2063.8 m -1978.6 2062.2 1976 2063.3 1976 2064.5 C -1976.1 2067.8 1975.5 2071.4 1976.4 2074.4 C -1975.7 2071.1 1975.9 2067.2 1976.2 2063.8 C -f -S -n -1996.8 2074.1 m -1998.3 2073.4 1999.7 2072.7 2001.2 2072 C -1999.7 2072.7 1998.3 2073.4 1996.8 2074.1 C -f -S -n -2001.6 2071.7 m -2002.9 2071.2 2004.2 2070.6 2005.5 2070 C -2004.2 2070.6 2002.9 2071.2 2001.6 2071.7 C -f -S -n -1981.5 2060.7 m -1980.2 2061.2 1978.9 2061.5 1977.9 2062.6 C -1978.9 2061.5 1980.2 2061.2 1981.5 2060.7 C -f -S -n -1982 2060.4 m -1982.7 2060.1 1983.6 2059.8 1984.4 2059.5 C -1983.6 2059.8 1982.7 2060.1 1982 2060.4 C -f -S -n -1952 2051.3 m -1950.8 2050.2 1949.7 2049.1 1948.6 2048 C -1949.7 2049.1 1950.8 2050.2 1952 2051.3 C -f -S -n -vmrs -1977.4 2047.7 m -1975.8 2047.8 1974.8 2046.1 1974.5 2045.3 C -1974.9 2044.4 1976 2044.5 1976.7 2044.8 C -1977.9 2045 1977 2048.4 1979.3 2047.5 C -1979.9 2047.5 1980.8 2048.6 1979.8 2049.2 C -1978.2 2050.4 1980.8 2049.5 1980.3 2049.4 C -1981.4 2049.8 1980.3 2048.4 1980.3 2048 C -1979.8 2047.5 1979 2046.6 1978.4 2046.5 C -1977.3 2045.9 1977.2 2043.3 1975.2 2044.6 C -1974.7 2045.3 1973.6 2045 1973.3 2045.8 C -1975 2046.3 1975.8 2049.8 1978.1 2049.4 C -1978.4 2050.9 1978.7 2048.5 1977.9 2049.2 C -1977.7 2048.7 1977.2 2047.8 1977.4 2047.7 C -[0 0.5 0.5 0.2] vc -f -0.4 w -2 J -2 M -S -n -1957.2 2048.9 m -1958.7 2048.2 1960.1 2047.5 1961.6 2046.8 C -1960.1 2047.5 1958.7 2048.2 1957.2 2048.9 C -[0 0.2 1 0] vc -f -S -n -1958 2048.9 m -1960.4 2047.1 1961.1 2047.1 1958 2048.9 C -[0 0.4 1 0] vc -f -S -n -1966.1 2044.8 m -1967.6 2044.1 1969 2043.4 1970.4 2042.7 C -1969 2043.4 1967.6 2044.1 1966.1 2044.8 C -f -S -n -1970.9 2042.4 m -1972.2 2041.9 1973.5 2041.3 1974.8 2040.8 C -1973.5 2041.3 1972.2 2041.9 1970.9 2042.4 C -f -S -n -2012 2034.5 m -2010.4 2034.6 2009.3 2032.9 2009.1 2032.1 C -2009.4 2031 2010.3 2031.3 2011.2 2031.6 C -2012.5 2031.8 2011.6 2035.2 2013.9 2034.3 C -2014.4 2034.3 2015.4 2035.4 2014.4 2036 C -2012.7 2037.2 2015.3 2036.3 2014.8 2036.2 C -2015.9 2036.6 2014.8 2035.2 2014.8 2034.8 C -2014.4 2034.3 2013.6 2033.4 2012.9 2033.3 C -2011.5 2031 2009.3 2029.4 2007.4 2028 C -2007.5 2026.5 2007.3 2027.9 2007.2 2028.3 C -2007.9 2028.8 2008.7 2029.1 2009.3 2030 C -2009.6 2030.7 2009 2031.9 2008.4 2031.6 C -2006.7 2031 2007.7 2028 2005 2028.8 C -2004.8 2028.6 2004.3 2028.2 2003.8 2028.3 C -2006.6 2030.4 2008.9 2033.7 2011.2 2036.2 C -2011.8 2036.4 2012.9 2035.8 2012.9 2036.7 C -2012.7 2036.1 2011.8 2035 2012 2034.5 C -[0 0.5 0.5 0.2] vc -f -S -n -1981.2 2005.2 m -1979.7 2005.3 1978.6 2003.6 1978.4 2002.8 C -1978.7 2001.8 1979.6 2002.1 1980.5 2002.4 C -1981.8 2002.5 1980.9 2005.9 1983.2 2005 C -1983.7 2005.1 1984.7 2006.1 1983.6 2006.7 C -1982 2007.9 1984.6 2007 1984.1 2006.9 C -1985.2 2007.3 1984.1 2006 1984.1 2005.5 C -1983.6 2005 1982.9 2004.1 1982.2 2004 C -1980.8 2001.7 1978.6 2000.1 1976.7 1998.8 C -1976.7 1997.2 1976.6 1998.6 1976.4 1999 C -1977.2 1999.5 1978 1999.8 1978.6 2000.7 C -1978.8 2001.5 1978.3 2002.7 1977.6 2002.4 C -1976 2001.8 1977 1998.7 1974.3 1999.5 C -1974.1 1999.3 1973.6 1998.9 1973.1 1999 C -1975.8 2001.2 1978.2 2004.4 1980.5 2006.9 C -1981.1 2007.1 1982.1 2006.5 1982.2 2007.4 C -1982 2006.8 1981.1 2005.7 1981.2 2005.2 C -f -S -n -1966.8 1976.4 m -1969.4 1973 1974.4 1974.6 1976.2 1970.4 C -1972.7 1974 1968 1975.1 1964 1977.4 C -1960.9 1979.9 1957.1 1981.8 1953.9 1982.7 C -1958.4 1981.1 1962.6 1978.8 1966.8 1976.4 C -[0.18 0.18 0 0.78] vc -f -S -n -1948.4 2093.8 m -1949.8 2093.1 1951.2 2092.5 1952.7 2091.9 C -1951.2 2092.5 1949.8 2093.1 1948.4 2093.8 C -[0 0.2 1 0] vc -f -S -n -1948.1 2093.6 m -1947.3 2092.8 1946.5 2091.9 1945.7 2091.2 C -1946.5 2091.9 1947.3 2092.8 1948.1 2093.6 C -f -S -n -vmrs -1942.1 2087.8 m -1943.5 2088.4 1944.3 2089.5 1945.2 2090.7 C -1944.8 2089.3 1943.3 2088.3 1942.1 2087.8 C -[0 0.2 1 0] vc -f -0.4 w -2 J -2 M -S -n -1933.5 2078.4 m -1933.5 2078 1933.2 2079 1933.7 2079.4 C -1935 2080.4 1936.2 2081.3 1937.1 2082.8 C -1936.7 2080.7 1933.7 2080.7 1933.5 2078.4 C -f -S -n -1982.9 2080.6 m -1984.4 2079.9 1985.8 2079.3 1987.2 2078.7 C -1985.8 2079.3 1984.4 2079.9 1982.9 2080.6 C -f -S -n -1982.7 2080.4 m -1981.9 2079.6 1981.1 2078.7 1980.3 2078 C -1981.1 2078.7 1981.9 2079.6 1982.7 2080.4 C -f -S -n -1977.4 2075.1 m -1977.9 2075.3 1979.1 2076.4 1979.8 2077.5 C -1979 2076.8 1978.7 2075.1 1977.4 2075.1 C -f -S -n -1952.2 2051.3 m -1953.6 2050.7 1955.1 2050.1 1956.5 2049.4 C -1955.1 2050.1 1953.6 2050.7 1952.2 2051.3 C -f -S -n -1952 2051.1 m -1951.2 2050.3 1950.3 2049.5 1949.6 2048.7 C -1950.3 2049.5 1951.2 2050.3 1952 2051.1 C -f -S -n -1946 2045.3 m -1947.3 2045.9 1948.1 2047 1949.1 2048.2 C -1948.6 2046.8 1947.1 2045.8 1946 2045.3 C -f -S -n -1937.3 2036 m -1937.4 2035.5 1937 2036.5 1937.6 2036.9 C -1938.8 2037.9 1940.1 2038.8 1940.9 2040.3 C -1940.6 2038.2 1937.6 2038.2 1937.3 2036 C -f -S -n -1935.2 2073.2 m -1936.4 2069.9 1935.8 2061.8 1935.6 2056.4 C -1935.8 2055.9 1936.3 2055.7 1936.1 2055.2 C -1935.7 2054.7 1935 2055 1934.4 2054.9 C -1934.4 2061.5 1934.4 2068.7 1934.4 2074.6 C -1935.7 2075.1 1936 2073.7 1935.2 2073.2 C -[0 0.01 1 0] vc -f -S -n -vmrs -1939 2030.7 m -1940.3 2027.4 1939.7 2019.3 1939.5 2013.9 C -1939.7 2013.5 1940.1 2013.2 1940 2012.7 C -1939.5 2012.3 1938.8 2012.5 1938.3 2012.4 C -1938.3 2019 1938.3 2026.2 1938.3 2032.1 C -1939.5 2032.7 1939.8 2031.2 1939 2030.7 C -[0 0.01 1 0] vc -f -0.4 w -2 J -2 M -S -n -1975.2 2077.2 m -1975.3 2077.3 1975.4 2077.4 1975.5 2077.5 C -1974.7 2073.2 1974.9 2067.5 1975.2 2063.6 C -1975.4 2064 1974.6 2063.9 1974.8 2064.3 C -1974.9 2069.9 1974.3 2076.5 1975.2 2081.1 C -1974.9 2079.9 1974.9 2078.4 1975.2 2077.2 C -[0.92 0.92 0 0.67] vc -f -S -n -1930.8 2067.4 m -1931.5 2070.1 1929.6 2072.1 1930.6 2074.6 C -1931 2072.6 1930.8 2069.8 1930.8 2067.4 C -f -S -n -2010 2050.1 m -2009.8 2050.5 2009.5 2050.9 2009.3 2051.1 C -2009.5 2056.7 2008.9 2063.3 2009.8 2067.9 C -2009.5 2062.1 2009.3 2054.7 2010 2050.1 C -f -S -n -1930.1 2060.9 m -1929.3 2057.1 1930.7 2054.8 1929.9 2051.3 C -1930.2 2050.2 1931.1 2049.6 1931.8 2049.2 C -1931.4 2049.6 1930.4 2049.5 1930.1 2050.1 C -1928.4 2054.8 1933.4 2063.5 1925.3 2064.3 C -1927.2 2063.9 1928.5 2062.1 1930.1 2060.9 C -[0.07 0.06 0 0.58] vc -f -S -n -1929.6 2061.2 m -1929.6 2057.6 1929.6 2054.1 1929.6 2050.6 C -1930 2049.9 1930.5 2049.4 1931.1 2049.2 C -1930 2048.6 1930.5 2050.2 1929.4 2049.6 C -1928 2054.4 1932.8 2063 1925.3 2064 C -1926.9 2063.3 1928.3 2062.4 1929.6 2061.2 C -[0.4 0.4 0 0] vc -f -S -n -1930.8 2061.6 m -1930.5 2058 1931.6 2054 1930.8 2051.3 C -1930.3 2054.5 1930.9 2058.5 1930.4 2061.9 C -1930.5 2061.2 1931 2062.2 1930.8 2061.6 C -[0.92 0.92 0 0.67] vc -f -S -n -1941.2 2045.1 m -1939.7 2042.6 1937.3 2041.2 1935.4 2039.3 C -1934.2 2040 1933.7 2036.4 1934 2039.3 C -1934.9 2040.1 1936.1 2039.9 1936.8 2040.8 C -1935.3 2044.2 1942.3 2041.7 1939.5 2046 C -1937.1 2048.5 1940.5 2045.6 1941.2 2045.1 C -f -S -n -1910 2045.8 m -1910 2039.4 1910 2033 1910 2026.6 C -1910 2033 1910 2039.4 1910 2045.8 C -f -S -n -1978.8 2022.3 m -1979.1 2021.7 1979.4 2020.4 1978.6 2021.6 C -1978.6 2026.9 1978.6 2033 1978.6 2037.6 C -1979.2 2037 1979.1 2038.2 1979.1 2038.6 C -1978.7 2033.6 1978.9 2026.8 1978.8 2022.3 C -f -S -n -vmrs -2026.1 2041.2 m -2026.1 2034.8 2026.1 2028.3 2026.1 2021.8 C -2026.1 2028.5 2026.3 2035.4 2025.9 2042 C -2024.4 2042.9 2022.9 2044.1 2021.3 2044.8 C -2023.1 2044 2025.1 2042.8 2026.1 2041.2 C -[0.07 0.06 0 0.58] vc -f -0.4 w -2 J -2 M -S -n -2026.4 2021.8 m -2026.3 2028.5 2026.5 2035.4 2026.1 2042 C -2025.6 2042.8 2024.7 2042.7 2024.2 2043.4 C -2024.7 2042.7 2025.5 2042.7 2026.1 2042.2 C -2026.5 2035.5 2026.3 2027.9 2026.4 2021.8 C -[0.4 0.4 0 0] vc -f -S -n -2025.6 2038.4 m -2025.6 2033 2025.6 2027.6 2025.6 2022.3 C -2025.6 2027.6 2025.6 2033 2025.6 2038.4 C -[0.92 0.92 0 0.67] vc -f -S -n -1934 2023.5 m -1934 2024.7 1933.8 2026 1934.2 2027.1 C -1934 2025.5 1934.7 2024.6 1934 2023.5 C -f -S -n -1928.2 2023.5 m -1928 2024.6 1927.4 2023.1 1926.8 2023.2 C -1926.2 2021 1921.4 2019.3 1923.2 2018 C -1922.7 2016.5 1923.2 2019.3 1922.2 2018.2 C -1924.4 2020.4 1926.2 2023.3 1928.9 2024.9 C -1927.9 2024.2 1929.8 2023.5 1928.2 2023.5 C -[0.18 0.18 0 0.78] vc -f -S -n -1934 2019.2 m -1932 2019.6 1930.8 2022.6 1928.7 2021.8 C -1924.5 2016.5 1918.2 2011.8 1914 2006.7 C -1914 2005.7 1914 2004.6 1914 2003.6 C -1913.6 2004.3 1913.9 2005.8 1913.8 2006.9 C -1919 2012.4 1924.1 2016.5 1929.2 2022.3 C -1931 2021.7 1932.2 2019.8 1934 2019.2 C -f -S -n -1928.7 2024.9 m -1926.3 2022.7 1924.1 2020.4 1921.7 2018.2 C -1924.1 2020.4 1926.3 2022.7 1928.7 2024.9 C -[0.65 0.65 0 0.42] vc -f -S -n -1914.3 2006.7 m -1918.7 2011.8 1924.5 2016.4 1928.9 2021.6 C -1924.2 2016.1 1919 2012.1 1914.3 2006.7 C -[0.07 0.06 0 0.58] vc -f -S -n -1924.8 2020.8 m -1921.2 2016.9 1925.6 2022.5 1926 2021.1 C -1924.2 2021 1926.7 2019.6 1924.8 2020.8 C -[0.92 0.92 0 0.67] vc -f -S -n -1934 2018.4 m -1933.2 2014.7 1934.5 2012.3 1933.7 2008.8 C -1934 2007.8 1935 2007.2 1935.6 2006.7 C -1935.3 2007.1 1934.3 2007 1934 2007.6 C -1932.2 2012.3 1937.2 2021 1929.2 2021.8 C -1931.1 2021.4 1932.3 2019.6 1934 2018.4 C -[0.07 0.06 0 0.58] vc -f -S -n -vmrs -1933.5 2018.7 m -1933.5 2015.1 1933.5 2011.7 1933.5 2008.1 C -1933.8 2007.4 1934.3 2006.9 1934.9 2006.7 C -1933.8 2006.1 1934.3 2007.7 1933.2 2007.2 C -1931.9 2012 1936.7 2020.5 1929.2 2021.6 C -1930.7 2020.8 1932.2 2019.9 1933.5 2018.7 C -[0.4 0.4 0 0] vc -f -0.4 w -2 J -2 M -S -n -1934.7 2019.2 m -1934.3 2015.6 1935.4 2011.5 1934.7 2008.8 C -1934.1 2012 1934.7 2016 1934.2 2019.4 C -1934.4 2018.7 1934.8 2019.8 1934.7 2019.2 C -[0.92 0.92 0 0.67] vc -f -S -n -1917.6 2013.6 m -1917.8 2011.1 1916.8 2014.2 1917.2 2012.2 C -1916.3 2012.9 1914.8 2011.8 1914.3 2010.8 C -1914.2 2010.5 1914.4 2010.4 1914.5 2010.3 C -1913.9 2008.8 1913.9 2011.9 1914.3 2012 C -1916.3 2012 1917.6 2013.6 1916.7 2015.6 C -1913.7 2017.4 1919.6 2014.8 1917.6 2013.6 C -f -S -n -1887.2 2015.3 m -1887.2 2008.9 1887.2 2002.5 1887.2 1996.1 C -1887.2 2002.5 1887.2 2008.9 1887.2 2015.3 C -f -S -n -1916.7 2014.4 m -1917 2012.1 1913 2013 1913.8 2010.8 C -1912.1 2009.8 1910.9 2009.4 1910.7 2007.9 C -1910.4 2010.6 1913.4 2010.4 1914 2012.4 C -1914.9 2012.8 1916.6 2012.9 1916.4 2014.4 C -1916.9 2015.1 1914.5 2016.6 1916.2 2015.8 C -1916.4 2015.3 1916.7 2015 1916.7 2014.4 C -[0.65 0.65 0 0.42] vc -f -S -n -1914 2009.3 m -1912.8 2010.9 1909.6 2005.3 1911.9 2009.8 C -1912.3 2009.6 1913.6 2010.2 1914 2009.3 C -[0.92 0.92 0 0.67] vc -f -S -n -1951.2 1998.8 m -1949 1996.4 1951.5 1994 1950.3 1991.8 C -1949.1 1989.1 1954 1982.7 1948.8 1981.2 C -1949.2 1981.5 1951 1982.4 1950.8 1983.6 C -1951.9 1988.6 1947.1 1986.5 1948.1 1990.4 C -1948.5 1990.3 1948.7 1990.7 1948.6 1991.1 C -1949 1992.5 1947.3 1991.9 1948.1 1992.5 C -1947.1 1992.7 1945.7 1993.5 1945.2 1994.7 C -1944.5 1996.8 1947.7 2000.5 1943.8 2001.4 C -1943.4 2002 1943.7 2004 1942.4 2004.5 C -1945.2 2002.2 1948.9 2000.9 1951.2 1998.8 C -f -S -n -1994.9 1993 m -1995.1 1996.5 1994.5 2000.3 1995.4 2003.6 C -1994.5 2000.3 1995.1 1996.5 1994.9 1993 C -f -S -n -1913.8 2003.3 m -1913.8 1996.9 1913.8 1990.5 1913.8 1984.1 C -1913.8 1990.5 1913.8 1996.9 1913.8 2003.3 C -f -S -n -1941.9 1998 m -1940.5 1997.3 1940.7 1999.4 1940.7 2000 C -1942.8 2001.3 1942.6 1998.8 1941.9 1998 C -[0 0 0 0] vc -f -S -n -vmrs -1942.1 1999.2 m -1942.2 1998.9 1941.8 1998.8 1941.6 1998.5 C -1940.4 1998 1940.7 1999.7 1940.7 2000 C -1941.6 2000.3 1942.6 2000.4 1942.1 1999.2 C -[0.92 0.92 0 0.67] vc -f -0.4 w -2 J -2 M -S -n -1940 1997.1 m -1939.8 1996 1939.7 1995.9 1939.2 1995.2 C -1939.1 1995.3 1938.5 1997.9 1937.8 1996.4 C -1938 1997.3 1939.4 1998.6 1940 1997.1 C -f -S -n -1911.2 1995.9 m -1911.2 1991.6 1911.3 1987.2 1911.4 1982.9 C -1911.3 1987.2 1911.2 1991.6 1911.2 1995.9 C -f -S -n -1947.2 1979.1 m -1945.1 1978.8 1944.6 1975.7 1942.4 1975 C -1940.5 1972.6 1942.2 1973.7 1942.4 1975.7 C -1945.8 1975.5 1944.2 1979.8 1947.6 1979.6 C -1948.3 1982.3 1948.5 1980 1947.2 1979.1 C -f -S -n -1939.5 1973.3 m -1940.1 1972.6 1939.8 1974.2 1940.2 1973.1 C -1939.1 1972.8 1938.8 1968.5 1935.9 1969.7 C -1937.4 1969.2 1938.5 1970.6 1939 1971.4 C -1939.2 1972.7 1938.6 1973.9 1939.5 1973.3 C -f -S -n -1975.2 2073.2 m -1975.2 2070.2 1975.2 2067.2 1975.2 2064.3 C -1975.2 2067.2 1975.2 2070.2 1975.2 2073.2 C -[0.18 0.18 0 0.78] vc -f -S -n -1929.9 2065.7 m -1928.1 2065.6 1926 2068.8 1924.1 2066.9 C -1918.1 2060.9 1912.9 2055.7 1907.1 2049.9 C -1906.7 2047.1 1906.9 2043.9 1906.8 2041 C -1906.8 2043.9 1906.8 2046.8 1906.8 2049.6 C -1913.2 2055.5 1918.7 2061.9 1925.1 2067.6 C -1927.1 2067.9 1928.6 2064.4 1930.1 2066.2 C -1929.7 2070.3 1929.9 2074.7 1929.9 2078.9 C -1929.6 2074.4 1930.5 2070.1 1929.9 2065.7 C -[0.07 0.06 0 0.58] vc -f -S -n -1930.1 2061.6 m -1928.1 2062.1 1927 2065.1 1924.8 2064.3 C -1920.7 2058.9 1914.4 2054.3 1910.2 2049.2 C -1910.2 2048.1 1910.2 2047.1 1910.2 2046 C -1909.8 2046.8 1910 2048.3 1910 2049.4 C -1915.1 2054.9 1920.3 2059 1925.3 2064.8 C -1927.1 2064.2 1928.4 2062.3 1930.1 2061.6 C -[0.18 0.18 0 0.78] vc -f -S -n -1932 2049.9 m -1932.3 2050.3 1932 2050.4 1932.8 2050.4 C -1932 2050.4 1932.2 2049.2 1931.3 2049.6 C -1931.4 2050.5 1930.3 2050.4 1930.4 2051.3 C -1931.1 2051.1 1930.7 2049.4 1932 2049.9 C -f -S -n -1938.3 2046 m -1936.3 2046.8 1935.2 2047.2 1934.2 2048.9 C -1935.3 2047.7 1936.8 2046.2 1938.3 2046 C -[0.4 0.4 0 0] vc -f -S -n -vmrs -1938.3 2047 m -1937.9 2046.9 1936.6 2047.1 1936.1 2048 C -1936.5 2047.5 1937.3 2046.7 1938.3 2047 C -[0.18 0.18 0 0.78] vc -f -0.4 w -2 J -2 M -S -n -1910.2 2043.2 m -1910.1 2037.5 1910 2031.8 1910 2026.1 C -1910 2031.8 1910.1 2037.5 1910.2 2043.2 C -f -S -n -1933.5 2032.1 m -1933.7 2035.2 1932.8 2035.8 1933.7 2038.6 C -1933.3 2036.6 1934.6 2018 1933.5 2032.1 C -f -S -n -1907.3 2021.8 m -1906.6 2025.9 1909.4 2032.6 1903.2 2034 C -1902.8 2034.1 1902.4 2033.9 1902 2033.8 C -1897.9 2028.5 1891.6 2023.8 1887.4 2018.7 C -1887.4 2017.7 1887.4 2016.6 1887.4 2015.6 C -1887 2016.3 1887.2 2017.8 1887.2 2018.9 C -1892.3 2024.4 1897.5 2028.5 1902.5 2034.3 C -1904.3 2033.6 1905.7 2032 1907.3 2030.9 C -1907.3 2027.9 1907.3 2024.9 1907.3 2021.8 C -f -S -n -1933.7 2023.2 m -1932 2021.7 1931.1 2024.9 1929.4 2024.9 C -1931.2 2024.7 1932.4 2021.5 1933.7 2023.2 C -f -S -n -1989.2 2024.4 m -1987.4 2023.7 1985.8 2022.2 1985.1 2020.4 C -1984.6 2020.1 1986 2018.9 1985.1 2019.2 C -1985.6 2020.8 1984.1 2019.4 1984.6 2021.1 C -1986.3 2022.3 1988.1 2025.3 1989.2 2024.4 C -f -S -n -1904.4 2031.9 m -1903 2029.7 1905.3 2027.7 1904.2 2025.9 C -1904.5 2025 1903.7 2023 1904 2021.3 C -1904 2022.3 1903.2 2022 1902.5 2022 C -1901.3 2022.3 1902.2 2020.1 1901.6 2019.6 C -1902.5 2019.8 1902.6 2018.3 1903.5 2018.9 C -1903.7 2021.8 1905.6 2016.8 1905.6 2020.6 C -1905.9 2020 1906.3 2020.8 1906.1 2021.1 C -1905.8 2022.7 1906.7 2020.4 1906.4 2019.9 C -1906.4 2018.5 1908.2 2017.8 1906.8 2016.5 C -1906.9 2015.7 1907.7 2017.1 1907.1 2016.3 C -1908.5 2015.8 1910.3 2015.1 1911.6 2016 C -1912.2 2016.2 1911.9 2018 1911.6 2018 C -1914.5 2017.1 1910.4 2013.6 1913.3 2013.4 C -1912.4 2011.3 1910.5 2011.8 1909.5 2010 C -1910 2010.5 1909 2010.8 1908.8 2011.2 C -1907.5 2009.9 1906.1 2011.7 1904.9 2011.5 C -1904.7 2010.9 1904.3 2010.5 1904.4 2009.8 C -1905 2010.2 1904.6 2008.6 1905.4 2008.1 C -1906.6 2007.5 1907.7 2008.4 1908.5 2007.4 C -1908.9 2008.5 1909.7 2008.1 1909 2007.2 C -1908.1 2006.5 1905.9 2007.3 1905.4 2007.4 C -1903.9 2007.3 1905.2 2008.5 1904.2 2008.4 C -1904.6 2009.9 1902.8 2010.3 1902.3 2010.5 C -1901.5 2009.9 1900.4 2010 1899.4 2010 C -1898.6 2011.2 1898.2 2013.4 1896.5 2013.4 C -1896 2012.9 1894.4 2012.9 1893.6 2012.9 C -1893.1 2013.9 1892.9 2015.5 1891.5 2016 C -1890.3 2016.1 1889.2 2014 1888.6 2015.8 C -1890 2016 1891 2016.9 1892.9 2016.5 C -1894.1 2017.2 1892.8 2018.3 1893.2 2018.9 C -1892.6 2018.9 1891.1 2019.8 1890.5 2020.6 C -1891.1 2023.6 1893.2 2019.8 1893.9 2022.5 C -1894.1 2023.3 1892.7 2023.6 1893.9 2024 C -1894.2 2024.3 1897.4 2023.8 1896.5 2026.1 C -1896 2025.6 1897.4 2028.1 1897.5 2027.1 C -1898.4 2027.4 1899.3 2027 1899.6 2028.5 C -1899.5 2028.6 1899.4 2028.8 1899.2 2028.8 C -1899.3 2029.2 1899.6 2029.8 1900.1 2030.2 C -1900.4 2029.6 1901 2030 1901.8 2030.2 C -1903.1 2032.1 1900.4 2031.5 1902.8 2033.1 C -1903.3 2032.7 1904.5 2032 1904.4 2031.9 C -[0.21 0.21 0 0] vc -f -S -n -1909.2 2019.4 m -1908.8 2020.3 1910.2 2019.8 1909.2 2019.2 C -1908.3 2019.3 1907.6 2020.2 1907.6 2021.3 C -1908.5 2021 1907.6 2019 1909.2 2019.4 C -[0.18 0.18 0 0.78] vc -f -S -n -1915.5 2015.6 m -1913.5 2016.3 1912.4 2016.8 1911.4 2018.4 C -1912.5 2017.2 1914 2015.7 1915.5 2015.6 C -[0.4 0.4 0 0] vc -f -S -n -1915.5 2016.5 m -1915.1 2016.4 1913.8 2016.6 1913.3 2017.5 C -1913.7 2017 1914.5 2016.2 1915.5 2016.5 C -[0.18 0.18 0 0.78] vc -f -S -n -vmrs -1887.4 2012.7 m -1887.3 2007 1887.2 2001.3 1887.2 1995.6 C -1887.2 2001.3 1887.3 2007 1887.4 2012.7 C -[0.18 0.18 0 0.78] vc -f -0.4 w -2 J -2 M -S -n -1935.9 2007.4 m -1936.2 2007.8 1935.8 2007.9 1936.6 2007.9 C -1935.9 2007.9 1936.1 2006.7 1935.2 2007.2 C -1935.2 2008.1 1934.1 2007.9 1934.2 2008.8 C -1935 2008.7 1934.6 2006.9 1935.9 2007.4 C -f -S -n -1942.1 2003.6 m -1940.1 2004.3 1939.1 2004.8 1938 2006.4 C -1939.1 2005.2 1940.6 2003.7 1942.1 2003.6 C -[0.4 0.4 0 0] vc -f -S -n -1942.1 2004.5 m -1941.8 2004.4 1940.4 2004.6 1940 2005.5 C -1940.4 2005 1941.2 2004.2 1942.1 2004.5 C -[0.18 0.18 0 0.78] vc -f -S -n -1914 2000.7 m -1914 1995 1913.9 1989.3 1913.8 1983.6 C -1913.9 1989.3 1914 1995 1914 2000.7 C -f -S -n -1941.6 1998.3 m -1943.4 2001.9 1942.4 1996 1940.9 1998.3 C -1941.2 1998.3 1941.4 1998.3 1941.6 1998.3 C -f -S -n -1954.8 1989.9 m -1953.9 1989.6 1954.7 1991.6 1953.9 1991.1 C -1954.5 1993.1 1953.6 1998 1954.6 1993.2 C -1954 1992.2 1954.7 1990.7 1954.8 1989.9 C -f -S -n -1947.6 1992.5 m -1946.2 1993.5 1944.9 1993 1944.8 1994.7 C -1945.5 1994 1947 1992.2 1947.6 1992.5 C -f -S -n -1910.7 1982.2 m -1910.3 1981.8 1909.7 1982 1909.2 1982 C -1909.7 1982 1910.3 1981.9 1910.7 1982.2 C -1911 1987.1 1910 1992.6 1910.7 1997.3 C -1910.7 1992.3 1910.7 1987.2 1910.7 1982.2 C -[0.65 0.65 0 0.42] vc -f -S -n -1910.9 1992.8 m -1910.9 1991.3 1910.9 1989.7 1910.9 1988.2 C -1910.9 1989.7 1910.9 1991.3 1910.9 1992.8 C -[0.18 0.18 0 0.78] vc -f -S -n -vmrs -1953.6 1983.6 m -1954.1 1985.3 1953.2 1988.6 1954.8 1989.4 C -1954.1 1987.9 1954.4 1985.4 1953.6 1983.6 C -[0.18 0.18 0 0.78] vc -f -0.4 w -2 J -2 M -S -n -1910.7 1982 m -1911.6 1982.9 1911 1984.4 1911.2 1985.6 C -1911 1984.4 1911.6 1982.9 1910.7 1982 C -f -S -n -1947.2 1979.6 m -1947.5 1980.6 1948.3 1980.6 1947.4 1979.6 C -1946.2 1979.4 1945.7 1978.8 1947.2 1979.6 C -f -S -n -1930.4 2061.4 m -1930.4 2058 1930.4 2053.5 1930.4 2051.1 C -1930.7 2054.6 1929.8 2057.4 1930.1 2061.2 C -1929.5 2061.9 1929.7 2061.2 1930.4 2061.4 C -[0.65 0.65 0 0.42] vc -f -S -n -1939.5 2044.8 m -1940 2041.5 1935.2 2044.3 1936.4 2040.8 C -1934.9 2040.9 1934.1 2039.7 1933.5 2038.6 C -1933.3 2035.4 1933.2 2040 1934 2040.3 C -1936.2 2040.6 1936.3 2043.6 1938.5 2043.4 C -1939.7 2044.2 1939.4 2045.6 1938.3 2046.5 C -1939.1 2046.6 1939.6 2045.6 1939.5 2044.8 C -f -S -n -1910.4 2045.3 m -1910.4 2039.5 1910.4 2033.6 1910.4 2027.8 C -1910.4 2033.6 1910.4 2039.5 1910.4 2045.3 C -f -S -n -1906.8 2030.9 m -1907.6 2026.8 1905 2020.8 1909 2018.7 C -1906.5 2018.9 1906.8 2022.4 1906.8 2024.7 C -1906.4 2028.2 1907.9 2032 1903 2033.8 C -1902.2 2034 1903.8 2033.4 1904.2 2033.1 C -1905.1 2032.4 1905.9 2031.5 1906.8 2030.9 C -[0.07 0.06 0 0.58] vc -f -S -n -1907.1 2030.7 m -1907.1 2028.8 1907.1 2027 1907.1 2025.2 C -1907.1 2027 1907.1 2028.8 1907.1 2030.7 C -[0.65 0.65 0 0.42] vc -f -S -n -1932 2023.2 m -1932.2 2023.6 1931.7 2023.7 1931.6 2024 C -1932 2023.7 1932.3 2022.8 1933 2023 C -1933.9 2024.3 1933.3 2026.2 1933.5 2027.8 C -1933.5 2026.4 1934.9 2022.2 1932 2023.2 C -f -S -n -2026.1 2021.6 m -2026.1 2020.8 2026.1 2019.9 2026.1 2019.2 C -2026.1 2019.9 2026.1 2020.8 2026.1 2021.6 C -f -S -n -vmrs -1934.2 2018.9 m -1934.2 2015.5 1934.2 2011 1934.2 2008.6 C -1934.5 2012.1 1933.7 2014.9 1934 2018.7 C -1933.4 2019.5 1933.5 2018.7 1934.2 2018.9 C -[0.65 0.65 0 0.42] vc -f -0.4 w -2 J -2 M -S -n -1887.6 2014.8 m -1887.6 2009 1887.6 2003.1 1887.6 1997.3 C -1887.6 2003.1 1887.6 2009 1887.6 2014.8 C -f -S -n -1914.3 2002.8 m -1914.3 1997 1914.3 1991.1 1914.3 1985.3 C -1914.3 1991.1 1914.3 1997 1914.3 2002.8 C -f -S -n -1995.4 1992.3 m -1995.4 1991.5 1995.4 1990.7 1995.4 1989.9 C -1995.4 1990.7 1995.4 1991.5 1995.4 1992.3 C -f -S -n -1896 1988.4 m -1896.9 1988 1897.8 1987.7 1898.7 1987.2 C -1897.8 1987.7 1896.9 1988 1896 1988.4 C -f -S -n -1899.4 1986.8 m -1900.4 1986.3 1901.3 1985.8 1902.3 1985.3 C -1901.3 1985.8 1900.4 1986.3 1899.4 1986.8 C -f -S -n -1902.8 1985.1 m -1905.2 1984 1905.2 1984 1902.8 1985.1 C -f -S -n -1949.1 1983.4 m -1950.2 1984.4 1947.8 1984.6 1949.3 1985.1 C -1949.5 1984.4 1949.6 1984.1 1949.1 1983.4 C -[0.07 0.06 0 0.58] vc -f -S -n -1906.1 1983.4 m -1908.6 1982 1908.6 1982 1906.1 1983.4 C -[0.65 0.65 0 0.42] vc -f -S -n -1922.7 1976.4 m -1923.6 1976 1924.4 1975.7 1925.3 1975.2 C -1924.4 1975.7 1923.6 1976 1922.7 1976.4 C -f -S -n -vmrs -1926 1974.8 m -1927 1974.3 1928 1973.8 1928.9 1973.3 C -1928 1973.8 1927 1974.3 1926 1974.8 C -[0.65 0.65 0 0.42] vc -f -0.4 w -2 J -2 M -S -n -1929.4 1973.1 m -1931.9 1972 1931.9 1972 1929.4 1973.1 C -f -S -n -1932.8 1971.4 m -1935.3 1970 1935.3 1970 1932.8 1971.4 C -f -S -n -1949.6 2097.2 m -1951.1 2096.4 1952.6 2095.5 1954.1 2094.8 C -1952.6 2095.5 1951.1 2096.4 1949.6 2097.2 C -[0.07 0.06 0 0.58] vc -f -S -n -1955.1 2094.3 m -1956.7 2093.5 1958.3 2092.7 1959.9 2091.9 C -1958.3 2092.7 1956.7 2093.5 1955.1 2094.3 C -f -S -n -1960.4 2091.6 m -1961.3 2091.2 1962.1 2090.9 1963 2090.4 C -1962.1 2090.9 1961.3 2091.2 1960.4 2091.6 C -f -S -n -1963.5 2090.2 m -1964.4 2089.7 1965.2 2089.2 1966.1 2088.8 C -1965.2 2089.2 1964.4 2089.7 1963.5 2090.2 C -f -S -n -1966.6 2088.5 m -1969.5 2087.1 1972.4 2085.8 1975.2 2084.4 C -1972.4 2085.8 1969.5 2087.1 1966.6 2088.5 C -f -S -n -1965.2 2086.1 m -1965.9 2085.7 1966.8 2085.3 1967.6 2084.9 C -1966.8 2085.3 1965.9 2085.7 1965.2 2086.1 C -f -S -n -1968.3 2084.7 m -1969.2 2084.3 1970 2083.9 1970.9 2083.5 C -1970 2083.9 1969.2 2084.3 1968.3 2084.7 C -f -S -n -vmrs -1984.1 2084 m -1985.6 2083.2 1987.2 2082.3 1988.7 2081.6 C -1987.2 2082.3 1985.6 2083.2 1984.1 2084 C -[0.07 0.06 0 0.58] vc -f -0.4 w -2 J -2 M -S -n -1976 2078.7 m -1978.1 2080.1 1980 2082 1982 2083.7 C -1980 2081.9 1977.9 2080.3 1976 2078.2 C -1975.5 2079.9 1975.8 2081.9 1975.7 2083.7 C -1975.8 2082 1975.5 2080.2 1976 2078.7 C -f -S -n -1989.6 2081.1 m -1991.3 2080.3 1992.8 2079.5 1994.4 2078.7 C -1992.8 2079.5 1991.3 2080.3 1989.6 2081.1 C -f -S -n -1933.2 2074.6 m -1932.4 2076.2 1932.8 2077.5 1933 2078.7 C -1933 2077.6 1932.9 2074.8 1933.2 2074.6 C -f -S -n -1994.9 2078.4 m -1995.8 2078 1996.7 2077.7 1997.6 2077.2 C -1996.7 2077.7 1995.8 2078 1994.9 2078.4 C -f -S -n -1998 2077 m -1998.9 2076.5 1999.8 2076 2000.7 2075.6 C -1999.8 2076 1998.9 2076.5 1998 2077 C -f -S -n -2001.2 2075.3 m -2004 2073.9 2006.9 2072.6 2009.8 2071.2 C -2006.9 2072.6 2004 2073.9 2001.2 2075.3 C -f -S -n -1980.5 2060.7 m -1979.9 2060.7 1976.7 2062.8 1975.7 2064.5 C -1975.7 2067.5 1975.7 2070.5 1975.7 2073.4 C -1976.3 2068.7 1973.9 2061.6 1980.5 2060.7 C -f -S -n -1999.7 2072.9 m -2000.5 2072.5 2001.3 2072.1 2002.1 2071.7 C -2001.3 2072.1 2000.5 2072.5 1999.7 2072.9 C -f -S -n -2002.8 2071.5 m -2003.7 2071.1 2004.6 2070.7 2005.5 2070.3 C -2004.6 2070.7 2003.7 2071.1 2002.8 2071.5 C -f -S -n -vmrs -2015.1 2047.5 m -2014.4 2047.5 2011.2 2049.6 2010.3 2051.3 C -2010.3 2057.7 2010.3 2064.1 2010.3 2070.5 C -2010.3 2063.9 2010.1 2057.1 2010.5 2050.6 C -2012 2049.3 2013.5 2048.3 2015.1 2047.5 C -[0.07 0.06 0 0.58] vc -f -0.4 w -2 J -2 M -S -n -1910.4 2049.2 m -1914.8 2054.3 1920.7 2058.9 1925.1 2064 C -1920.4 2058.6 1915.1 2054.6 1910.4 2049.2 C -f -S -n -1988.2 2057.3 m -1989.1 2056.8 1989.9 2056.2 1990.8 2055.6 C -1989.9 2056.2 1989.1 2056.8 1988.2 2057.3 C -f -S -n -1991.6 2051.3 m -1991.6 2046.3 1991.6 2041.2 1991.6 2036.2 C -1991.6 2041.2 1991.6 2046.3 1991.6 2051.3 C -f -S -n -1935.6 2047.5 m -1932.9 2051.7 1939.7 2043.8 1935.6 2047.5 C -f -S -n -1938.8 2043.9 m -1938.1 2043.3 1938.2 2043.7 1937.3 2043.4 C -1938.7 2043 1938.2 2044.9 1939 2045.3 C -1938.2 2045.3 1938.7 2046.6 1937.8 2046.5 C -1939.1 2046.2 1939.1 2044.5 1938.8 2043.9 C -f -S -n -1972.4 2045.6 m -1973.4 2045 1974.5 2044.4 1975.5 2043.9 C -1974.5 2044.4 1973.4 2045 1972.4 2045.6 C -f -S -n -1969 2043.6 m -1969.8 2043.2 1970.6 2042.9 1971.4 2042.4 C -1970.6 2042.9 1969.8 2043.2 1969 2043.6 C -f -S -n -1972.1 2042.2 m -1973 2041.8 1973.9 2041.4 1974.8 2041 C -1973.9 2041.4 1973 2041.8 1972.1 2042.2 C -f -S -n -1906.6 2035 m -1905 2034.7 1904.8 2036.6 1903.5 2036.9 C -1904.9 2037 1905.8 2033.4 1907.1 2035.7 C -1907.1 2037.2 1907.1 2038.6 1907.1 2040 C -1906.9 2038.4 1907.5 2036.4 1906.6 2035 C -f -S -n -vmrs -1937.1 2032.1 m -1936.2 2033.7 1936.6 2035 1936.8 2036.2 C -1936.8 2035.1 1936.8 2032.4 1937.1 2032.1 C -[0.07 0.06 0 0.58] vc -f -0.4 w -2 J -2 M -S -n -1887.6 2018.7 m -1892 2023.8 1897.9 2028.4 1902.3 2033.6 C -1897.6 2028.1 1892.3 2024.1 1887.6 2018.7 C -f -S -n -1999.7 2031.4 m -1998.7 2030.3 1997.6 2029.2 1996.6 2028 C -1997.6 2029.2 1998.7 2030.3 1999.7 2031.4 C -f -S -n -1912.8 2017 m -1910.6 2021.1 1913.6 2015.3 1914.5 2016 C -1914 2016.3 1913.4 2016.7 1912.8 2017 C -f -S -n -1939.5 2005 m -1936.7 2009.2 1943.6 2001.3 1939.5 2005 C -f -S -n -1942.6 2001.4 m -1941.9 2000.8 1942 2001.2 1941.2 2000.9 C -1942.5 2000.6 1942.1 2002.4 1942.8 2002.8 C -1942 2002.8 1942.5 2004.1 1941.6 2004 C -1943 2003.7 1942.9 2002.1 1942.6 2001.4 C -f -S -n -2006.2 2000.7 m -2005.4 2001.5 2004 2002.8 2004 2002.8 C -2004.5 2002.4 2005.5 2001.4 2006.2 2000.7 C -f -S -n -1998.5 2001.6 m -1997.7 2002 1996.8 2002.4 1995.9 2002.6 C -1995.5 1999.3 1995.7 1995.7 1995.6 1992.3 C -1995.6 1995.7 1995.6 1999.2 1995.6 2002.6 C -1996.6 2002.4 1997.7 2002.2 1998.5 2001.6 C -[0.4 0.4 0 0] vc -f -S -n -1996.1 2002.8 m -1995.9 2002.8 1995.8 2002.8 1995.6 2002.8 C -1995.2 1999.5 1995.5 1995.9 1995.4 1992.5 C -1995.4 1995.9 1995.4 1999.4 1995.4 2002.8 C -1996.4 2003.1 1998.2 2001.6 1996.1 2002.8 C -[0.07 0.06 0 0.58] vc -f -S -n -1969 2002.1 m -1968 2001 1966.9 1999.9 1965.9 1998.8 C -1966.9 1999.9 1968 2001 1969 2002.1 C -f -S -n -vmrs -2000 2001.2 m -2002.1 2000 2004.1 1998.9 2006.2 1997.8 C -2004.1 1998.9 2002.1 2000 2000 2001.2 C -[0.07 0.06 0 0.58] vc -f -0.4 w -2 J -2 M -S -n -1895.8 1984.8 m -1898.3 1983.6 1900.8 1982.3 1903.2 1981 C -1900.8 1982.3 1898.3 1983.6 1895.8 1984.8 C -f -S -n -1905.2 1980.3 m -1906.4 1979.9 1907.6 1979.5 1908.8 1979.1 C -1907.6 1979.5 1906.4 1979.9 1905.2 1980.3 C -f -S -n -1964.7 1977.4 m -1963.8 1977.5 1962.5 1980.2 1960.8 1980 C -1962.5 1980.2 1963.3 1978 1964.7 1977.4 C -f -S -n -1952 1979.6 m -1955.2 1979.2 1955.2 1979.2 1952 1979.6 C -f -S -n -1937.8 1966.4 m -1941.2 1969.5 1946.1 1976.4 1951.5 1979.3 C -1946.1 1976.7 1942.8 1970.4 1937.8 1966.4 C -f -S -n -1911.9 1978.6 m -1914.3 1977.4 1916.7 1976.2 1919.1 1975 C -1916.7 1976.2 1914.3 1977.4 1911.9 1978.6 C -f -S -n -1975.5 1971.4 m -1974.6 1972.2 1973.3 1973.6 1973.3 1973.6 C -1973.7 1973.1 1974.8 1972.1 1975.5 1971.4 C -f -S -n -1922.4 1972.8 m -1924.9 1971.6 1927.4 1970.3 1929.9 1969 C -1927.4 1970.3 1924.9 1971.6 1922.4 1972.8 C -f -S -n -1969.2 1971.9 m -1971.1 1970.9 1972.9 1969.8 1974.8 1968.8 C -1972.9 1969.8 1971.1 1970.9 1969.2 1971.9 C -f -S -n -vmrs -1931.8 1968.3 m -1933 1967.9 1934.2 1967.5 1935.4 1967.1 C -1934.2 1967.5 1933 1967.9 1931.8 1968.3 C -[0.07 0.06 0 0.58] vc -f -0.4 w -2 J -2 M -S -n -1940.7 2072.4 m -1941.5 2072.4 1942.3 2072.3 1943.1 2072.2 C -1942.3 2072.3 1941.5 2072.4 1940.7 2072.4 C -[0 0 0 0.18] vc -f -S -n -1948.6 2069.3 m -1947 2069.5 1945.7 2068.9 1944.8 2069.8 C -1945.9 2068.5 1948.4 2070.2 1948.6 2069.3 C -f -S -n -1954.6 2066.4 m -1954.7 2067.9 1955.6 2067.3 1955.6 2068.8 C -1955.4 2067.8 1956 2066.6 1954.6 2066.4 C -f -S -n -1929.2 2061.2 m -1927.8 2062.1 1926.3 2064.1 1924.8 2063.3 C -1926.3 2064.6 1928 2062 1929.2 2061.2 C -f -S -n -1924.4 2067.4 m -1918.5 2061.6 1912.7 2055.9 1906.8 2050.1 C -1912.7 2055.9 1918.5 2061.6 1924.4 2067.4 C -[0.4 0.4 0 0] vc -f -S -n -1924.6 2062.8 m -1923.9 2062.1 1923.2 2061.2 1922.4 2060.4 C -1923.2 2061.2 1923.9 2062.1 1924.6 2062.8 C -[0 0 0 0.18] vc -f -S -n -1919.3 2057.3 m -1917.5 2055.6 1915.7 2053.8 1913.8 2052 C -1915.7 2053.8 1917.5 2055.6 1919.3 2057.3 C -f -S -n -1929.2 2055.2 m -1929.2 2054.2 1929.2 2053.2 1929.2 2052.3 C -1929.2 2053.2 1929.2 2054.2 1929.2 2055.2 C -f -S -n -1926.3 2049.6 m -1925.4 2049 1925.4 2050.5 1924.4 2050.4 C -1925.3 2051.3 1924.5 2051.9 1925.6 2052.5 C -1926.9 2052.6 1926 2050.6 1926.3 2049.6 C -f -S -n -vmrs -1911.2 2046.8 m -1910.1 2048.9 1911.9 2050.1 1913.1 2051.3 C -1912.1 2049.9 1910.6 2048.8 1911.2 2046.8 C -[0 0 0 0.18] vc -f -0.4 w -2 J -2 M -S -n -1934 2048.7 m -1932.6 2048.7 1930.1 2047.7 1929.6 2049.4 C -1930.9 2048.6 1933.3 2049 1934 2048.7 C -f -S -n -1980 2048.4 m -1979.5 2046.8 1976.3 2047.9 1977.2 2045.6 C -1976.8 2045.1 1976.1 2044.7 1975.2 2044.8 C -1973.7 2046 1976.3 2046.4 1976.7 2047.5 C -1977.8 2047.2 1978.2 2050 1979.6 2049.2 C -1980 2049 1979.6 2048.6 1980 2048.4 C -f -S -n -1938.3 2045.6 m -1938.2 2044.4 1936.8 2043.8 1935.9 2043.4 C -1936.4 2044.4 1939.1 2044.3 1937.6 2045.8 C -1937 2046.1 1935.9 2046.1 1935.9 2046.8 C -1936.7 2046.3 1937.8 2046.2 1938.3 2045.6 C -f -S -n -1932.5 2040 m -1932.8 2038.1 1932 2038.9 1932.3 2040.3 C -1933.1 2040.3 1932.7 2041.7 1933.7 2041.5 C -1933.1 2041 1932.9 2040.5 1932.5 2040 C -f -S -n -2014.6 2035.2 m -2014.1 2033.6 2010.9 2034.7 2011.7 2032.4 C -2011.3 2031.9 2009.4 2030.7 2009.3 2032.1 C -2009.5 2033.7 2012.9 2033.8 2012.4 2035.7 C -2013 2036.4 2014.2 2036.5 2014.6 2035.2 C -f -S -n -1906.4 2030.7 m -1905 2031.6 1903.5 2033.6 1902 2032.8 C -1903.4 2034 1905.6 2031.4 1906.4 2030.7 C -f -S -n -1901.8 2037.2 m -1899.5 2034.8 1897.2 2032.5 1894.8 2030.2 C -1897.2 2032.5 1899.5 2034.8 1901.8 2037.2 C -[0.4 0.4 0 0] vc -f -S -n -1901.8 2032.4 m -1901.1 2031.6 1900.4 2030.7 1899.6 2030 C -1900.4 2030.7 1901.1 2031.6 1901.8 2032.4 C -[0 0 0 0.18] vc -f -S -n -1944.5 2030 m -1945.3 2029.9 1946.1 2029.8 1946.9 2029.7 C -1946.1 2029.8 1945.3 2029.9 1944.5 2030 C -f -S -n -vmrs -1997.8 2027.8 m -1997.7 2027.9 1997.6 2028.1 1997.3 2028 C -1997.4 2029.1 1998.5 2029.5 1999.2 2030 C -2000.1 2029.5 1998.9 2028 1997.8 2027.8 C -[0 0 0 0.18] vc -f -0.4 w -2 J -2 M -S -n -1906.4 2029.2 m -1906.4 2026.6 1906.4 2024 1906.4 2021.3 C -1906.4 2024 1906.4 2026.6 1906.4 2029.2 C -f -S -n -2006.2 2025.9 m -2006 2025.9 2005.8 2025.8 2005.7 2025.6 C -2005.7 2025.5 2005.7 2025.3 2005.7 2025.2 C -2004.6 2025.8 2002.7 2024.7 2001.9 2026.1 C -2001.9 2027.9 2007.8 2029.2 2006.2 2025.9 C -[0 0 0 0] vc -f -S -n -1952.4 2026.8 m -1950.9 2027 1949.6 2026.4 1948.6 2027.3 C -1949.7 2026.1 1952.2 2027.7 1952.4 2026.8 C -[0 0 0 0.18] vc -f -S -n -1896.5 2026.8 m -1894.7 2025.1 1892.9 2023.3 1891 2021.6 C -1892.9 2023.3 1894.7 2025.1 1896.5 2026.8 C -f -S -n -1958.4 2024 m -1958.5 2025.5 1959.4 2024.8 1959.4 2026.4 C -1959.3 2025.3 1959.8 2024.1 1958.4 2024 C -f -S -n -1903.5 2019.2 m -1902.6 2018.6 1902.6 2020 1901.6 2019.9 C -1902.5 2020.8 1901.7 2021.4 1902.8 2022 C -1904.1 2022.2 1903.2 2020.1 1903.5 2019.2 C -f -S -n -1933 2018.7 m -1931.7 2019.6 1930.1 2021.6 1928.7 2020.8 C -1930.1 2022.1 1931.8 2019.5 1933 2018.7 C -f -S -n -1888.4 2016.3 m -1887.3 2018.4 1889.1 2019.6 1890.3 2020.8 C -1889.3 2019.5 1887.8 2018.3 1888.4 2016.3 C -f -S -n -1928.4 2020.4 m -1927.7 2019.6 1927 2018.7 1926.3 2018 C -1927 2018.7 1927.7 2019.6 1928.4 2020.4 C -f -S -n -vmrs -1911.2 2018.2 m -1909.8 2018.3 1907.3 2017.2 1906.8 2018.9 C -1908.1 2018.1 1910.5 2018.6 1911.2 2018.2 C -[0 0 0 0.18] vc -f -0.4 w -2 J -2 M -S -n -1915.5 2015.1 m -1915.4 2013.9 1914 2013.3 1913.1 2012.9 C -1913.6 2013.9 1916.3 2013.8 1914.8 2015.3 C -1914.2 2015.6 1913.1 2015.6 1913.1 2016.3 C -1913.9 2015.9 1915 2015.7 1915.5 2015.1 C -f -S -n -1923.2 2014.8 m -1921.3 2013.1 1919.5 2011.3 1917.6 2009.6 C -1919.5 2011.3 1921.3 2013.1 1923.2 2014.8 C -f -S -n -1933 2012.7 m -1933 2011.7 1933 2010.8 1933 2009.8 C -1933 2010.8 1933 2011.7 1933 2012.7 C -f -S -n -1909.7 2008.1 m -1908.9 2009.2 1910.1 2009.9 1910.4 2011 C -1911.1 2010.7 1908.9 2009.7 1909.7 2008.1 C -f -S -n -1930.1 2007.2 m -1929.2 2006.6 1929.2 2008 1928.2 2007.9 C -1929.1 2008.8 1928.4 2009.4 1929.4 2010 C -1930.7 2010.2 1929.9 2008.1 1930.1 2007.2 C -f -S -n -1915 2004.3 m -1914 2006.4 1915.7 2007.6 1916.9 2008.8 C -1915.9 2007.5 1914.4 2006.3 1915 2004.3 C -f -S -n -1937.8 2006.2 m -1936.4 2006.3 1934 2005.2 1933.5 2006.9 C -1934.7 2006.1 1937.1 2006.6 1937.8 2006.2 C -f -S -n -1983.9 2006 m -1983.3 2004.3 1980.2 2005.4 1981 2003.1 C -1980.6 2002.7 1978.7 2001.5 1978.6 2002.8 C -1978.8 2004.4 1982.1 2004.5 1981.7 2006.4 C -1982.3 2007.2 1983.5 2007.2 1983.9 2006 C -f -S -n -1942.1 2003.1 m -1942 2001.9 1940.6 2001.3 1939.7 2000.9 C -1940.2 2001.9 1943 2001.8 1941.4 2003.3 C -1940.9 2003.6 1939.7 2003.6 1939.7 2004.3 C -1940.5 2003.9 1941.6 2003.7 1942.1 2003.1 C -f -S -n -vmrs -1967.1 1998.5 m -1967 1998.6 1966.8 1998.8 1966.6 1998.8 C -1966.7 1999.8 1967.8 2000.2 1968.5 2000.7 C -1969.4 2000.2 1968.2 1998.8 1967.1 1998.5 C -[0 0 0 0.18] vc -f -0.4 w -2 J -2 M -S -n -1936.4 1997.6 m -1936.7 1995.6 1935.8 1996.4 1936.1 1997.8 C -1936.9 1997.9 1936.5 1999.2 1937.6 1999 C -1937 1998.5 1936.8 1998 1936.4 1997.6 C -f -S -n -1975.5 1996.6 m -1975.2 1996.7 1975.1 1996.5 1975 1996.4 C -1975 1996.2 1975 1996.1 1975 1995.9 C -1973.9 1996.5 1972 1995.5 1971.2 1996.8 C -1971.2 1998.6 1977 1999.9 1975.5 1996.6 C -[0 0 0 0] vc -f -S -n -1949.3 2097.4 m -1950.3 2096.9 1951.2 2096.4 1952.2 2096 C -1951.2 2096.4 1950.3 2096.9 1949.3 2097.4 C -[0.4 0.4 0 0] vc -f -S -n -1960.8 2091.6 m -1961.7 2091.2 1962.6 2090.9 1963.5 2090.4 C -1962.6 2090.9 1961.7 2091.2 1960.8 2091.6 C -f -S -n -1964.4 2090 m -1965.7 2089.2 1967 2088.5 1968.3 2087.8 C -1967 2088.5 1965.7 2089.2 1964.4 2090 C -f -S -n -1976 2083.7 m -1976.3 2082.3 1975.2 2079.1 1976.9 2079.4 C -1978.8 2080.7 1980.3 2082.9 1982.2 2084.2 C -1980.6 2083.1 1978.2 2080.2 1976 2078.9 C -1975.6 2081.2 1977 2084.9 1973.8 2085.4 C -1972.2 2086.1 1970.7 2087 1969 2087.6 C -1971.4 2086.5 1974.1 2085.6 1976 2083.7 C -f -S -n -1983.9 2084.2 m -1984.8 2083.7 1985.8 2083.2 1986.8 2082.8 C -1985.8 2083.2 1984.8 2083.7 1983.9 2084.2 C -f -S -n -1995.4 2078.4 m -1996.3 2078 1997.1 2077.7 1998 2077.2 C -1997.1 2077.7 1996.3 2078 1995.4 2078.4 C -f -S -n -1999 2076.8 m -2000.3 2076 2001.6 2075.3 2002.8 2074.6 C -2001.6 2075.3 2000.3 2076 1999 2076.8 C -f -S -n -vmrs -1929.6 2065.7 m -1930.1 2065.6 1929.8 2068.6 1929.9 2070 C -1929.8 2068.6 1930.1 2067 1929.6 2065.7 C -[0.4 0.4 0 0] vc -f -0.4 w -2 J -2 M -S -n -1906.6 2049.4 m -1906.6 2046.7 1906.6 2043.9 1906.6 2041.2 C -1906.6 2043.9 1906.6 2046.7 1906.6 2049.4 C -f -S -n -2016 2047.5 m -2014.8 2048 2013.5 2048.3 2012.4 2049.4 C -2013.5 2048.3 2014.8 2048 2016 2047.5 C -f -S -n -2016.5 2047.2 m -2017.3 2046.9 2018.1 2046.6 2018.9 2046.3 C -2018.1 2046.6 2017.3 2046.9 2016.5 2047.2 C -f -S -n -1912.4 2028.5 m -1911.8 2032.4 1912.4 2037.2 1911.9 2041.2 C -1911.5 2037.2 1911.7 2032.9 1911.6 2028.8 C -1911.6 2033.5 1911.6 2038.9 1911.6 2042.9 C -1912.5 2042.2 1911.6 2043.9 1912.6 2043.6 C -1912.9 2039.3 1913.1 2033.3 1912.4 2028.5 C -[0.21 0.21 0 0] vc -f -S -n -1906.8 2040.8 m -1906.8 2039 1906.8 2037.2 1906.8 2035.5 C -1906.8 2037.2 1906.8 2039 1906.8 2040.8 C -[0.4 0.4 0 0] vc -f -S -n -1905.9 2035.2 m -1904.9 2036.4 1903.7 2037.2 1902.3 2037.4 C -1903.7 2037.2 1904.9 2036.4 1905.9 2035.2 C -f -S -n -1906.1 2031.2 m -1907 2031.1 1906.4 2028 1906.6 2030.7 C -1905.5 2032.1 1904 2032.8 1902.5 2033.6 C -1903.9 2033.2 1905 2032.1 1906.1 2031.2 C -f -S -n -1908.3 2018.7 m -1905.2 2018.6 1907.1 2023.2 1906.6 2025.4 C -1906.8 2023 1905.9 2019.5 1908.3 2018.7 C -f -S -n -1889.6 1998 m -1889 2001.9 1889.6 2006.7 1889.1 2010.8 C -1888.7 2006.7 1888.9 2002.4 1888.8 1998.3 C -1888.8 2003 1888.8 2008.4 1888.8 2012.4 C -1889.7 2011.7 1888.8 2013.4 1889.8 2013.2 C -1890.1 2008.8 1890.3 2002.8 1889.6 1998 C -[0.21 0.21 0 0] vc -f -S -n -vmrs -1999 2001.4 m -2001 2000.3 2003 1999.2 2005 1998 C -2003 1999.2 2001 2000.3 1999 2001.4 C -[0.4 0.4 0 0] vc -f -0.4 w -2 J -2 M -S -n -1916.2 1986 m -1915.7 1989.9 1916.3 1994.7 1915.7 1998.8 C -1915.3 1994.7 1915.5 1990.4 1915.5 1986.3 C -1915.5 1991 1915.5 1996.4 1915.5 2000.4 C -1916.3 1999.7 1915.5 2001.4 1916.4 2001.2 C -1916.7 1996.8 1917 1990.8 1916.2 1986 C -[0.21 0.21 0 0] vc -f -S -n -1886.9 1989.6 m -1887.8 1989.2 1888.7 1988.9 1889.6 1988.4 C -1888.7 1988.9 1887.8 1989.2 1886.9 1989.6 C -[0.4 0.4 0 0] vc -f -S -n -1892.4 1986.8 m -1895.1 1985.1 1897.9 1983.6 1900.6 1982 C -1897.9 1983.6 1895.1 1985.1 1892.4 1986.8 C -f -S -n -1907.3 1979.3 m -1908.5 1978.9 1909.7 1978.5 1910.9 1978.1 C -1909.7 1978.5 1908.5 1978.9 1907.3 1979.3 C -f -S -n -1938.5 1966.6 m -1942.6 1970.1 1945.9 1976.4 1951.7 1979.1 C -1946.2 1976.1 1943.1 1970.9 1938.5 1966.6 C -f -S -n -1955.1 1978.6 m -1955.9 1978.2 1956.7 1977.8 1957.5 1977.4 C -1956.7 1977.8 1955.9 1978.2 1955.1 1978.6 C -f -S -n -1913.6 1977.6 m -1914.5 1977.2 1915.3 1976.9 1916.2 1976.4 C -1915.3 1976.9 1914.5 1977.2 1913.6 1977.6 C -f -S -n -1919.1 1974.8 m -1921.8 1973.1 1924.5 1971.6 1927.2 1970 C -1924.5 1971.6 1921.8 1973.1 1919.1 1974.8 C -f -S -n -1963.5 1974.5 m -1964.5 1974 1965.6 1973.4 1966.6 1972.8 C -1965.6 1973.4 1964.5 1974 1963.5 1974.5 C -f -S -n -vmrs -1967.8 1972.4 m -1970 1971.2 1972.1 1970 1974.3 1968.8 C -1972.1 1970 1970 1971.2 1967.8 1972.4 C -[0.4 0.4 0 0] vc -f -0.4 w -2 J -2 M -S -n -1934 1967.3 m -1935.2 1966.9 1936.4 1966.5 1937.6 1966.1 C -1936.4 1966.5 1935.2 1966.9 1934 1967.3 C -f -S -n -1928.9 2061.2 m -1928.9 2059.2 1928.9 2057.3 1928.9 2055.4 C -1928.9 2057.3 1928.9 2059.2 1928.9 2061.2 C -[0.21 0.21 0 0] vc -f -S -n -1917.2 2047 m -1917.8 2046.5 1919.6 2046.8 1920 2047.2 C -1920 2046.5 1920.9 2046.8 1921 2046.3 C -1921.9 2047.3 1921.3 2044.1 1921.5 2044.1 C -1919.7 2044.8 1915.7 2043.5 1916.2 2046 C -1916.2 2048.3 1917 2045.9 1917.2 2047 C -[0 0 0 0] vc -f -S -n -1922 2044.1 m -1923.5 2043.2 1927 2045.4 1927.5 2042.9 C -1927.1 2042.6 1927.3 2040.9 1927.2 2041.5 C -1924.9 2042.3 1920.9 2040.6 1922 2044.1 C -f -S -n -1934.9 2043.9 m -1935.2 2043.4 1934.4 2042.7 1934 2042.2 C -1933.2 2041.8 1932.4 2042.8 1932.8 2043.2 C -1932.9 2044 1934.3 2043.3 1934.9 2043.9 C -f -S -n -1906.1 2030.7 m -1906.1 2028.8 1906.1 2027 1906.1 2025.2 C -1906.1 2027 1906.1 2028.8 1906.1 2030.7 C -[0.21 0.21 0 0] vc -f -S -n -1932.8 2018.7 m -1932.8 2016.8 1932.8 2014.8 1932.8 2012.9 C -1932.8 2014.8 1932.8 2016.8 1932.8 2018.7 C -f -S -n -1894.4 2016.5 m -1895 2016 1896.8 2016.3 1897.2 2016.8 C -1897.2 2016 1898.1 2016.3 1898.2 2015.8 C -1899.1 2016.8 1898.5 2013.6 1898.7 2013.6 C -1896.9 2014.4 1892.9 2013 1893.4 2015.6 C -1893.4 2017.8 1894.2 2015.4 1894.4 2016.5 C -[0 0 0 0] vc -f -S -n -1899.2 2013.6 m -1900.7 2012.7 1904.2 2014.9 1904.7 2012.4 C -1904.3 2012.1 1904.5 2010.5 1904.4 2011 C -1902.1 2011.8 1898.1 2010.1 1899.2 2013.6 C -f -S -n -vmrs -1912.1 2013.4 m -1912.4 2012.9 1911.6 2012.3 1911.2 2011.7 C -1910.4 2011.4 1909.6 2012.3 1910 2012.7 C -1910.1 2013.5 1911.5 2012.9 1912.1 2013.4 C -[0 0 0 0] vc -f -0.4 w -2 J -2 M -S -n -1921 2004.5 m -1921.6 2004 1923.4 2004.3 1923.9 2004.8 C -1923.8 2004 1924.8 2004.3 1924.8 2003.8 C -1925.7 2004.8 1925.1 2001.6 1925.3 2001.6 C -1923.6 2002.4 1919.6 2001 1920 2003.6 C -1920 2005.8 1920.8 2003.4 1921 2004.5 C -f -S -n -1925.8 2001.6 m -1927.3 2000.7 1930.8 2002.9 1931.3 2000.4 C -1930.9 2000.1 1931.1 1998.5 1931.1 1999 C -1928.7 1999.8 1924.8 1998.1 1925.8 2001.6 C -f -S -n -1938.8 2001.4 m -1939 2000.9 1938.2 2000.3 1937.8 1999.7 C -1937.1 1999.4 1936.2 2000.3 1936.6 2000.7 C -1936.7 2001.5 1938.1 2000.9 1938.8 2001.4 C -f -S -n -1908.6691 2008.1348 m -1897.82 2010.0477 L -1894.1735 1989.3671 L -1905.0226 1987.4542 L -1908.6691 2008.1348 L -n -q -_bfh -%%IncludeResource: font Symbol -_efh -{ -f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont -1895.041763 1994.291153 m -0 0 32 0 0 (l) ts -} -true -[0 0 0 1]sts -Q -1979.2185 1991.7809 m -1960.6353 1998.5452 L -1953.4532 1978.8124 L -1972.0363 1972.0481 L -1979.2185 1991.7809 L -n -q -_bfh -%%IncludeResource: font Symbol -_efh -{ -f0 [18.793335 -6.84082 6.84021 18.793335 0 0] makesetfont -1955.163254 1983.510773 m -0 0 32 0 0 (\256) ts -} -true -[0 0 0 1]sts -Q -1952.1544 2066.5423 m -1938.0739 2069.025 L -1934.4274 2048.3444 L -1948.5079 2045.8617 L -1952.1544 2066.5423 L -n -q -_bfh -%%IncludeResource: font Symbol -_efh -{ -f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont -1935.29567 2053.268433 m -0 0 32 0 0 (") ts -} -true -[0 0 0 1]sts -Q -1931.7231 2043.621 m -1919.3084 2048.14 L -1910.6898 2024.4607 L -1923.1046 2019.9417 L -1931.7231 2043.621 L -n -q -_bfh -%%IncludeResource: font Symbol -_efh -{ -f0 [22.552002 -8.208984 8.208252 22.552002 0 0] makesetfont -1912.741867 2030.098648 m -0 0 32 0 0 (=) ts -} -true -[0 0 0 1]sts -Q -1944 2024.5 m -1944 2014 L -0.8504 w -0 J -3.863693 M -[0 0 0 1] vc -false setoverprint -S -n -1944.25 2019.1673 m -1952.5 2015.9173 L -S -n -1931.0787 2124.423 m -1855.5505 2043.4285 L -1871.0419 2013.0337 L -1946.5701 2094.0282 L -1931.0787 2124.423 L -n -q -_bfh -%%IncludeResource: font ZapfHumanist601BT-Bold -_efh -{ -f1 [22.155762 23.759277 -14.753906 28.947754 0 0] makesetfont -1867.35347 2020.27063 m -0 0 32 0 0 (Isabelle) ts -} -true -[0 0 0 1]sts -Q -1933.5503 1996.9547 m -1922.7012 1998.8677 L -1919.0547 1978.1871 L -1929.9038 1976.2741 L -1933.5503 1996.9547 L -n -q -_bfh -%%IncludeResource: font Symbol -_efh -{ -f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont -1919.922913 1983.111069 m -0 0 32 0 0 (b) ts -} -true -[0 0 0 1]sts -Q -2006.3221 2025.7184 m -1993.8573 2027.9162 L -1990.2108 2007.2356 L -2002.6756 2005.0378 L -2006.3221 2025.7184 L -n -q -_bfh -%%IncludeResource: font Symbol -_efh -{ -f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont -1991.07901 2012.159653 m -0 0 32 0 0 (a) ts -} -true -[0 0 0 1]sts -Q -vmrs -2030.0624 2094.056 m -1956.3187 2120.904 L -1956.321 2095.3175 L -2030.0647 2068.4695 L -2030.0624 2094.056 L -n -q -_bfh -%%IncludeResource: font ZapfHumanist601BT-Bold -_efh -{ -f1 [22.898804 -8.336792 -0.002197 24.368408 0 0] makesetfont -1956.320496 2101.409561 m -0 0 32 0 0 (Isar) ts -} -true -[0 0 0 1]sts -Q -vmr -vmr -end -%%Trailer -%%DocumentNeededResources: font Symbol -%%+ font ZapfHumanist601BT-Bold -%%DocumentFonts: Symbol -%%+ ZapfHumanist601BT-Bold -%%DocumentNeededFonts: Symbol -%%+ ZapfHumanist601BT-Bold diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/isabelle_isar.pdf Binary file doc-src/IsarAdvanced/Functions/isabelle_isar.pdf has changed diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/mathpartir.sty --- a/doc-src/IsarAdvanced/Functions/mathpartir.sty Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,421 +0,0 @@ -% Mathpartir --- Math Paragraph for Typesetting Inference Rules -% -% Copyright (C) 2001, 2002, 2003, 2004, 2005 Didier Rémy -% -% Author : Didier Remy -% Version : 1.2.0 -% Bug Reports : to author -% Web Site : http://pauillac.inria.fr/~remy/latex/ -% -% Mathpartir is free software; you can redistribute it and/or modify -% it under the terms of the GNU General Public License as published by -% the Free Software Foundation; either version 2, or (at your option) -% any later version. -% -% Mathpartir is distributed in the hope that it will be useful, -% but WITHOUT ANY WARRANTY; without even the implied warranty of -% MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the -% GNU General Public License for more details -% (http://pauillac.inria.fr/~remy/license/GPL). -% -%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% -% File mathpartir.sty (LaTeX macros) -%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% - -\NeedsTeXFormat{LaTeX2e} -\ProvidesPackage{mathpartir} - [2005/12/20 version 1.2.0 Math Paragraph for Typesetting Inference Rules] - -%% - -%% Identification -%% Preliminary declarations - -\RequirePackage {keyval} - -%% Options -%% More declarations - -%% PART I: Typesetting maths in paragraphe mode - -\newdimen \mpr@tmpdim - -% To ensure hevea \hva compatibility, \hva should expands to nothing -% in mathpar or in inferrule -\let \mpr@hva \empty - -%% normal paragraph parametters, should rather be taken dynamically -\def \mpr@savepar {% - \edef \MathparNormalpar - {\noexpand \lineskiplimit \the\lineskiplimit - \noexpand \lineskip \the\lineskip}% - } - -\def \mpr@rulelineskip {\lineskiplimit=0.3em\lineskip=0.2em plus 0.1em} -\def \mpr@lesslineskip {\lineskiplimit=0.6em\lineskip=0.5em plus 0.2em} -\def \mpr@lineskip {\lineskiplimit=1.2em\lineskip=1.2em plus 0.2em} -\let \MathparLineskip \mpr@lineskip -\def \mpr@paroptions {\MathparLineskip} -\let \mpr@prebindings \relax - -\newskip \mpr@andskip \mpr@andskip 2em plus 0.5fil minus 0.5em - -\def \mpr@goodbreakand - {\hskip -\mpr@andskip \penalty -1000\hskip \mpr@andskip} -\def \mpr@and {\hskip \mpr@andskip} -\def \mpr@andcr {\penalty 50\mpr@and} -\def \mpr@cr {\penalty -10000\mpr@and} -\def \mpr@eqno #1{\mpr@andcr #1\hskip 0em plus -1fil \penalty 10} - -\def \mpr@bindings {% - \let \and \mpr@andcr - \let \par \mpr@andcr - \let \\\mpr@cr - \let \eqno \mpr@eqno - \let \hva \mpr@hva - } -\let \MathparBindings \mpr@bindings - -% \@ifundefined {ignorespacesafterend} -% {\def \ignorespacesafterend {\aftergroup \ignorespaces} - -\newenvironment{mathpar}[1][] - {$$\mpr@savepar \parskip 0em \hsize \linewidth \centering - \vbox \bgroup \mpr@prebindings \mpr@paroptions #1\ifmmode $\else - \noindent $\displaystyle\fi - \MathparBindings} - {\unskip \ifmmode $\fi\egroup $$\ignorespacesafterend} - -% \def \math@mathpar #1{\setbox0 \hbox {$\displaystyle #1$}\ifnum -% \wd0 < \hsize $$\box0$$\else \bmathpar #1\emathpar \fi} - -%%% HOV BOXES - -\def \mathvbox@ #1{\hbox \bgroup \mpr@normallineskip - \vbox \bgroup \tabskip 0em \let \\ \cr - \halign \bgroup \hfil $##$\hfil\cr #1\crcr \egroup \egroup - \egroup} - -\def \mathhvbox@ #1{\setbox0 \hbox {\let \\\qquad $#1$}\ifnum \wd0 < \hsize - \box0\else \mathvbox {#1}\fi} - - -%% Part II -- operations on lists - -\newtoks \mpr@lista -\newtoks \mpr@listb - -\long \def\mpr@cons #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter -{#2}\edef #2{\the \mpr@lista \the \mpr@listb}} - -\long \def\mpr@snoc #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter -{#2}\edef #2{\the \mpr@listb\the\mpr@lista}} - -\long \def \mpr@concat#1=#2\mpr@to#3{\mpr@lista \expandafter {#2}\mpr@listb -\expandafter {#3}\edef #1{\the \mpr@listb\the\mpr@lista}} - -\def \mpr@head #1\mpr@to #2{\expandafter \mpr@head@ #1\mpr@head@ #1#2} -\long \def \mpr@head@ #1#2\mpr@head@ #3#4{\def #4{#1}\def#3{#2}} - -\def \mpr@flatten #1\mpr@to #2{\expandafter \mpr@flatten@ #1\mpr@flatten@ #1#2} -\long \def \mpr@flatten@ \\#1\\#2\mpr@flatten@ #3#4{\def #4{#1}\def #3{\\#2}} - -\def \mpr@makelist #1\mpr@to #2{\def \mpr@all {#1}% - \mpr@lista {\\}\mpr@listb \expandafter {\mpr@all}\edef \mpr@all {\the - \mpr@lista \the \mpr@listb \the \mpr@lista}\let #2\empty - \def \mpr@stripof ##1##2\mpr@stripend{\def \mpr@stripped{##2}}\loop - \mpr@flatten \mpr@all \mpr@to \mpr@one - \expandafter \mpr@snoc \mpr@one \mpr@to #2\expandafter \mpr@stripof - \mpr@all \mpr@stripend - \ifx \mpr@stripped \empty \let \mpr@isempty 0\else \let \mpr@isempty 1\fi - \ifx 1\mpr@isempty - \repeat -} - -\def \mpr@rev #1\mpr@to #2{\let \mpr@tmp \empty - \def \\##1{\mpr@cons ##1\mpr@to \mpr@tmp}#1\let #2\mpr@tmp} - -%% Part III -- Type inference rules - -\newif \if@premisse -\newbox \mpr@hlist -\newbox \mpr@vlist -\newif \ifmpr@center \mpr@centertrue -\def \mpr@htovlist {% - \setbox \mpr@hlist - \hbox {\strut - \ifmpr@center \hskip -0.5\wd\mpr@hlist\fi - \unhbox \mpr@hlist}% - \setbox \mpr@vlist - \vbox {\if@premisse \box \mpr@hlist \unvbox \mpr@vlist - \else \unvbox \mpr@vlist \box \mpr@hlist - \fi}% -} -% OLD version -% \def \mpr@htovlist {% -% \setbox \mpr@hlist -% \hbox {\strut \hskip -0.5\wd\mpr@hlist \unhbox \mpr@hlist}% -% \setbox \mpr@vlist -% \vbox {\if@premisse \box \mpr@hlist \unvbox \mpr@vlist -% \else \unvbox \mpr@vlist \box \mpr@hlist -% \fi}% -% } - -\def \mpr@item #1{$\displaystyle #1$} -\def \mpr@sep{2em} -\def \mpr@blank { } -\def \mpr@hovbox #1#2{\hbox - \bgroup - \ifx #1T\@premissetrue - \else \ifx #1B\@premissefalse - \else - \PackageError{mathpartir} - {Premisse orientation should either be T or B} - {Fatal error in Package}% - \fi \fi - \def \@test {#2}\ifx \@test \mpr@blank\else - \setbox \mpr@hlist \hbox {}% - \setbox \mpr@vlist \vbox {}% - \if@premisse \let \snoc \mpr@cons \else \let \snoc \mpr@snoc \fi - \let \@hvlist \empty \let \@rev \empty - \mpr@tmpdim 0em - \expandafter \mpr@makelist #2\mpr@to \mpr@flat - \if@premisse \mpr@rev \mpr@flat \mpr@to \@rev \else \let \@rev \mpr@flat \fi - \def \\##1{% - \def \@test {##1}\ifx \@test \empty - \mpr@htovlist - \mpr@tmpdim 0em %%% last bug fix not extensively checked - \else - \setbox0 \hbox{\mpr@item {##1}}\relax - \advance \mpr@tmpdim by \wd0 - %\mpr@tmpdim 1.02\mpr@tmpdim - \ifnum \mpr@tmpdim < \hsize - \ifnum \wd\mpr@hlist > 0 - \if@premisse - \setbox \mpr@hlist - \hbox {\unhbox0 \hskip \mpr@sep \unhbox \mpr@hlist}% - \else - \setbox \mpr@hlist - \hbox {\unhbox \mpr@hlist \hskip \mpr@sep \unhbox0}% - \fi - \else - \setbox \mpr@hlist \hbox {\unhbox0}% - \fi - \else - \ifnum \wd \mpr@hlist > 0 - \mpr@htovlist - \mpr@tmpdim \wd0 - \fi - \setbox \mpr@hlist \hbox {\unhbox0}% - \fi - \advance \mpr@tmpdim by \mpr@sep - \fi - }% - \@rev - \mpr@htovlist - \ifmpr@center \hskip \wd\mpr@vlist\fi \box \mpr@vlist - \fi - \egroup -} - -%%% INFERENCE RULES - -\@ifundefined{@@over}{% - \let\@@over\over % fallback if amsmath is not loaded - \let\@@overwithdelims\overwithdelims - \let\@@atop\atop \let\@@atopwithdelims\atopwithdelims - \let\@@above\above \let\@@abovewithdelims\abovewithdelims - }{} - -%% The default - -\def \mpr@@fraction #1#2{\hbox {\advance \hsize by -0.5em - $\displaystyle {#1\mpr@over #2}$}} -\let \mpr@fraction \mpr@@fraction - -%% A generic solution to arrow - -\def \mpr@make@fraction #1#2#3#4#5{\hbox {% - \def \mpr@tail{#1}% - \def \mpr@body{#2}% - \def \mpr@head{#3}% - \setbox1=\hbox{$#4$}\setbox2=\hbox{$#5$}% - \setbox3=\hbox{$\mkern -3mu\mpr@body\mkern -3mu$}% - \setbox3=\hbox{$\mkern -3mu \mpr@body\mkern -3mu$}% - \dimen0=\dp1\advance\dimen0 by \ht3\relax\dp1\dimen0\relax - \dimen0=\ht2\advance\dimen0 by \dp3\relax\ht2\dimen0\relax - \setbox0=\hbox {$\box1 \@@atop \box2$}% - \dimen0=\wd0\box0 - \box0 \hskip -\dimen0\relax - \hbox to \dimen0 {$% - \mathrel{\mpr@tail}\joinrel - \xleaders\hbox{\copy3}\hfil\joinrel\mathrel{\mpr@head}% - $}}} - -%% Old stuff should be removed in next version -\def \mpr@@reduce #1#2{\hbox - {$\lower 0.01pt \mpr@@fraction {#1}{#2}\mkern -15mu\rightarrow$}} -\def \mpr@@rewrite #1#2#3{\hbox - {$\lower 0.01pt \mpr@@fraction {#2}{#3}\mkern -8mu#1$}} -\def \mpr@infercenter #1{\vcenter {\mpr@hovbox{T}{#1}}} - -\def \mpr@empty {} -\def \mpr@inferrule - {\bgroup - \ifnum \linewidth<\hsize \hsize \linewidth\fi - \mpr@rulelineskip - \let \and \qquad - \let \hva \mpr@hva - \let \@rulename \mpr@empty - \let \@rule@options \mpr@empty - \let \mpr@over \@@over - \mpr@inferrule@} -\newcommand {\mpr@inferrule@}[3][] - {\everymath={\displaystyle}% - \def \@test {#2}\ifx \empty \@test - \setbox0 \hbox {$\vcenter {\mpr@hovbox{B}{#3}}$}% - \else - \def \@test {#3}\ifx \empty \@test - \setbox0 \hbox {$\vcenter {\mpr@hovbox{T}{#2}}$}% - \else - \setbox0 \mpr@fraction {\mpr@hovbox{T}{#2}}{\mpr@hovbox{B}{#3}}% - \fi \fi - \def \@test {#1}\ifx \@test\empty \box0 - \else \vbox -%%% Suggestion de Francois pour les etiquettes longues -%%% {\hbox to \wd0 {\RefTirName {#1}\hfil}\box0}\fi - {\hbox {\RefTirName {#1}}\box0}\fi - \egroup} - -\def \mpr@vdotfil #1{\vbox to #1{\leaders \hbox{$\cdot$} \vfil}} - -% They are two forms -% \inferrule [label]{[premisses}{conclusions} -% or -% \inferrule* [options]{[premisses}{conclusions} -% -% Premisses and conclusions are lists of elements separated by \\ -% Each \\ produces a break, attempting horizontal breaks if possible, -% and vertical breaks if needed. -% -% An empty element obtained by \\\\ produces a vertical break in all cases. -% -% The former rule is aligned on the fraction bar. -% The optional label appears on top of the rule -% The second form to be used in a derivation tree is aligned on the last -% line of its conclusion -% -% The second form can be parameterized, using the key=val interface. The -% folloiwng keys are recognized: -% -% width set the width of the rule to val -% narrower set the width of the rule to val\hsize -% before execute val at the beginning/left -% lab put a label [Val] on top of the rule -% lskip add negative skip on the right -% left put a left label [Val] -% Left put a left label [Val], ignoring its width -% right put a right label [Val] -% Right put a right label [Val], ignoring its width -% leftskip skip negative space on the left-hand side -% rightskip skip negative space on the right-hand side -% vdots lift the rule by val and fill vertical space with dots -% after execute val at the end/right -% -% Note that most options must come in this order to avoid strange -% typesetting (in particular leftskip must preceed left and Left and -% rightskip must follow Right or right; vdots must come last -% or be only followed by rightskip. -% - -%% Keys that make sence in all kinds of rules -\def \mprset #1{\setkeys{mprset}{#1}} -\define@key {mprset}{flushleft}[]{\mpr@centerfalse} -\define@key {mprset}{center}[]{\mpr@centertrue} -\define@key {mprset}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite} -\define@key {mprset}{myfraction}[]{\let \mpr@fraction #1} -\define@key {mprset}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}} - -\newbox \mpr@right -\define@key {mpr}{flushleft}[]{\mpr@centerfalse} -\define@key {mpr}{center}[]{\mpr@centertrue} -\define@key {mpr}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite} -\define@key {mpr}{myfraction}[]{\let \mpr@fraction #1} -\define@key {mpr}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}} -\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax - \advance \hsize by -\wd0\box0} -\define@key {mpr}{width}{\hsize #1} -\define@key {mpr}{sep}{\def\mpr@sep{#1}} -\define@key {mpr}{before}{#1} -\define@key {mpr}{lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}} -\define@key {mpr}{Lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}} -\define@key {mpr}{narrower}{\hsize #1\hsize} -\define@key {mpr}{leftskip}{\hskip -#1} -\define@key {mpr}{reduce}[]{\let \mpr@fraction \mpr@@reduce} -\define@key {mpr}{rightskip} - {\setbox \mpr@right \hbox {\unhbox \mpr@right \hskip -#1}} -\define@key {mpr}{LEFT}{\setbox0 \hbox {$#1$}\relax - \advance \hsize by -\wd0\box0} -\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax - \advance \hsize by -\wd0\box0} -\define@key {mpr}{Left}{\llap{$\TirName {#1}\;$}} -\define@key {mpr}{right} - {\setbox0 \hbox {$\;\TirName {#1}$}\relax \advance \hsize by -\wd0 - \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}} -\define@key {mpr}{RIGHT} - {\setbox0 \hbox {$#1$}\relax \advance \hsize by -\wd0 - \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}} -\define@key {mpr}{Right} - {\setbox \mpr@right \hbox {\unhbox \mpr@right \rlap {$\;\TirName {#1}$}}} -\define@key {mpr}{vdots}{\def \mpr@vdots {\@@atop \mpr@vdotfil{#1}}} -\define@key {mpr}{after}{\edef \mpr@after {\mpr@after #1}} - -\newdimen \rule@dimen -\newcommand \mpr@inferstar@ [3][]{\setbox0 - \hbox {\let \mpr@rulename \mpr@empty \let \mpr@vdots \relax - \setbox \mpr@right \hbox{}% - $\setkeys{mpr}{#1}% - \ifx \mpr@rulename \mpr@empty \mpr@inferrule {#2}{#3}\else - \mpr@inferrule [{\mpr@rulename}]{#2}{#3}\fi - \box \mpr@right \mpr@vdots$} - \setbox1 \hbox {\strut} - \rule@dimen \dp0 \advance \rule@dimen by -\dp1 - \raise \rule@dimen \box0} - -\def \mpr@infer {\@ifnextchar *{\mpr@inferstar}{\mpr@inferrule}} -\newcommand \mpr@err@skipargs[3][]{} -\def \mpr@inferstar*{\ifmmode - \let \@do \mpr@inferstar@ - \else - \let \@do \mpr@err@skipargs - \PackageError {mathpartir} - {\string\inferrule* can only be used in math mode}{}% - \fi \@do} - - -%%% Exports - -% Envirnonment mathpar - -\let \inferrule \mpr@infer - -% make a short name \infer is not already defined -\@ifundefined {infer}{\let \infer \mpr@infer}{} - -\def \TirNameStyle #1{\small \textsc{#1}} -\def \tir@name #1{\hbox {\small \TirNameStyle{#1}}} -\let \TirName \tir@name -\let \DefTirName \TirName -\let \RefTirName \TirName - -%%% Other Exports - -% \let \listcons \mpr@cons -% \let \listsnoc \mpr@snoc -% \let \listhead \mpr@head -% \let \listmake \mpr@makelist - - - - -\endinput diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Functions/style.sty --- a/doc-src/IsarAdvanced/Functions/style.sty Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,62 +0,0 @@ - -%% $Id$ - -%% toc -\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} -\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} - -%% references -\newcommand{\secref}[1]{\S\ref{#1}} -\newcommand{\chref}[1]{chapter~\ref{#1}} -\newcommand{\figref}[1]{figure~\ref{#1}} - -%% glossary -\renewcommand{\glossary}[2]{\nomenclature{\bf #1}{#2}} -\newcommand{\seeglossary}[1]{\emph{#1}} -\newcommand{\glossaryname}{Glossary} -\renewcommand{\nomname}{\glossaryname} -\renewcommand{\pagedeclaration}[1]{\nobreak\quad\dotfill~page~\bold{#1}} - -%% index -\newcommand{\indexml}[1]{\index{\emph{#1}|bold}} -\newcommand{\indexmltype}[1]{\index{\emph{#1} (type)|bold}} -\newcommand{\indexmlstructure}[1]{\index{\emph{#1} (structure)|bold}} -\newcommand{\indexmlfunctor}[1]{\index{\emph{#1} (functor)|bold}} - -%% math -\newcommand{\text}[1]{\mbox{#1}} -\newcommand{\isasymvartheta}{\isamath{\theta}} -\newcommand{\isactrlvec}[1]{\emph{$\overline{#1}$}} - -\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} - -\pagestyle{headings} -\sloppy -\binperiod -\underscoreon - -\renewcommand{\isadigit}[1]{\isamath{#1}} - -\newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup} - -\isafoldtag{FIXME} -\isakeeptag{mlref} -\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small} -\renewcommand{\endisatagmlref}{\endgroup} - -\newcommand{\isasymGUESS}{\isakeyword{guess}} -\newcommand{\isasymOBTAIN}{\isakeyword{obtain}} -\newcommand{\isasymTHEORY}{\isakeyword{theory}} -\newcommand{\isasymUSES}{\isakeyword{uses}} -\newcommand{\isasymEND}{\isakeyword{end}} -\newcommand{\isasymCONSTS}{\isakeyword{consts}} -\newcommand{\isasymDEFS}{\isakeyword{defs}} -\newcommand{\isasymTHEOREM}{\isakeyword{theorem}} -\newcommand{\isasymDEFINITION}{\isakeyword{definition}} - -\isabellestyle{it} - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "implementation" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarAdvanced/Makefile.in --- a/doc-src/IsarAdvanced/Makefile.in Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,13 +0,0 @@ -# $Id$ -# - -include ../../Makefile.in - -SEDINDEX = ../../sedindex -FIXBOOKMARKS = perl -pi ../../fixbookmarks.pl - -isabelle_isar.eps: - test -r isabelle_isar.eps || ln -s ../../gfx/isabelle_isar.eps . - -isabelle_isar.pdf: - test -r isabelle_isar.pdf || ln -s ../../gfx/isabelle_isar.pdf . diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/IsaMakefile --- a/doc-src/IsarImplementation/IsaMakefile Wed Mar 04 11:05:02 2009 +0100 +++ b/doc-src/IsarImplementation/IsaMakefile Wed Mar 04 11:05:29 2009 +0100 @@ -21,9 +21,10 @@ Thy: $(LOG)/Pure-Thy.gz -$(LOG)/Pure-Thy.gz: Thy/ROOT.ML Thy/base.thy Thy/integration.thy Thy/isar.thy \ - Thy/locale.thy Thy/logic.thy Thy/prelim.thy Thy/proof.thy Thy/tactic.thy \ - Thy/ML.thy ../antiquote_setup.ML +$(LOG)/Pure-Thy.gz: Thy/ROOT.ML Thy/Base.thy Thy/Integration.thy \ + Thy/Isar.thy Thy/Local_Theory.thy Thy/Logic.thy Thy/Prelim.thy \ + Thy/Proof.thy Thy/Syntax.thy Thy/Tactic.thy Thy/ML.thy \ + ../antiquote_setup.ML @$(USEDIR) Pure Thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Makefile --- a/doc-src/IsarImplementation/Makefile Wed Mar 04 11:05:02 2009 +0100 +++ b/doc-src/IsarImplementation/Makefile Wed Mar 04 11:05:29 2009 +0100 @@ -1,6 +1,3 @@ -# -# $Id$ -# ## targets @@ -11,16 +8,14 @@ include ../Makefile.in -MAKEGLOSSARY = ./makeglossary - NAME = implementation -FILES = implementation.tex intro.tex Thy/document/prelim.tex \ - Thy/document/logic.tex Thy/document/tactic.tex \ - Thy/document/proof.tex Thy/document/locale.tex \ - Thy/document/integration.tex style.sty ../iman.sty ../extra.sty \ - ../isar.sty ../isabelle.sty ../isabellesym.sty ../pdfsetup.sty \ - ../manual.bib ../proof.sty +FILES = ../extra.sty ../iman.sty ../isabelle.sty ../isabellesym.sty \ + ../isar.sty ../manual.bib ../pdfsetup.sty ../proof.sty \ + Thy/document/Integration.tex Thy/document/Local_Theory.tex \ + Thy/document/Logic.tex Thy/document/Prelim.tex \ + Thy/document/Proof.tex Thy/document/Syntax.tex \ + Thy/document/Tactic.tex implementation.tex style.sty dvi: $(NAME).dvi @@ -29,7 +24,6 @@ $(BIBTEX) $(NAME) $(LATEX) $(NAME) $(LATEX) $(NAME) - $(MAKEGLOSSARY) $(NAME) $(SEDINDEX) $(NAME) $(LATEX) $(NAME) $(LATEX) $(NAME) @@ -41,7 +35,6 @@ $(BIBTEX) $(NAME) $(PDFLATEX) $(NAME) $(PDFLATEX) $(NAME) - $(MAKEGLOSSARY) $(NAME) $(SEDINDEX) $(NAME) $(FIXBOOKMARKS) $(NAME).out $(PDFLATEX) $(NAME) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/Local_Theory.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarImplementation/Thy/Local_Theory.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,168 @@ +theory Local_Theory +imports Base +begin + +chapter {* Local theory specifications *} + +text {* + A \emph{local theory} combines aspects of both theory and proof + context (cf.\ \secref{sec:context}), such that definitional + specifications may be given relatively to parameters and + assumptions. A local theory is represented as a regular proof + context, augmented by administrative data about the \emph{target + context}. + + The target is usually derived from the background theory by adding + local @{text "\"} and @{text "\"} elements, plus + suitable modifications of non-logical context data (e.g.\ a special + type-checking discipline). Once initialized, the target is ready to + absorb definitional primitives: @{text "\"} for terms and + @{text "\"} for theorems. Such definitions may get + transformed in a target-specific way, but the programming interface + hides such details. + + Isabelle/Pure provides target mechanisms for locales, type-classes, + type-class instantiations, and general overloading. In principle, + users can implement new targets as well, but this rather arcane + discipline is beyond the scope of this manual. In contrast, + implementing derived definitional packages to be used within a local + theory context is quite easy: the interfaces are even simpler and + more abstract than the underlying primitives for raw theories. + + Many definitional packages for local theories are available in + Isabelle. Although a few old packages only work for global + theories, the local theory interface is already the standard way of + implementing definitional packages in Isabelle. +*} + + +section {* Definitional elements *} + +text {* + There are separate elements @{text "\ c \ t"} for terms, and + @{text "\ b = thm"} for theorems. Types are treated + implicitly, according to Hindley-Milner discipline (cf.\ + \secref{sec:variables}). These definitional primitives essentially + act like @{text "let"}-bindings within a local context that may + already contain earlier @{text "let"}-bindings and some initial + @{text "\"}-bindings. Thus we gain \emph{dependent definitions} + that are relative to an initial axiomatic context. The following + diagram illustrates this idea of axiomatic elements versus + definitional elements: + + \begin{center} + \begin{tabular}{|l|l|l|} + \hline + & @{text "\"}-binding & @{text "let"}-binding \\ + \hline + types & fixed @{text "\"} & arbitrary @{text "\"} \\ + terms & @{text "\ x :: \"} & @{text "\ c \ t"} \\ + theorems & @{text "\ a: A"} & @{text "\ b = \<^BG>B\<^EN>"} \\ + \hline + \end{tabular} + \end{center} + + A user package merely needs to produce suitable @{text "\"} + and @{text "\"} elements according to the application. For + example, a package for inductive definitions might first @{text + "\"} a certain predicate as some fixed-point construction, + then @{text "\"} a proven result about monotonicity of the + functor involved here, and then produce further derived concepts via + additional @{text "\"} and @{text "\"} elements. + + The cumulative sequence of @{text "\"} and @{text "\"} + produced at package runtime is managed by the local theory + infrastructure by means of an \emph{auxiliary context}. Thus the + system holds up the impression of working within a fully abstract + situation with hypothetical entities: @{text "\ c \ t"} + always results in a literal fact @{text "\<^BG>c \ t\<^EN>"}, where + @{text "c"} is a fixed variable @{text "c"}. The details about + global constants, name spaces etc. are handled internally. + + So the general structure of a local theory is a sandwich of three + layers: + + \begin{center} + \framebox{\quad auxiliary context \quad\framebox{\quad target context \quad\framebox{\quad background theory\quad}}} + \end{center} + + \noindent When a definitional package is finished, the auxiliary + context is reset to the target context. The target now holds + definitions for terms and theorems that stem from the hypothetical + @{text "\"} and @{text "\"} elements, transformed by + the particular target policy (see + \cite[\S4--5]{Haftmann-Wenzel:2009} for details). +*} + +text %mlref {* + \begin{mldecls} + @{index_ML_type local_theory: Proof.context} \\ + @{index_ML TheoryTarget.init: "string option -> theory -> local_theory"} \\[1ex] + @{index_ML LocalTheory.define: "string -> + (binding * mixfix) * (Attrib.binding * term) -> local_theory -> + (term * (string * thm)) * local_theory"} \\ + @{index_ML LocalTheory.note: "string -> + Attrib.binding * thm list -> local_theory -> + (string * thm list) * local_theory"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML_type local_theory} represents local theories. Although + this is merely an alias for @{ML_type Proof.context}, it is + semantically a subtype of the same: a @{ML_type local_theory} holds + target information as special context data. Subtyping means that + any value @{text "lthy:"}~@{ML_type local_theory} can be also used + with operations on expecting a regular @{text "ctxt:"}~@{ML_type + Proof.context}. + + \item @{ML TheoryTarget.init}~@{text "NONE thy"} initializes a + trivial local theory from the given background theory. + Alternatively, @{text "SOME name"} may be given to initialize a + @{command locale} or @{command class} context (a fully-qualified + internal name is expected here). This is useful for experimentation + --- normally the Isar toplevel already takes care to initialize the + local theory context. + + \item @{ML LocalTheory.define}~@{text "kind ((b, mx), (a, rhs)) + lthy"} defines a local entity according to the specification that is + given relatively to the current @{text "lthy"} context. In + particular the term of the RHS may refer to earlier local entities + from the auxiliary context, or hypothetical parameters from the + target context. The result is the newly defined term (which is + always a fixed variable with exactly the same name as specified for + the LHS), together with an equational theorem that states the + definition as a hypothetical fact. + + Unless an explicit name binding is given for the RHS, the resulting + fact will be called @{text "b_def"}. Any given attributes are + applied to that same fact --- immediately in the auxiliary context + \emph{and} in any transformed versions stemming from target-specific + policies or any later interpretations of results from the target + context (think of @{command locale} and @{command interpretation}, + for example). This means that attributes should be usually plain + declarations such as @{attribute simp}, while non-trivial rules like + @{attribute simplified} are better avoided. + + The @{text kind} determines the theorem kind tag of the resulting + fact. Typical examples are @{ML Thm.definitionK}, @{ML + Thm.theoremK}, or @{ML Thm.internalK}. + + \item @{ML LocalTheory.note}~@{text "kind (a, ths) lthy"} is + analogous to @{ML LocalTheory.define}, but defines facts instead of + terms. There is also a slightly more general variant @{ML + LocalTheory.notes} that defines several facts (with attribute + expressions) simultaneously. + + This is essentially the internal version of the @{command lemmas} + command, or @{command declare} if an empty name binding is given. + + \end{description} +*} + + +section {* Morphisms and declarations *} + +text FIXME + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/ROOT.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/Syntax.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarImplementation/Thy/Syntax.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,9 @@ +theory Syntax +imports Base +begin + +chapter {* Syntax and type-checking *} + +text FIXME + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/base.thy --- a/doc-src/IsarImplementation/Thy/base.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,9 +0,0 @@ - -(* $Id$ *) - -theory base -imports Pure -uses "../../antiquote_setup.ML" -begin - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/document/Local_Theory.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarImplementation/Thy/document/Local_Theory.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,220 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Local{\isacharunderscore}Theory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Local{\isacharunderscore}Theory\isanewline +\isakeyword{imports}\ Base\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupchapter{Local theory specifications% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +A \emph{local theory} combines aspects of both theory and proof + context (cf.\ \secref{sec:context}), such that definitional + specifications may be given relatively to parameters and + assumptions. A local theory is represented as a regular proof + context, augmented by administrative data about the \emph{target + context}. + + The target is usually derived from the background theory by adding + local \isa{{\isasymFIX}} and \isa{{\isasymASSUME}} elements, plus + suitable modifications of non-logical context data (e.g.\ a special + type-checking discipline). Once initialized, the target is ready to + absorb definitional primitives: \isa{{\isasymDEFINE}} for terms and + \isa{{\isasymNOTE}} for theorems. Such definitions may get + transformed in a target-specific way, but the programming interface + hides such details. + + Isabelle/Pure provides target mechanisms for locales, type-classes, + type-class instantiations, and general overloading. In principle, + users can implement new targets as well, but this rather arcane + discipline is beyond the scope of this manual. In contrast, + implementing derived definitional packages to be used within a local + theory context is quite easy: the interfaces are even simpler and + more abstract than the underlying primitives for raw theories. + + Many definitional packages for local theories are available in + Isabelle. Although a few old packages only work for global + theories, the local theory interface is already the standard way of + implementing definitional packages in Isabelle.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Definitional elements% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +There are separate elements \isa{{\isasymDEFINE}\ c\ {\isasymequiv}\ t} for terms, and + \isa{{\isasymNOTE}\ b\ {\isacharequal}\ thm} for theorems. Types are treated + implicitly, according to Hindley-Milner discipline (cf.\ + \secref{sec:variables}). These definitional primitives essentially + act like \isa{let}-bindings within a local context that may + already contain earlier \isa{let}-bindings and some initial + \isa{{\isasymlambda}}-bindings. Thus we gain \emph{dependent definitions} + that are relative to an initial axiomatic context. The following + diagram illustrates this idea of axiomatic elements versus + definitional elements: + + \begin{center} + \begin{tabular}{|l|l|l|} + \hline + & \isa{{\isasymlambda}}-binding & \isa{let}-binding \\ + \hline + types & fixed \isa{{\isasymalpha}} & arbitrary \isa{{\isasymbeta}} \\ + terms & \isa{{\isasymFIX}\ x\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}} & \isa{{\isasymDEFINE}\ c\ {\isasymequiv}\ t} \\ + theorems & \isa{{\isasymASSUME}\ a{\isacharcolon}\ A} & \isa{{\isasymNOTE}\ b\ {\isacharequal}\ \isactrlBG B\isactrlEN } \\ + \hline + \end{tabular} + \end{center} + + A user package merely needs to produce suitable \isa{{\isasymDEFINE}} + and \isa{{\isasymNOTE}} elements according to the application. For + example, a package for inductive definitions might first \isa{{\isasymDEFINE}} a certain predicate as some fixed-point construction, + then \isa{{\isasymNOTE}} a proven result about monotonicity of the + functor involved here, and then produce further derived concepts via + additional \isa{{\isasymDEFINE}} and \isa{{\isasymNOTE}} elements. + + The cumulative sequence of \isa{{\isasymDEFINE}} and \isa{{\isasymNOTE}} + produced at package runtime is managed by the local theory + infrastructure by means of an \emph{auxiliary context}. Thus the + system holds up the impression of working within a fully abstract + situation with hypothetical entities: \isa{{\isasymDEFINE}\ c\ {\isasymequiv}\ t} + always results in a literal fact \isa{\isactrlBG c\ {\isasymequiv}\ t\isactrlEN }, where + \isa{c} is a fixed variable \isa{c}. The details about + global constants, name spaces etc. are handled internally. + + So the general structure of a local theory is a sandwich of three + layers: + + \begin{center} + \framebox{\quad auxiliary context \quad\framebox{\quad target context \quad\framebox{\quad background theory\quad}}} + \end{center} + + \noindent When a definitional package is finished, the auxiliary + context is reset to the target context. The target now holds + definitions for terms and theorems that stem from the hypothetical + \isa{{\isasymDEFINE}} and \isa{{\isasymNOTE}} elements, transformed by + the particular target policy (see + \cite[\S4--5]{Haftmann-Wenzel:2009} for details).% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML type}{local\_theory}\verb|type local_theory = Proof.context| \\ + \indexdef{}{ML}{TheoryTarget.init}\verb|TheoryTarget.init: string option -> theory -> local_theory| \\[1ex] + \indexdef{}{ML}{LocalTheory.define}\verb|LocalTheory.define: string ->|\isasep\isanewline% +\verb| (binding * mixfix) * (Attrib.binding * term) -> local_theory ->|\isasep\isanewline% +\verb| (term * (string * thm)) * local_theory| \\ + \indexdef{}{ML}{LocalTheory.note}\verb|LocalTheory.note: string ->|\isasep\isanewline% +\verb| Attrib.binding * thm list -> local_theory ->|\isasep\isanewline% +\verb| (string * thm list) * local_theory| \\ + \end{mldecls} + + \begin{description} + + \item \verb|local_theory| represents local theories. Although + this is merely an alias for \verb|Proof.context|, it is + semantically a subtype of the same: a \verb|local_theory| holds + target information as special context data. Subtyping means that + any value \isa{lthy{\isacharcolon}}~\verb|local_theory| can be also used + with operations on expecting a regular \isa{ctxt{\isacharcolon}}~\verb|Proof.context|. + + \item \verb|TheoryTarget.init|~\isa{NONE\ thy} initializes a + trivial local theory from the given background theory. + Alternatively, \isa{SOME\ name} may be given to initialize a + \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}} or \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} context (a fully-qualified + internal name is expected here). This is useful for experimentation + --- normally the Isar toplevel already takes care to initialize the + local theory context. + + \item \verb|LocalTheory.define|~\isa{kind\ {\isacharparenleft}{\isacharparenleft}b{\isacharcomma}\ mx{\isacharparenright}{\isacharcomma}\ {\isacharparenleft}a{\isacharcomma}\ rhs{\isacharparenright}{\isacharparenright}\ lthy} defines a local entity according to the specification that is + given relatively to the current \isa{lthy} context. In + particular the term of the RHS may refer to earlier local entities + from the auxiliary context, or hypothetical parameters from the + target context. The result is the newly defined term (which is + always a fixed variable with exactly the same name as specified for + the LHS), together with an equational theorem that states the + definition as a hypothetical fact. + + Unless an explicit name binding is given for the RHS, the resulting + fact will be called \isa{b{\isacharunderscore}def}. Any given attributes are + applied to that same fact --- immediately in the auxiliary context + \emph{and} in any transformed versions stemming from target-specific + policies or any later interpretations of results from the target + context (think of \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}} and \hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}}, + for example). This means that attributes should be usually plain + declarations such as \hyperlink{attribute.simp}{\mbox{\isa{simp}}}, while non-trivial rules like + \hyperlink{attribute.simplified}{\mbox{\isa{simplified}}} are better avoided. + + The \isa{kind} determines the theorem kind tag of the resulting + fact. Typical examples are \verb|Thm.definitionK|, \verb|Thm.theoremK|, or \verb|Thm.internalK|. + + \item \verb|LocalTheory.note|~\isa{kind\ {\isacharparenleft}a{\isacharcomma}\ ths{\isacharparenright}\ lthy} is + analogous to \verb|LocalTheory.define|, but defines facts instead of + terms. There is also a slightly more general variant \verb|LocalTheory.notes| that defines several facts (with attribute + expressions) simultaneously. + + This is essentially the internal version of the \hyperlink{command.lemmas}{\mbox{\isa{\isacommand{lemmas}}}} + command, or \hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}} if an empty name binding is given. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsection{Morphisms and declarations% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +FIXME% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/document/ML.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/document/Syntax.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarImplementation/Thy/document/Syntax.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,48 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Syntax}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Syntax\isanewline +\isakeyword{imports}\ Base\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupchapter{Syntax and type-checking% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +FIXME% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/document/base.tex --- a/doc-src/IsarImplementation/Thy/document/base.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,32 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{base}% -% -\isadelimtheory -\isanewline -\isanewline -\isanewline -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ base\isanewline -\isakeyword{imports}\ Pure\isanewline -\isakeyword{uses}\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isacharslash}{\isachardot}{\isachardot}{\isacharslash}antiquote{\isacharunderscore}setup{\isachardot}ML{\isachardoublequoteclose}\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -\isanewline -% -\endisadelimtheory -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/document/integration.tex --- a/doc-src/IsarImplementation/Thy/document/integration.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,521 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{integration}% -% -\isadelimtheory -\isanewline -\isanewline -\isanewline -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ integration\ \isakeyword{imports}\ base\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{System integration% -} -\isamarkuptrue% -% -\isamarkupsection{Isar toplevel \label{sec:isar-toplevel}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The Isar toplevel may be considered the centeral hub of the - Isabelle/Isar system, where all key components and sub-systems are - integrated into a single read-eval-print loop of Isar commands. We - shall even incorporate the existing {\ML} toplevel of the compiler - and run-time system (cf.\ \secref{sec:ML-toplevel}). - - Isabelle/Isar departs from the original ``LCF system architecture'' - where {\ML} was really The Meta Language for defining theories and - conducting proofs. Instead, {\ML} now only serves as the - implementation language for the system (and user extensions), while - the specific Isar toplevel supports the concepts of theory and proof - development natively. This includes the graph structure of theories - and the block structure of proofs, support for unlimited undo, - facilities for tracing, debugging, timing, profiling etc. - - \medskip The toplevel maintains an implicit state, which is - transformed by a sequence of transitions -- either interactively or - in batch-mode. In interactive mode, Isar state transitions are - encapsulated as safe transactions, such that both failure and undo - are handled conveniently without destroying the underlying draft - theory (cf.~\secref{sec:context-theory}). In batch mode, - transitions operate in a linear (destructive) fashion, such that - error conditions abort the present attempt to construct a theory or - proof altogether. - - The toplevel state is a disjoint sum of empty \isa{toplevel}, or - \isa{theory}, or \isa{proof}. On entering the main Isar loop we - start with an empty toplevel. A theory is commenced by giving a - \isa{{\isasymTHEORY}} header; within a theory we may issue theory - commands such as \isa{{\isasymDEFINITION}}, or state a \isa{{\isasymTHEOREM}} to be proven. Now we are within a proof state, with a - rich collection of Isar proof commands for structured proof - composition, or unstructured proof scripts. When the proof is - concluded we get back to the theory, which is then updated by - storing the resulting fact. Further theory declarations or theorem - statements with proofs may follow, until we eventually conclude the - theory development by issuing \isa{{\isasymEND}}. The resulting theory - is then stored within the theory database and we are back to the - empty toplevel. - - In addition to these proper state transformations, there are also - some diagnostic commands for peeking at the toplevel state without - modifying it (e.g.\ \isakeyword{thm}, \isakeyword{term}, - \isakeyword{print-cases}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexmltype{Toplevel.state}\verb|type Toplevel.state| \\ - \indexml{Toplevel.UNDEF}\verb|Toplevel.UNDEF: exn| \\ - \indexml{Toplevel.is\_toplevel}\verb|Toplevel.is_toplevel: Toplevel.state -> bool| \\ - \indexml{Toplevel.theory\_of}\verb|Toplevel.theory_of: Toplevel.state -> theory| \\ - \indexml{Toplevel.proof\_of}\verb|Toplevel.proof_of: Toplevel.state -> Proof.state| \\ - \indexml{Toplevel.debug}\verb|Toplevel.debug: bool ref| \\ - \indexml{Toplevel.timing}\verb|Toplevel.timing: bool ref| \\ - \indexml{Toplevel.profiling}\verb|Toplevel.profiling: int ref| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Toplevel.state| represents Isar toplevel states, - which are normally manipulated through the concept of toplevel - transitions only (\secref{sec:toplevel-transition}). Also note that - a raw toplevel state is subject to the same linearity restrictions - as a theory context (cf.~\secref{sec:context-theory}). - - \item \verb|Toplevel.UNDEF| is raised for undefined toplevel - operations. Many operations work only partially for certain cases, - since \verb|Toplevel.state| is a sum type. - - \item \verb|Toplevel.is_toplevel|~\isa{state} checks for an empty - toplevel state. - - \item \verb|Toplevel.theory_of|~\isa{state} selects the theory of - a theory or proof (!), otherwise raises \verb|Toplevel.UNDEF|. - - \item \verb|Toplevel.proof_of|~\isa{state} selects the Isar proof - state if available, otherwise raises \verb|Toplevel.UNDEF|. - - \item \verb|set Toplevel.debug| makes the toplevel print further - details about internal error conditions, exceptions being raised - etc. - - \item \verb|set Toplevel.timing| makes the toplevel print timing - information for each Isar command being executed. - - \item \verb|Toplevel.profiling|~\verb|:=|~\isa{n} controls - low-level profiling of the underlying {\ML} runtime system. For - Poly/ML, \isa{n\ {\isacharequal}\ {\isadigit{1}}} means time and \isa{n\ {\isacharequal}\ {\isadigit{2}}} space - profiling. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Toplevel transitions \label{sec:toplevel-transition}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -An Isar toplevel transition consists of a partial function on the - toplevel state, with additional information for diagnostics and - error reporting: there are fields for command name, source position, - optional source text, as well as flags for interactive-only commands - (which issue a warning in batch-mode), printing of result state, - etc. - - The operational part is represented as the sequential union of a - list of partial functions, which are tried in turn until the first - one succeeds. This acts like an outer case-expression for various - alternative state transitions. For example, \isakeyword{qed} acts - differently for a local proofs vs.\ the global ending of the main - proof. - - Toplevel transitions are composed via transition transformers. - Internally, Isar commands are put together from an empty transition - extended by name and source position (and optional source text). It - is then left to the individual command parser to turn the given - concrete syntax into a suitable transition transformer that adjoin - actual operations on a theory or proof state etc.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Toplevel.print}\verb|Toplevel.print: Toplevel.transition -> Toplevel.transition| \\ - \indexml{Toplevel.no\_timing}\verb|Toplevel.no_timing: Toplevel.transition -> Toplevel.transition| \\ - \indexml{Toplevel.keep}\verb|Toplevel.keep: (Toplevel.state -> unit) ->|\isasep\isanewline% -\verb| Toplevel.transition -> Toplevel.transition| \\ - \indexml{Toplevel.theory}\verb|Toplevel.theory: (theory -> theory) ->|\isasep\isanewline% -\verb| Toplevel.transition -> Toplevel.transition| \\ - \indexml{Toplevel.theory\_to\_proof}\verb|Toplevel.theory_to_proof: (theory -> Proof.state) ->|\isasep\isanewline% -\verb| Toplevel.transition -> Toplevel.transition| \\ - \indexml{Toplevel.proof}\verb|Toplevel.proof: (Proof.state -> Proof.state) ->|\isasep\isanewline% -\verb| Toplevel.transition -> Toplevel.transition| \\ - \indexml{Toplevel.proofs}\verb|Toplevel.proofs: (Proof.state -> Proof.state Seq.seq) ->|\isasep\isanewline% -\verb| Toplevel.transition -> Toplevel.transition| \\ - \indexml{Toplevel.end\_proof}\verb|Toplevel.end_proof: (bool -> Proof.state -> Proof.context) ->|\isasep\isanewline% -\verb| Toplevel.transition -> Toplevel.transition| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Toplevel.print|~\isa{tr} sets the print flag, which - causes the toplevel loop to echo the result state (in interactive - mode). - - \item \verb|Toplevel.no_timing|~\isa{tr} indicates that the - transition should never show timing information, e.g.\ because it is - a diagnostic command. - - \item \verb|Toplevel.keep|~\isa{tr} adjoins a diagnostic - function. - - \item \verb|Toplevel.theory|~\isa{tr} adjoins a theory - transformer. - - \item \verb|Toplevel.theory_to_proof|~\isa{tr} adjoins a global - goal function, which turns a theory into a proof state. The theory - may be changed before entering the proof; the generic Isar goal - setup includes an argument that specifies how to apply the proven - result to the theory, when the proof is finished. - - \item \verb|Toplevel.proof|~\isa{tr} adjoins a deterministic - proof command, with a singleton result. - - \item \verb|Toplevel.proofs|~\isa{tr} adjoins a general proof - command, with zero or more result states (represented as a lazy - list). - - \item \verb|Toplevel.end_proof|~\isa{tr} adjoins a concluding - proof command, that returns the resulting theory, after storing the - resulting facts in the context etc. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Toplevel control% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -There are a few special control commands that modify the behavior - the toplevel itself, and only make sense in interactive mode. Under - normal circumstances, the user encounters these only implicitly as - part of the protocol between the Isabelle/Isar system and a - user-interface such as ProofGeneral. - - \begin{description} - - \item \isacommand{undo} follows the three-level hierarchy of empty - toplevel vs.\ theory vs.\ proof: undo within a proof reverts to the - previous proof context, undo after a proof reverts to the theory - before the initial goal statement, undo of a theory command reverts - to the previous theory value, undo of a theory header discontinues - the current theory development and removes it from the theory - database (\secref{sec:theory-database}). - - \item \isacommand{kill} aborts the current level of development: - kill in a proof context reverts to the theory before the initial - goal statement, kill in a theory context aborts the current theory - development, removing it from the database. - - \item \isacommand{exit} drops out of the Isar toplevel into the - underlying {\ML} toplevel (\secref{sec:ML-toplevel}). The Isar - toplevel state is preserved and may be continued later. - - \item \isacommand{quit} terminates the Isabelle/Isar process without - saving. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{ML toplevel \label{sec:ML-toplevel}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The {\ML} toplevel provides a read-compile-eval-print loop for {\ML} - values, types, structures, and functors. {\ML} declarations operate - on the global system state, which consists of the compiler - environment plus the values of {\ML} reference variables. There is - no clean way to undo {\ML} declarations, except for reverting to a - previously saved state of the whole Isabelle process. {\ML} input - is either read interactively from a TTY, or from a string (usually - within a theory text), or from a source file (usually loaded from a - theory). - - Whenever the {\ML} toplevel is active, the current Isabelle theory - context is passed as an internal reference variable. Thus {\ML} - code may access the theory context during compilation, it may even - change the value of a theory being under construction --- while - observing the usual linearity restrictions - (cf.~\secref{sec:context-theory}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{the\_context}\verb|the_context: unit -> theory| \\ - \indexml{Context.$>$$>$ }\verb|Context.>> : (Context.generic -> Context.generic) -> unit| \\ - \end{mldecls} - - \begin{description} - - \item \verb|the_context ()| refers to the theory context of the - {\ML} toplevel --- at compile time! {\ML} code needs to take care - to refer to \verb|the_context ()| correctly. Recall that - evaluation of a function body is delayed until actual runtime. - Moreover, persistent {\ML} toplevel bindings to an unfinished theory - should be avoided: code should either project out the desired - information immediately, or produce an explicit \verb|theory_ref| (cf.\ \secref{sec:context-theory}). - - \item \verb|Context.>>|~\isa{f} applies context transformation - \isa{f} to the implicit context of the {\ML} toplevel. - - \end{description} - - It is very important to note that the above functions are really - restricted to the compile time, even though the {\ML} compiler is - invoked at runtime! The majority of {\ML} code uses explicit - functional arguments of a theory or proof context instead. Thus it - may be invoked for an arbitrary context later on, without having to - worry about any operational details. - - \bigskip - - \begin{mldecls} - \indexml{Isar.main}\verb|Isar.main: unit -> unit| \\ - \indexml{Isar.loop}\verb|Isar.loop: unit -> unit| \\ - \indexml{Isar.state}\verb|Isar.state: unit -> Toplevel.state| \\ - \indexml{Isar.exn}\verb|Isar.exn: unit -> (exn * string) option| \\ - \indexml{Isar.context}\verb|Isar.context: unit -> Proof.context| \\ - \indexml{Isar.goal}\verb|Isar.goal: unit -> thm| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Isar.main ()| invokes the Isar toplevel from {\ML}, - initializing an empty toplevel state. - - \item \verb|Isar.loop ()| continues the Isar toplevel with the - current state, after having dropped out of the Isar toplevel loop. - - \item \verb|Isar.state ()| and \verb|Isar.exn ()| get current - toplevel state and error condition, respectively. This only works - after having dropped out of the Isar toplevel loop. - - \item \verb|Isar.context ()| produces the proof context from \verb|Isar.state ()|, analogous to \verb|Context.proof_of| - (\secref{sec:generic-context}). - - \item \verb|Isar.goal ()| picks the tactical goal from \verb|Isar.state ()|, represented as a theorem according to - \secref{sec:tactical-goals}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsection{Theory database \label{sec:theory-database}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The theory database maintains a collection of theories, together - with some administrative information about their original sources, - which are held in an external store (i.e.\ some directory within the - regular file system). - - The theory database is organized as a directed acyclic graph; - entries are referenced by theory name. Although some additional - interfaces allow to include a directory specification as well, this - is only a hint to the underlying theory loader. The internal theory - name space is flat! - - Theory \isa{A} is associated with the main theory file \isa{A}\verb,.thy,, which needs to be accessible through the theory - loader path. Any number of additional {\ML} source files may be - associated with each theory, by declaring these dependencies in the - theory header as \isa{{\isasymUSES}}, and loading them consecutively - within the theory context. The system keeps track of incoming {\ML} - sources and associates them with the current theory. The file - \isa{A}\verb,.ML, is loaded after a theory has been concluded, in - order to support legacy proof {\ML} proof scripts. - - The basic internal actions of the theory database are \isa{update}, \isa{outdate}, and \isa{remove}: - - \begin{itemize} - - \item \isa{update\ A} introduces a link of \isa{A} with a - \isa{theory} value of the same name; it asserts that the theory - sources are now consistent with that value; - - \item \isa{outdate\ A} invalidates the link of a theory database - entry to its sources, but retains the present theory value; - - \item \isa{remove\ A} deletes entry \isa{A} from the theory - database. - - \end{itemize} - - These actions are propagated to sub- or super-graphs of a theory - entry as expected, in order to preserve global consistency of the - state of all loaded theories with the sources of the external store. - This implies certain causalities between actions: \isa{update} - or \isa{outdate} of an entry will \isa{outdate} all - descendants; \isa{remove} will \isa{remove} all descendants. - - \medskip There are separate user-level interfaces to operate on the - theory database directly or indirectly. The primitive actions then - just happen automatically while working with the system. In - particular, processing a theory header \isa{{\isasymTHEORY}\ A\ {\isasymIMPORTS}\ B\isactrlsub {\isadigit{1}}\ {\isasymdots}\ B\isactrlsub n\ {\isasymBEGIN}} ensures that the - sub-graph of the collective imports \isa{B\isactrlsub {\isadigit{1}}\ {\isasymdots}\ B\isactrlsub n} - is up-to-date, too. Earlier theories are reloaded as required, with - \isa{update} actions proceeding in topological order according to - theory dependencies. There may be also a wave of implied \isa{outdate} actions for derived theory nodes until a stable situation - is achieved eventually.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{theory}\verb|theory: string -> theory| \\ - \indexml{use\_thy}\verb|use_thy: string -> unit| \\ - \indexml{use\_thys}\verb|use_thys: string list -> unit| \\ - \indexml{ThyInfo.touch\_thy}\verb|ThyInfo.touch_thy: string -> unit| \\ - \indexml{ThyInfo.remove\_thy}\verb|ThyInfo.remove_thy: string -> unit| \\[1ex] - \indexml{ThyInfo.begin\_theory}\verb|ThyInfo.begin_theory|\verb|: ... -> bool -> theory| \\ - \indexml{ThyInfo.end\_theory}\verb|ThyInfo.end_theory: theory -> unit| \\ - \indexml{ThyInfo.register\_theory}\verb|ThyInfo.register_theory: theory -> unit| \\[1ex] - \verb|datatype action = Update |\verb,|,\verb| Outdate |\verb,|,\verb| Remove| \\ - \indexml{ThyInfo.add\_hook}\verb|ThyInfo.add_hook: (ThyInfo.action -> string -> unit) -> unit| \\ - \end{mldecls} - - \begin{description} - - \item \verb|theory|~\isa{A} retrieves the theory value presently - associated with name \isa{A}. Note that the result might be - outdated. - - \item \verb|use_thy|~\isa{A} ensures that theory \isa{A} is fully - up-to-date wrt.\ the external file store, reloading outdated - ancestors as required. - - \item \verb|use_thys| is similar to \verb|use_thy|, but handles - several theories simultaneously. Thus it acts like processing the - import header of a theory, without performing the merge of the - result, though. - - \item \verb|ThyInfo.touch_thy|~\isa{A} performs and \isa{outdate} action - on theory \isa{A} and all descendants. - - \item \verb|ThyInfo.remove_thy|~\isa{A} deletes theory \isa{A} and all - descendants from the theory database. - - \item \verb|ThyInfo.begin_theory| is the basic operation behind a - \isa{{\isasymTHEORY}} header declaration. This is {\ML} functions is - normally not invoked directly. - - \item \verb|ThyInfo.end_theory| concludes the loading of a theory - proper and stores the result in the theory database. - - \item \verb|ThyInfo.register_theory|~\isa{text\ thy} registers an - existing theory value with the theory loader database. There is no - management of associated sources. - - \item \verb|ThyInfo.add_hook|~\isa{f} registers function \isa{f} as a hook for theory database actions. The function will be - invoked with the action and theory name being involved; thus derived - actions may be performed in associated system components, e.g.\ - maintaining the state of an editor for the theory sources. - - The kind and order of actions occurring in practice depends both on - user interactions and the internal process of resolving theory - imports. Hooks should not rely on a particular policy here! Any - exceptions raised by the hook are ignored. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/document/isar.tex --- a/doc-src/IsarImplementation/Thy/document/isar.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,91 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{isar}% -% -\isadelimtheory -\isanewline -\isanewline -\isanewline -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ isar\ \isakeyword{imports}\ base\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Isar proof texts% -} -\isamarkuptrue% -% -\isamarkupsection{Proof context% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -FIXME% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Proof state \label{sec:isar-proof-state}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -FIXME - -\glossary{Proof state}{The whole configuration of a structured proof, -consisting of a \seeglossary{proof context} and an optional -\seeglossary{structured goal}. Internally, an Isar proof state is -organized as a stack to accomodate block structure of proof texts. -For historical reasons, a low-level \seeglossary{tactical goal} is -occasionally called ``proof state'' as well.} - -\glossary{Structured goal}{FIXME} - -\glossary{Goal}{See \seeglossary{tactical goal} or \seeglossary{structured goal}. \norefpage}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Proof methods% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -FIXME% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Attributes% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -FIXME ?!% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/document/locale.tex --- a/doc-src/IsarImplementation/Thy/document/locale.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,73 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{locale}% -% -\isadelimtheory -\isanewline -\isanewline -\isanewline -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ {\isachardoublequoteopen}locale{\isachardoublequoteclose}\ \isakeyword{imports}\ base\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Structured specifications% -} -\isamarkuptrue% -% -\isamarkupsection{Specification elements% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -FIXME% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Type-inference% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -FIXME% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Local theories% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -FIXME - - \glossary{Local theory}{FIXME}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/document/logic.tex --- a/doc-src/IsarImplementation/Thy/document/logic.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,886 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{logic}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ logic\ \isakeyword{imports}\ base\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Primitive logic \label{ch:logic}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The logical foundations of Isabelle/Isar are that of the Pure logic, - which has been introduced as a natural-deduction framework in - \cite{paulson700}. This is essentially the same logic as ``\isa{{\isasymlambda}HOL}'' in the more abstract setting of Pure Type Systems (PTS) - \cite{Barendregt-Geuvers:2001}, although there are some key - differences in the specific treatment of simple types in - Isabelle/Pure. - - Following type-theoretic parlance, the Pure logic consists of three - levels of \isa{{\isasymlambda}}-calculus with corresponding arrows, \isa{{\isasymRightarrow}} for syntactic function space (terms depending on terms), \isa{{\isasymAnd}} for universal quantification (proofs depending on terms), and - \isa{{\isasymLongrightarrow}} for implication (proofs depending on proofs). - - Derivations are relative to a logical theory, which declares type - constructors, constants, and axioms. Theory declarations support - schematic polymorphism, which is strictly speaking outside the - logic.\footnote{This is the deeper logical reason, why the theory - context \isa{{\isasymTheta}} is separate from the proof context \isa{{\isasymGamma}} - of the core calculus.}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Types \label{sec:types}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The language of types is an uninterpreted order-sorted first-order - algebra; types are qualified by ordered type classes. - - \medskip A \emph{type class} is an abstract syntactic entity - declared in the theory context. The \emph{subclass relation} \isa{c\isactrlisub {\isadigit{1}}\ {\isasymsubseteq}\ c\isactrlisub {\isadigit{2}}} is specified by stating an acyclic - generating relation; the transitive closure is maintained - internally. The resulting relation is an ordering: reflexive, - transitive, and antisymmetric. - - A \emph{sort} is a list of type classes written as \isa{s\ {\isacharequal}\ {\isacharbraceleft}c\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ c\isactrlisub m{\isacharbraceright}}, which represents symbolic - intersection. Notationally, the curly braces are omitted for - singleton intersections, i.e.\ any class \isa{c} may be read as - a sort \isa{{\isacharbraceleft}c{\isacharbraceright}}. The ordering on type classes is extended to - sorts according to the meaning of intersections: \isa{{\isacharbraceleft}c\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}\ c\isactrlisub m{\isacharbraceright}\ {\isasymsubseteq}\ {\isacharbraceleft}d\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ d\isactrlisub n{\isacharbraceright}} iff - \isa{{\isasymforall}j{\isachardot}\ {\isasymexists}i{\isachardot}\ c\isactrlisub i\ {\isasymsubseteq}\ d\isactrlisub j}. The empty intersection - \isa{{\isacharbraceleft}{\isacharbraceright}} refers to the universal sort, which is the largest - element wrt.\ the sort order. The intersections of all (finitely - many) classes declared in the current theory are the minimal - elements wrt.\ the sort order. - - \medskip A \emph{fixed type variable} is a pair of a basic name - (starting with a \isa{{\isacharprime}} character) and a sort constraint, e.g.\ - \isa{{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ s{\isacharparenright}} which is usually printed as \isa{{\isasymalpha}\isactrlisub s}. - A \emph{schematic type variable} is a pair of an indexname and a - sort constraint, e.g.\ \isa{{\isacharparenleft}{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ s{\isacharparenright}} which is usually - printed as \isa{{\isacharquery}{\isasymalpha}\isactrlisub s}. - - Note that \emph{all} syntactic components contribute to the identity - of type variables, including the sort constraint. The core logic - handles type variables with the same name but different sorts as - different, although some outer layers of the system make it hard to - produce anything like this. - - A \emph{type constructor} \isa{{\isasymkappa}} is a \isa{k}-ary operator - on types declared in the theory. Type constructor application is - written postfix as \isa{{\isacharparenleft}{\isasymalpha}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlisub k{\isacharparenright}{\isasymkappa}}. For - \isa{k\ {\isacharequal}\ {\isadigit{0}}} the argument tuple is omitted, e.g.\ \isa{prop} - instead of \isa{{\isacharparenleft}{\isacharparenright}prop}. For \isa{k\ {\isacharequal}\ {\isadigit{1}}} the parentheses - are omitted, e.g.\ \isa{{\isasymalpha}\ list} instead of \isa{{\isacharparenleft}{\isasymalpha}{\isacharparenright}list}. - Further notation is provided for specific constructors, notably the - right-associative infix \isa{{\isasymalpha}\ {\isasymRightarrow}\ {\isasymbeta}} instead of \isa{{\isacharparenleft}{\isasymalpha}{\isacharcomma}\ {\isasymbeta}{\isacharparenright}fun}. - - A \emph{type} is defined inductively over type variables and type - constructors as follows: \isa{{\isasymtau}\ {\isacharequal}\ {\isasymalpha}\isactrlisub s\ {\isacharbar}\ {\isacharquery}{\isasymalpha}\isactrlisub s\ {\isacharbar}\ {\isacharparenleft}{\isasymtau}\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlsub k{\isacharparenright}{\isasymkappa}}. - - A \emph{type abbreviation} is a syntactic definition \isa{{\isacharparenleft}\isactrlvec {\isasymalpha}{\isacharparenright}{\isasymkappa}\ {\isacharequal}\ {\isasymtau}} of an arbitrary type expression \isa{{\isasymtau}} over - variables \isa{\isactrlvec {\isasymalpha}}. Type abbreviations appear as type - constructors in the syntax, but are expanded before entering the - logical core. - - A \emph{type arity} declares the image behavior of a type - constructor wrt.\ the algebra of sorts: \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}s\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ s\isactrlisub k{\isacharparenright}s} means that \isa{{\isacharparenleft}{\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlisub k{\isacharparenright}{\isasymkappa}} is - of sort \isa{s} if every argument type \isa{{\isasymtau}\isactrlisub i} is - of sort \isa{s\isactrlisub i}. Arity declarations are implicitly - completed, i.e.\ \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}\isactrlvec s{\isacharparenright}c} entails \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}\isactrlvec s{\isacharparenright}c{\isacharprime}} for any \isa{c{\isacharprime}\ {\isasymsupseteq}\ c}. - - \medskip The sort algebra is always maintained as \emph{coregular}, - which means that type arities are consistent with the subclass - relation: for any type constructor \isa{{\isasymkappa}}, and classes \isa{c\isactrlisub {\isadigit{1}}\ {\isasymsubseteq}\ c\isactrlisub {\isadigit{2}}}, and arities \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}\isactrlvec s\isactrlisub {\isadigit{1}}{\isacharparenright}c\isactrlisub {\isadigit{1}}} and \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}\isactrlvec s\isactrlisub {\isadigit{2}}{\isacharparenright}c\isactrlisub {\isadigit{2}}} holds \isa{\isactrlvec s\isactrlisub {\isadigit{1}}\ {\isasymsubseteq}\ \isactrlvec s\isactrlisub {\isadigit{2}}} component-wise. - - The key property of a coregular order-sorted algebra is that sort - constraints can be solved in a most general fashion: for each type - constructor \isa{{\isasymkappa}} and sort \isa{s} there is a most general - vector of argument sorts \isa{{\isacharparenleft}s\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ s\isactrlisub k{\isacharparenright}} such - that a type scheme \isa{{\isacharparenleft}{\isasymalpha}\isactrlbsub s\isactrlisub {\isadigit{1}}\isactrlesub {\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlbsub s\isactrlisub k\isactrlesub {\isacharparenright}{\isasymkappa}} is of sort \isa{s}. - Consequently, type unification has most general solutions (modulo - equivalence of sorts), so type-inference produces primary types as - expected \cite{nipkow-prehofer}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexmltype{class}\verb|type class| \\ - \indexmltype{sort}\verb|type sort| \\ - \indexmltype{arity}\verb|type arity| \\ - \indexmltype{typ}\verb|type typ| \\ - \indexml{map\_atyps}\verb|map_atyps: (typ -> typ) -> typ -> typ| \\ - \indexml{fold\_atyps}\verb|fold_atyps: (typ -> 'a -> 'a) -> typ -> 'a -> 'a| \\ - \end{mldecls} - \begin{mldecls} - \indexml{Sign.subsort}\verb|Sign.subsort: theory -> sort * sort -> bool| \\ - \indexml{Sign.of\_sort}\verb|Sign.of_sort: theory -> typ * sort -> bool| \\ - \indexml{Sign.add\_types}\verb|Sign.add_types: (string * int * mixfix) list -> theory -> theory| \\ - \indexml{Sign.add\_tyabbrs\_i}\verb|Sign.add_tyabbrs_i: |\isasep\isanewline% -\verb| (string * string list * typ * mixfix) list -> theory -> theory| \\ - \indexml{Sign.primitive\_class}\verb|Sign.primitive_class: string * class list -> theory -> theory| \\ - \indexml{Sign.primitive\_classrel}\verb|Sign.primitive_classrel: class * class -> theory -> theory| \\ - \indexml{Sign.primitive\_arity}\verb|Sign.primitive_arity: arity -> theory -> theory| \\ - \end{mldecls} - - \begin{description} - - \item \verb|class| represents type classes; this is an alias for - \verb|string|. - - \item \verb|sort| represents sorts; this is an alias for - \verb|class list|. - - \item \verb|arity| represents type arities; this is an alias for - triples of the form \isa{{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec s{\isacharcomma}\ s{\isacharparenright}} for \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}\isactrlvec s{\isacharparenright}s} described above. - - \item \verb|typ| represents types; this is a datatype with - constructors \verb|TFree|, \verb|TVar|, \verb|Type|. - - \item \verb|map_atyps|~\isa{f\ {\isasymtau}} applies the mapping \isa{f} - to all atomic types (\verb|TFree|, \verb|TVar|) occurring in \isa{{\isasymtau}}. - - \item \verb|fold_atyps|~\isa{f\ {\isasymtau}} iterates the operation \isa{f} over all occurrences of atomic types (\verb|TFree|, \verb|TVar|) - in \isa{{\isasymtau}}; the type structure is traversed from left to right. - - \item \verb|Sign.subsort|~\isa{thy\ {\isacharparenleft}s\isactrlisub {\isadigit{1}}{\isacharcomma}\ s\isactrlisub {\isadigit{2}}{\isacharparenright}} - tests the subsort relation \isa{s\isactrlisub {\isadigit{1}}\ {\isasymsubseteq}\ s\isactrlisub {\isadigit{2}}}. - - \item \verb|Sign.of_sort|~\isa{thy\ {\isacharparenleft}{\isasymtau}{\isacharcomma}\ s{\isacharparenright}} tests whether type - \isa{{\isasymtau}} is of sort \isa{s}. - - \item \verb|Sign.add_types|~\isa{{\isacharbrackleft}{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ k{\isacharcomma}\ mx{\isacharparenright}{\isacharcomma}\ {\isasymdots}{\isacharbrackright}} declares a new - type constructors \isa{{\isasymkappa}} with \isa{k} arguments and - optional mixfix syntax. - - \item \verb|Sign.add_tyabbrs_i|~\isa{{\isacharbrackleft}{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec {\isasymalpha}{\isacharcomma}\ {\isasymtau}{\isacharcomma}\ mx{\isacharparenright}{\isacharcomma}\ {\isasymdots}{\isacharbrackright}} - defines a new type abbreviation \isa{{\isacharparenleft}\isactrlvec {\isasymalpha}{\isacharparenright}{\isasymkappa}\ {\isacharequal}\ {\isasymtau}} with - optional mixfix syntax. - - \item \verb|Sign.primitive_class|~\isa{{\isacharparenleft}c{\isacharcomma}\ {\isacharbrackleft}c\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ c\isactrlisub n{\isacharbrackright}{\isacharparenright}} declares a new class \isa{c}, together with class - relations \isa{c\ {\isasymsubseteq}\ c\isactrlisub i}, for \isa{i\ {\isacharequal}\ {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ n}. - - \item \verb|Sign.primitive_classrel|~\isa{{\isacharparenleft}c\isactrlisub {\isadigit{1}}{\isacharcomma}\ c\isactrlisub {\isadigit{2}}{\isacharparenright}} declares the class relation \isa{c\isactrlisub {\isadigit{1}}\ {\isasymsubseteq}\ c\isactrlisub {\isadigit{2}}}. - - \item \verb|Sign.primitive_arity|~\isa{{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec s{\isacharcomma}\ s{\isacharparenright}} declares - the arity \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}\isactrlvec s{\isacharparenright}s}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsection{Terms \label{sec:terms}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\glossary{Term}{FIXME} - - The language of terms is that of simply-typed \isa{{\isasymlambda}}-calculus - with de-Bruijn indices for bound variables (cf.\ \cite{debruijn72} - or \cite{paulson-ml2}), with the types being determined determined - by the corresponding binders. In contrast, free variables and - constants are have an explicit name and type in each occurrence. - - \medskip A \emph{bound variable} is a natural number \isa{b}, - which accounts for the number of intermediate binders between the - variable occurrence in the body and its binding position. For - example, the de-Bruijn term \isa{{\isasymlambda}\isactrlbsub nat\isactrlesub {\isachardot}\ {\isasymlambda}\isactrlbsub nat\isactrlesub {\isachardot}\ {\isadigit{1}}\ {\isacharplus}\ {\isadigit{0}}} would - correspond to \isa{{\isasymlambda}x\isactrlbsub nat\isactrlesub {\isachardot}\ {\isasymlambda}y\isactrlbsub nat\isactrlesub {\isachardot}\ x\ {\isacharplus}\ y} in a named - representation. Note that a bound variable may be represented by - different de-Bruijn indices at different occurrences, depending on - the nesting of abstractions. - - A \emph{loose variable} is a bound variable that is outside the - scope of local binders. The types (and names) for loose variables - can be managed as a separate context, that is maintained as a stack - of hypothetical binders. The core logic operates on closed terms, - without any loose variables. - - A \emph{fixed variable} is a pair of a basic name and a type, e.g.\ - \isa{{\isacharparenleft}x{\isacharcomma}\ {\isasymtau}{\isacharparenright}} which is usually printed \isa{x\isactrlisub {\isasymtau}}. A - \emph{schematic variable} is a pair of an indexname and a type, - e.g.\ \isa{{\isacharparenleft}{\isacharparenleft}x{\isacharcomma}\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ {\isasymtau}{\isacharparenright}} which is usually printed as \isa{{\isacharquery}x\isactrlisub {\isasymtau}}. - - \medskip A \emph{constant} is a pair of a basic name and a type, - e.g.\ \isa{{\isacharparenleft}c{\isacharcomma}\ {\isasymtau}{\isacharparenright}} which is usually printed as \isa{c\isactrlisub {\isasymtau}}. Constants are declared in the context as polymorphic - families \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}}, meaning that all substitution instances - \isa{c\isactrlisub {\isasymtau}} for \isa{{\isasymtau}\ {\isacharequal}\ {\isasymsigma}{\isasymvartheta}} are valid. - - The vector of \emph{type arguments} of constant \isa{c\isactrlisub {\isasymtau}} - wrt.\ the declaration \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} is defined as the codomain of - the matcher \isa{{\isasymvartheta}\ {\isacharequal}\ {\isacharbraceleft}{\isacharquery}{\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymmapsto}\ {\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isacharquery}{\isasymalpha}\isactrlisub n\ {\isasymmapsto}\ {\isasymtau}\isactrlisub n{\isacharbraceright}} presented in canonical order \isa{{\isacharparenleft}{\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlisub n{\isacharparenright}}. Within a given theory context, - there is a one-to-one correspondence between any constant \isa{c\isactrlisub {\isasymtau}} and the application \isa{c{\isacharparenleft}{\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlisub n{\isacharparenright}} of its type arguments. For example, with \isa{plus\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}}, the instance \isa{plus\isactrlbsub nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat\isactrlesub } corresponds to \isa{plus{\isacharparenleft}nat{\isacharparenright}}. - - Constant declarations \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} may contain sort constraints - for type variables in \isa{{\isasymsigma}}. These are observed by - type-inference as expected, but \emph{ignored} by the core logic. - This means the primitive logic is able to reason with instances of - polymorphic constants that the user-level type-checker would reject - due to violation of type class restrictions. - - \medskip An \emph{atomic} term is either a variable or constant. A - \emph{term} is defined inductively over atomic terms, with - abstraction and application as follows: \isa{t\ {\isacharequal}\ b\ {\isacharbar}\ x\isactrlisub {\isasymtau}\ {\isacharbar}\ {\isacharquery}x\isactrlisub {\isasymtau}\ {\isacharbar}\ c\isactrlisub {\isasymtau}\ {\isacharbar}\ {\isasymlambda}\isactrlisub {\isasymtau}{\isachardot}\ t\ {\isacharbar}\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}}. - Parsing and printing takes care of converting between an external - representation with named bound variables. Subsequently, we shall - use the latter notation instead of internal de-Bruijn - representation. - - The inductive relation \isa{t\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}} assigns a (unique) type to a - term according to the structure of atomic terms, abstractions, and - applicatins: - \[ - \infer{\isa{a\isactrlisub {\isasymtau}\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}}}{} - \qquad - \infer{\isa{{\isacharparenleft}{\isasymlambda}x\isactrlsub {\isasymtau}{\isachardot}\ t{\isacharparenright}\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymsigma}}}{\isa{t\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}}} - \qquad - \infer{\isa{t\ u\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}}}{\isa{t\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymsigma}} & \isa{u\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}}} - \] - A \emph{well-typed term} is a term that can be typed according to these rules. - - Typing information can be omitted: type-inference is able to - reconstruct the most general type of a raw term, while assigning - most general types to all of its variables and constants. - Type-inference depends on a context of type constraints for fixed - variables, and declarations for polymorphic constants. - - The identity of atomic terms consists both of the name and the type - component. This means that different variables \isa{x\isactrlbsub {\isasymtau}\isactrlisub {\isadigit{1}}\isactrlesub } and \isa{x\isactrlbsub {\isasymtau}\isactrlisub {\isadigit{2}}\isactrlesub } may become the same after type - instantiation. Some outer layers of the system make it hard to - produce variables of the same name, but different types. In - contrast, mixed instances of polymorphic constants occur frequently. - - \medskip The \emph{hidden polymorphism} of a term \isa{t\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} - is the set of type variables occurring in \isa{t}, but not in - \isa{{\isasymsigma}}. This means that the term implicitly depends on type - arguments that are not accounted in the result type, i.e.\ there are - different type instances \isa{t{\isasymvartheta}\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} and \isa{t{\isasymvartheta}{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} with the same type. This slightly - pathological situation notoriously demands additional care. - - \medskip A \emph{term abbreviation} is a syntactic definition \isa{c\isactrlisub {\isasymsigma}\ {\isasymequiv}\ t} of a closed term \isa{t} of type \isa{{\isasymsigma}}, - without any hidden polymorphism. A term abbreviation looks like a - constant in the syntax, but is expanded before entering the logical - core. Abbreviations are usually reverted when printing terms, using - \isa{t\ {\isasymrightarrow}\ c\isactrlisub {\isasymsigma}} as rules for higher-order rewriting. - - \medskip Canonical operations on \isa{{\isasymlambda}}-terms include \isa{{\isasymalpha}{\isasymbeta}{\isasymeta}}-conversion: \isa{{\isasymalpha}}-conversion refers to capture-free - renaming of bound variables; \isa{{\isasymbeta}}-conversion contracts an - abstraction applied to an argument term, substituting the argument - in the body: \isa{{\isacharparenleft}{\isasymlambda}x{\isachardot}\ b{\isacharparenright}a} becomes \isa{b{\isacharbrackleft}a{\isacharslash}x{\isacharbrackright}}; \isa{{\isasymeta}}-conversion contracts vacuous application-abstraction: \isa{{\isasymlambda}x{\isachardot}\ f\ x} becomes \isa{f}, provided that the bound variable - does not occur in \isa{f}. - - Terms are normally treated modulo \isa{{\isasymalpha}}-conversion, which is - implicit in the de-Bruijn representation. Names for bound variables - in abstractions are maintained separately as (meaningless) comments, - mostly for parsing and printing. Full \isa{{\isasymalpha}{\isasymbeta}{\isasymeta}}-conversion is - commonplace in various standard operations (\secref{sec:obj-rules}) - that are based on higher-order unification and matching.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexmltype{term}\verb|type term| \\ - \indexml{op aconv}\verb|op aconv: term * term -> bool| \\ - \indexml{map\_types}\verb|map_types: (typ -> typ) -> term -> term| \\ - \indexml{fold\_types}\verb|fold_types: (typ -> 'a -> 'a) -> term -> 'a -> 'a| \\ - \indexml{map\_aterms}\verb|map_aterms: (term -> term) -> term -> term| \\ - \indexml{fold\_aterms}\verb|fold_aterms: (term -> 'a -> 'a) -> term -> 'a -> 'a| \\ - \end{mldecls} - \begin{mldecls} - \indexml{fastype\_of}\verb|fastype_of: term -> typ| \\ - \indexml{lambda}\verb|lambda: term -> term -> term| \\ - \indexml{betapply}\verb|betapply: term * term -> term| \\ - \indexml{Sign.declare\_const}\verb|Sign.declare_const: Properties.T -> (binding * typ) * mixfix ->|\isasep\isanewline% -\verb| theory -> term * theory| \\ - \indexml{Sign.add\_abbrev}\verb|Sign.add_abbrev: string -> Properties.T -> binding * term ->|\isasep\isanewline% -\verb| theory -> (term * term) * theory| \\ - \indexml{Sign.const\_typargs}\verb|Sign.const_typargs: theory -> string * typ -> typ list| \\ - \indexml{Sign.const\_instance}\verb|Sign.const_instance: theory -> string * typ list -> typ| \\ - \end{mldecls} - - \begin{description} - - \item \verb|term| represents de-Bruijn terms, with comments in - abstractions, and explicitly named free variables and constants; - this is a datatype with constructors \verb|Bound|, \verb|Free|, \verb|Var|, \verb|Const|, \verb|Abs|, \verb|op $|. - - \item \isa{t}~\verb|aconv|~\isa{u} checks \isa{{\isasymalpha}}-equivalence of two terms. This is the basic equality relation - on type \verb|term|; raw datatype equality should only be used - for operations related to parsing or printing! - - \item \verb|map_types|~\isa{f\ t} applies the mapping \isa{f} to all types occurring in \isa{t}. - - \item \verb|fold_types|~\isa{f\ t} iterates the operation \isa{f} over all occurrences of types in \isa{t}; the term - structure is traversed from left to right. - - \item \verb|map_aterms|~\isa{f\ t} applies the mapping \isa{f} - to all atomic terms (\verb|Bound|, \verb|Free|, \verb|Var|, \verb|Const|) occurring in \isa{t}. - - \item \verb|fold_aterms|~\isa{f\ t} iterates the operation \isa{f} over all occurrences of atomic terms (\verb|Bound|, \verb|Free|, - \verb|Var|, \verb|Const|) in \isa{t}; the term structure is - traversed from left to right. - - \item \verb|fastype_of|~\isa{t} determines the type of a - well-typed term. This operation is relatively slow, despite the - omission of any sanity checks. - - \item \verb|lambda|~\isa{a\ b} produces an abstraction \isa{{\isasymlambda}a{\isachardot}\ b}, where occurrences of the atomic term \isa{a} in the - body \isa{b} are replaced by bound variables. - - \item \verb|betapply|~\isa{{\isacharparenleft}t{\isacharcomma}\ u{\isacharparenright}} produces an application \isa{t\ u}, with topmost \isa{{\isasymbeta}}-conversion if \isa{t} is an - abstraction. - - \item \verb|Sign.declare_const|~\isa{properties\ {\isacharparenleft}{\isacharparenleft}c{\isacharcomma}\ {\isasymsigma}{\isacharparenright}{\isacharcomma}\ mx{\isacharparenright}} - declares a new constant \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} with optional mixfix - syntax. - - \item \verb|Sign.add_abbrev|~\isa{print{\isacharunderscore}mode\ properties\ {\isacharparenleft}c{\isacharcomma}\ t{\isacharparenright}} - introduces a new term abbreviation \isa{c\ {\isasymequiv}\ t}. - - \item \verb|Sign.const_typargs|~\isa{thy\ {\isacharparenleft}c{\isacharcomma}\ {\isasymtau}{\isacharparenright}} and \verb|Sign.const_instance|~\isa{thy\ {\isacharparenleft}c{\isacharcomma}\ {\isacharbrackleft}{\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlisub n{\isacharbrackright}{\isacharparenright}} - convert between two representations of polymorphic constants: full - type instance vs.\ compact type arguments form. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsection{Theorems \label{sec:thms}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\glossary{Proposition}{FIXME A \seeglossary{term} of - \seeglossary{type} \isa{prop}. Internally, there is nothing - special about propositions apart from their type, but the concrete - syntax enforces a clear distinction. Propositions are structured - via implication \isa{A\ {\isasymLongrightarrow}\ B} or universal quantification \isa{{\isasymAnd}x{\isachardot}\ B\ x} --- anything else is considered atomic. The canonical - form for propositions is that of a \seeglossary{Hereditary Harrop - Formula}. FIXME} - - \glossary{Theorem}{A proven proposition within a certain theory and - proof context, formally \isa{{\isasymGamma}\ {\isasymturnstile}\isactrlsub {\isasymTheta}\ {\isasymphi}}; both contexts are - rarely spelled out explicitly. Theorems are usually normalized - according to the \seeglossary{HHF} format. FIXME} - - \glossary{Fact}{Sometimes used interchangeably for - \seeglossary{theorem}. Strictly speaking, a list of theorems, - essentially an extra-logical conjunction. Facts emerge either as - local assumptions, or as results of local goal statements --- both - may be simultaneous, hence the list representation. FIXME} - - \glossary{Schematic variable}{FIXME} - - \glossary{Fixed variable}{A variable that is bound within a certain - proof context; an arbitrary-but-fixed entity within a portion of - proof text. FIXME} - - \glossary{Free variable}{Synonymous for \seeglossary{fixed - variable}. FIXME} - - \glossary{Bound variable}{FIXME} - - \glossary{Variable}{See \seeglossary{schematic variable}, - \seeglossary{fixed variable}, \seeglossary{bound variable}, or - \seeglossary{type variable}. The distinguishing feature of - different variables is their binding scope. FIXME} - - A \emph{proposition} is a well-typed term of type \isa{prop}, a - \emph{theorem} is a proven proposition (depending on a context of - hypotheses and the background theory). Primitive inferences include - plain natural deduction rules for the primary connectives \isa{{\isasymAnd}} and \isa{{\isasymLongrightarrow}} of the framework. There is also a builtin - notion of equality/equivalence \isa{{\isasymequiv}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Primitive connectives and rules \label{sec:prim-rules}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The theory \isa{Pure} contains constant declarations for the - primitive connectives \isa{{\isasymAnd}}, \isa{{\isasymLongrightarrow}}, and \isa{{\isasymequiv}} of - the logical framework, see \figref{fig:pure-connectives}. The - derivability judgment \isa{A\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ A\isactrlisub n\ {\isasymturnstile}\ B} is - defined inductively by the primitive inferences given in - \figref{fig:prim-rules}, with the global restriction that the - hypotheses must \emph{not} contain any schematic variables. The - builtin equality is conceptually axiomatized as shown in - \figref{fig:pure-equality}, although the implementation works - directly with derived inferences. - - \begin{figure}[htb] - \begin{center} - \begin{tabular}{ll} - \isa{all\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}{\isasymalpha}\ {\isasymRightarrow}\ prop{\isacharparenright}\ {\isasymRightarrow}\ prop} & universal quantification (binder \isa{{\isasymAnd}}) \\ - \isa{{\isasymLongrightarrow}\ {\isacharcolon}{\isacharcolon}\ prop\ {\isasymRightarrow}\ prop\ {\isasymRightarrow}\ prop} & implication (right associative infix) \\ - \isa{{\isasymequiv}\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ prop} & equality relation (infix) \\ - \end{tabular} - \caption{Primitive connectives of Pure}\label{fig:pure-connectives} - \end{center} - \end{figure} - - \begin{figure}[htb] - \begin{center} - \[ - \infer[\isa{{\isacharparenleft}axiom{\isacharparenright}}]{\isa{{\isasymturnstile}\ A}}{\isa{A\ {\isasymin}\ {\isasymTheta}}} - \qquad - \infer[\isa{{\isacharparenleft}assume{\isacharparenright}}]{\isa{A\ {\isasymturnstile}\ A}}{} - \] - \[ - \infer[\isa{{\isacharparenleft}{\isasymAnd}{\isacharunderscore}intro{\isacharparenright}}]{\isa{{\isasymGamma}\ {\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ b{\isacharbrackleft}x{\isacharbrackright}}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ b{\isacharbrackleft}x{\isacharbrackright}} & \isa{x\ {\isasymnotin}\ {\isasymGamma}}} - \qquad - \infer[\isa{{\isacharparenleft}{\isasymAnd}{\isacharunderscore}elim{\isacharparenright}}]{\isa{{\isasymGamma}\ {\isasymturnstile}\ b{\isacharbrackleft}a{\isacharbrackright}}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ b{\isacharbrackleft}x{\isacharbrackright}}} - \] - \[ - \infer[\isa{{\isacharparenleft}{\isasymLongrightarrow}{\isacharunderscore}intro{\isacharparenright}}]{\isa{{\isasymGamma}\ {\isacharminus}\ A\ {\isasymturnstile}\ A\ {\isasymLongrightarrow}\ B}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B}} - \qquad - \infer[\isa{{\isacharparenleft}{\isasymLongrightarrow}{\isacharunderscore}elim{\isacharparenright}}]{\isa{{\isasymGamma}\isactrlsub {\isadigit{1}}\ {\isasymunion}\ {\isasymGamma}\isactrlsub {\isadigit{2}}\ {\isasymturnstile}\ B}}{\isa{{\isasymGamma}\isactrlsub {\isadigit{1}}\ {\isasymturnstile}\ A\ {\isasymLongrightarrow}\ B} & \isa{{\isasymGamma}\isactrlsub {\isadigit{2}}\ {\isasymturnstile}\ A}} - \] - \caption{Primitive inferences of Pure}\label{fig:prim-rules} - \end{center} - \end{figure} - - \begin{figure}[htb] - \begin{center} - \begin{tabular}{ll} - \isa{{\isasymturnstile}\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ b{\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}\ a\ {\isasymequiv}\ b{\isacharbrackleft}a{\isacharbrackright}} & \isa{{\isasymbeta}}-conversion \\ - \isa{{\isasymturnstile}\ x\ {\isasymequiv}\ x} & reflexivity \\ - \isa{{\isasymturnstile}\ x\ {\isasymequiv}\ y\ {\isasymLongrightarrow}\ P\ x\ {\isasymLongrightarrow}\ P\ y} & substitution \\ - \isa{{\isasymturnstile}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ f\ x\ {\isasymequiv}\ g\ x{\isacharparenright}\ {\isasymLongrightarrow}\ f\ {\isasymequiv}\ g} & extensionality \\ - \isa{{\isasymturnstile}\ {\isacharparenleft}A\ {\isasymLongrightarrow}\ B{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}B\ {\isasymLongrightarrow}\ A{\isacharparenright}\ {\isasymLongrightarrow}\ A\ {\isasymequiv}\ B} & logical equivalence \\ - \end{tabular} - \caption{Conceptual axiomatization of Pure equality}\label{fig:pure-equality} - \end{center} - \end{figure} - - The introduction and elimination rules for \isa{{\isasymAnd}} and \isa{{\isasymLongrightarrow}} are analogous to formation of dependently typed \isa{{\isasymlambda}}-terms representing the underlying proof objects. Proof terms - are irrelevant in the Pure logic, though; they cannot occur within - propositions. The system provides a runtime option to record - explicit proof terms for primitive inferences. Thus all three - levels of \isa{{\isasymlambda}}-calculus become explicit: \isa{{\isasymRightarrow}} for - terms, and \isa{{\isasymAnd}{\isacharslash}{\isasymLongrightarrow}} for proofs (cf.\ - \cite{Berghofer-Nipkow:2000:TPHOL}). - - Observe that locally fixed parameters (as in \isa{{\isasymAnd}{\isacharunderscore}intro}) need - not be recorded in the hypotheses, because the simple syntactic - types of Pure are always inhabitable. ``Assumptions'' \isa{x\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}} for type-membership are only present as long as some \isa{x\isactrlisub {\isasymtau}} occurs in the statement body.\footnote{This is the key - difference to ``\isa{{\isasymlambda}HOL}'' in the PTS framework - \cite{Barendregt-Geuvers:2001}, where hypotheses \isa{x\ {\isacharcolon}\ A} are - treated uniformly for propositions and types.} - - \medskip The axiomatization of a theory is implicitly closed by - forming all instances of type and term variables: \isa{{\isasymturnstile}\ A{\isasymvartheta}} holds for any substitution instance of an axiom - \isa{{\isasymturnstile}\ A}. By pushing substitutions through derivations - inductively, we also get admissible \isa{generalize} and \isa{instance} rules as shown in \figref{fig:subst-rules}. - - \begin{figure}[htb] - \begin{center} - \[ - \infer{\isa{{\isasymGamma}\ {\isasymturnstile}\ B{\isacharbrackleft}{\isacharquery}{\isasymalpha}{\isacharbrackright}}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B{\isacharbrackleft}{\isasymalpha}{\isacharbrackright}} & \isa{{\isasymalpha}\ {\isasymnotin}\ {\isasymGamma}}} - \quad - \infer[\quad\isa{{\isacharparenleft}generalize{\isacharparenright}}]{\isa{{\isasymGamma}\ {\isasymturnstile}\ B{\isacharbrackleft}{\isacharquery}x{\isacharbrackright}}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B{\isacharbrackleft}x{\isacharbrackright}} & \isa{x\ {\isasymnotin}\ {\isasymGamma}}} - \] - \[ - \infer{\isa{{\isasymGamma}\ {\isasymturnstile}\ B{\isacharbrackleft}{\isasymtau}{\isacharbrackright}}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B{\isacharbrackleft}{\isacharquery}{\isasymalpha}{\isacharbrackright}}} - \quad - \infer[\quad\isa{{\isacharparenleft}instantiate{\isacharparenright}}]{\isa{{\isasymGamma}\ {\isasymturnstile}\ B{\isacharbrackleft}t{\isacharbrackright}}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B{\isacharbrackleft}{\isacharquery}x{\isacharbrackright}}} - \] - \caption{Admissible substitution rules}\label{fig:subst-rules} - \end{center} - \end{figure} - - Note that \isa{instantiate} does not require an explicit - side-condition, because \isa{{\isasymGamma}} may never contain schematic - variables. - - In principle, variables could be substituted in hypotheses as well, - but this would disrupt the monotonicity of reasoning: deriving - \isa{{\isasymGamma}{\isasymvartheta}\ {\isasymturnstile}\ B{\isasymvartheta}} from \isa{{\isasymGamma}\ {\isasymturnstile}\ B} is - correct, but \isa{{\isasymGamma}{\isasymvartheta}\ {\isasymsupseteq}\ {\isasymGamma}} does not necessarily hold: - the result belongs to a different proof context. - - \medskip An \emph{oracle} is a function that produces axioms on the - fly. Logically, this is an instance of the \isa{axiom} rule - (\figref{fig:prim-rules}), but there is an operational difference. - The system always records oracle invocations within derivations of - theorems. Tracing plain axioms (and named theorems) is optional. - - Axiomatizations should be limited to the bare minimum, typically as - part of the initial logical basis of an object-logic formalization. - Later on, theories are usually developed in a strictly definitional - fashion, by stating only certain equalities over new constants. - - A \emph{simple definition} consists of a constant declaration \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} together with an axiom \isa{{\isasymturnstile}\ c\ {\isasymequiv}\ t}, where \isa{t\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} is a closed term without any hidden polymorphism. The RHS - may depend on further defined constants, but not \isa{c} itself. - Definitions of functions may be presented as \isa{c\ \isactrlvec x\ {\isasymequiv}\ t} instead of the puristic \isa{c\ {\isasymequiv}\ {\isasymlambda}\isactrlvec x{\isachardot}\ t}. - - An \emph{overloaded definition} consists of a collection of axioms - for the same constant, with zero or one equations \isa{c{\isacharparenleft}{\isacharparenleft}\isactrlvec {\isasymalpha}{\isacharparenright}{\isasymkappa}{\isacharparenright}\ {\isasymequiv}\ t} for each type constructor \isa{{\isasymkappa}} (for - distinct variables \isa{\isactrlvec {\isasymalpha}}). The RHS may mention - previously defined constants as above, or arbitrary constants \isa{d{\isacharparenleft}{\isasymalpha}\isactrlisub i{\isacharparenright}} for some \isa{{\isasymalpha}\isactrlisub i} projected from \isa{\isactrlvec {\isasymalpha}}. Thus overloaded definitions essentially work by - primitive recursion over the syntactic structure of a single type - argument.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexmltype{ctyp}\verb|type ctyp| \\ - \indexmltype{cterm}\verb|type cterm| \\ - \indexml{Thm.ctyp\_of}\verb|Thm.ctyp_of: theory -> typ -> ctyp| \\ - \indexml{Thm.cterm\_of}\verb|Thm.cterm_of: theory -> term -> cterm| \\ - \end{mldecls} - \begin{mldecls} - \indexmltype{thm}\verb|type thm| \\ - \indexml{proofs}\verb|proofs: int ref| \\ - \indexml{Thm.assume}\verb|Thm.assume: cterm -> thm| \\ - \indexml{Thm.forall\_intr}\verb|Thm.forall_intr: cterm -> thm -> thm| \\ - \indexml{Thm.forall\_elim}\verb|Thm.forall_elim: cterm -> thm -> thm| \\ - \indexml{Thm.implies\_intr}\verb|Thm.implies_intr: cterm -> thm -> thm| \\ - \indexml{Thm.implies\_elim}\verb|Thm.implies_elim: thm -> thm -> thm| \\ - \indexml{Thm.generalize}\verb|Thm.generalize: string list * string list -> int -> thm -> thm| \\ - \indexml{Thm.instantiate}\verb|Thm.instantiate: (ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm| \\ - \indexml{Thm.axiom}\verb|Thm.axiom: theory -> string -> thm| \\ - \indexml{Thm.add\_oracle}\verb|Thm.add_oracle: bstring * ('a -> cterm) -> theory|\isasep\isanewline% -\verb| -> (string * ('a -> thm)) * theory| \\ - \end{mldecls} - \begin{mldecls} - \indexml{Theory.add\_axioms\_i}\verb|Theory.add_axioms_i: (binding * term) list -> theory -> theory| \\ - \indexml{Theory.add\_deps}\verb|Theory.add_deps: string -> string * typ -> (string * typ) list -> theory -> theory| \\ - \indexml{Theory.add\_defs\_i}\verb|Theory.add_defs_i: bool -> bool -> (binding * term) list -> theory -> theory| \\ - \end{mldecls} - - \begin{description} - - \item \verb|ctyp| and \verb|cterm| represent certified types - and terms, respectively. These are abstract datatypes that - guarantee that its values have passed the full well-formedness (and - well-typedness) checks, relative to the declarations of type - constructors, constants etc. in the theory. - - \item \verb|ctyp_of|~\isa{thy\ {\isasymtau}} and \verb|cterm_of|~\isa{thy\ t} explicitly checks types and terms, respectively. This also - involves some basic normalizations, such expansion of type and term - abbreviations from the theory context. - - Re-certification is relatively slow and should be avoided in tight - reasoning loops. There are separate operations to decompose - certified entities (including actual theorems). - - \item \verb|thm| represents proven propositions. This is an - abstract datatype that guarantees that its values have been - constructed by basic principles of the \verb|Thm| module. - Every \verb|thm| value contains a sliding back-reference to the - enclosing theory, cf.\ \secref{sec:context-theory}. - - \item \verb|proofs| determines the detail of proof recording within - \verb|thm| values: \verb|0| records only oracles, \verb|1| records - oracles, axioms and named theorems, \verb|2| records full proof - terms. - - \item \verb|Thm.assume|, \verb|Thm.forall_intr|, \verb|Thm.forall_elim|, \verb|Thm.implies_intr|, and \verb|Thm.implies_elim| - correspond to the primitive inferences of \figref{fig:prim-rules}. - - \item \verb|Thm.generalize|~\isa{{\isacharparenleft}\isactrlvec {\isasymalpha}{\isacharcomma}\ \isactrlvec x{\isacharparenright}} - corresponds to the \isa{generalize} rules of - \figref{fig:subst-rules}. Here collections of type and term - variables are generalized simultaneously, specified by the given - basic names. - - \item \verb|Thm.instantiate|~\isa{{\isacharparenleft}\isactrlvec {\isasymalpha}\isactrlisub s{\isacharcomma}\ \isactrlvec x\isactrlisub {\isasymtau}{\isacharparenright}} corresponds to the \isa{instantiate} rules - of \figref{fig:subst-rules}. Type variables are substituted before - term variables. Note that the types in \isa{\isactrlvec x\isactrlisub {\isasymtau}} - refer to the instantiated versions. - - \item \verb|Thm.axiom|~\isa{thy\ name} retrieves a named - axiom, cf.\ \isa{axiom} in \figref{fig:prim-rules}. - - \item \verb|Thm.add_oracle|~\isa{{\isacharparenleft}name{\isacharcomma}\ oracle{\isacharparenright}} produces a named - oracle rule, essentially generating arbitrary axioms on the fly, - cf.\ \isa{axiom} in \figref{fig:prim-rules}. - - \item \verb|Theory.add_axioms_i|~\isa{{\isacharbrackleft}{\isacharparenleft}name{\isacharcomma}\ A{\isacharparenright}{\isacharcomma}\ {\isasymdots}{\isacharbrackright}} declares - arbitrary propositions as axioms. - - \item \verb|Theory.add_deps|~\isa{name\ c\isactrlisub {\isasymtau}\ \isactrlvec d\isactrlisub {\isasymsigma}} declares dependencies of a named specification - for constant \isa{c\isactrlisub {\isasymtau}}, relative to existing - specifications for constants \isa{\isactrlvec d\isactrlisub {\isasymsigma}}. - - \item \verb|Theory.add_defs_i|~\isa{unchecked\ overloaded\ {\isacharbrackleft}{\isacharparenleft}name{\isacharcomma}\ c\ \isactrlvec x\ {\isasymequiv}\ t{\isacharparenright}{\isacharcomma}\ {\isasymdots}{\isacharbrackright}} states a definitional axiom for an existing - constant \isa{c}. Dependencies are recorded (cf.\ \verb|Theory.add_deps|), unless the \isa{unchecked} option is set. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Auxiliary definitions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Theory \isa{Pure} provides a few auxiliary definitions, see - \figref{fig:pure-aux}. These special constants are normally not - exposed to the user, but appear in internal encodings. - - \begin{figure}[htb] - \begin{center} - \begin{tabular}{ll} - \isa{conjunction\ {\isacharcolon}{\isacharcolon}\ prop\ {\isasymRightarrow}\ prop\ {\isasymRightarrow}\ prop} & (infix \isa{{\isacharampersand}}) \\ - \isa{{\isasymturnstile}\ A\ {\isacharampersand}\ B\ {\isasymequiv}\ {\isacharparenleft}{\isasymAnd}C{\isachardot}\ {\isacharparenleft}A\ {\isasymLongrightarrow}\ B\ {\isasymLongrightarrow}\ C{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isacharparenright}} \\[1ex] - \isa{prop\ {\isacharcolon}{\isacharcolon}\ prop\ {\isasymRightarrow}\ prop} & (prefix \isa{{\isacharhash}}, suppressed) \\ - \isa{{\isacharhash}A\ {\isasymequiv}\ A} \\[1ex] - \isa{term\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}\ {\isasymRightarrow}\ prop} & (prefix \isa{TERM}) \\ - \isa{term\ x\ {\isasymequiv}\ {\isacharparenleft}{\isasymAnd}A{\isachardot}\ A\ {\isasymLongrightarrow}\ A{\isacharparenright}} \\[1ex] - \isa{TYPE\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}\ itself} & (prefix \isa{TYPE}) \\ - \isa{{\isacharparenleft}unspecified{\isacharparenright}} \\ - \end{tabular} - \caption{Definitions of auxiliary connectives}\label{fig:pure-aux} - \end{center} - \end{figure} - - Derived conjunction rules include introduction \isa{A\ {\isasymLongrightarrow}\ B\ {\isasymLongrightarrow}\ A\ {\isacharampersand}\ B}, and destructions \isa{A\ {\isacharampersand}\ B\ {\isasymLongrightarrow}\ A} and \isa{A\ {\isacharampersand}\ B\ {\isasymLongrightarrow}\ B}. - Conjunction allows to treat simultaneous assumptions and conclusions - uniformly. For example, multiple claims are intermediately - represented as explicit conjunction, but this is refined into - separate sub-goals before the user continues the proof; the final - result is projected into a list of theorems (cf.\ - \secref{sec:tactical-goals}). - - The \isa{prop} marker (\isa{{\isacharhash}}) makes arbitrarily complex - propositions appear as atomic, without changing the meaning: \isa{{\isasymGamma}\ {\isasymturnstile}\ A} and \isa{{\isasymGamma}\ {\isasymturnstile}\ {\isacharhash}A} are interchangeable. See - \secref{sec:tactical-goals} for specific operations. - - The \isa{term} marker turns any well-typed term into a derivable - proposition: \isa{{\isasymturnstile}\ TERM\ t} holds unconditionally. Although - this is logically vacuous, it allows to treat terms and proofs - uniformly, similar to a type-theoretic framework. - - The \isa{TYPE} constructor is the canonical representative of - the unspecified type \isa{{\isasymalpha}\ itself}; it essentially injects the - language of types into that of terms. There is specific notation - \isa{TYPE{\isacharparenleft}{\isasymtau}{\isacharparenright}} for \isa{TYPE\isactrlbsub {\isasymtau}\ itself\isactrlesub }. - Although being devoid of any particular meaning, the \isa{TYPE{\isacharparenleft}{\isasymtau}{\isacharparenright}} accounts for the type \isa{{\isasymtau}} within the term - language. In particular, \isa{TYPE{\isacharparenleft}{\isasymalpha}{\isacharparenright}} may be used as formal - argument in primitive definitions, in order to circumvent hidden - polymorphism (cf.\ \secref{sec:terms}). For example, \isa{c\ TYPE{\isacharparenleft}{\isasymalpha}{\isacharparenright}\ {\isasymequiv}\ A{\isacharbrackleft}{\isasymalpha}{\isacharbrackright}} defines \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}\ itself\ {\isasymRightarrow}\ prop} in terms of - a proposition \isa{A} that depends on an additional type - argument, which is essentially a predicate on types.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Conjunction.intr}\verb|Conjunction.intr: thm -> thm -> thm| \\ - \indexml{Conjunction.elim}\verb|Conjunction.elim: thm -> thm * thm| \\ - \indexml{Drule.mk\_term}\verb|Drule.mk_term: cterm -> thm| \\ - \indexml{Drule.dest\_term}\verb|Drule.dest_term: thm -> cterm| \\ - \indexml{Logic.mk\_type}\verb|Logic.mk_type: typ -> term| \\ - \indexml{Logic.dest\_type}\verb|Logic.dest_type: term -> typ| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Conjunction.intr| derives \isa{A\ {\isacharampersand}\ B} from \isa{A} and \isa{B}. - - \item \verb|Conjunction.elim| derives \isa{A} and \isa{B} - from \isa{A\ {\isacharampersand}\ B}. - - \item \verb|Drule.mk_term| derives \isa{TERM\ t}. - - \item \verb|Drule.dest_term| recovers term \isa{t} from \isa{TERM\ t}. - - \item \verb|Logic.mk_type|~\isa{{\isasymtau}} produces the term \isa{TYPE{\isacharparenleft}{\isasymtau}{\isacharparenright}}. - - \item \verb|Logic.dest_type|~\isa{TYPE{\isacharparenleft}{\isasymtau}{\isacharparenright}} recovers the type - \isa{{\isasymtau}}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsection{Object-level rules \label{sec:obj-rules}% -} -\isamarkuptrue% -% -\isadelimFIXME -% -\endisadelimFIXME -% -\isatagFIXME -% -\begin{isamarkuptext}% -FIXME - - A \emph{rule} is any Pure theorem in HHF normal form; there is a - separate calculus for rule composition, which is modeled after - Gentzen's Natural Deduction \cite{Gentzen:1935}, but allows - rules to be nested arbitrarily, similar to \cite{extensions91}. - - Normally, all theorems accessible to the user are proper rules. - Low-level inferences are occasional required internally, but the - result should be always presented in canonical form. The higher - interfaces of Isabelle/Isar will always produce proper rules. It is - important to maintain this invariant in add-on applications! - - There are two main principles of rule composition: \isa{resolution} (i.e.\ backchaining of rules) and \isa{by{\isacharminus}assumption} (i.e.\ closing a branch); both principles are - combined in the variants of \isa{elim{\isacharminus}resolution} and \isa{dest{\isacharminus}resolution}. Raw \isa{composition} is occasionally - useful as well, also it is strictly speaking outside of the proper - rule calculus. - - Rules are treated modulo general higher-order unification, which is - unification modulo the equational theory of \isa{{\isasymalpha}{\isasymbeta}{\isasymeta}}-conversion - on \isa{{\isasymlambda}}-terms. Moreover, propositions are understood modulo - the (derived) equivalence \isa{{\isacharparenleft}A\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ B\ x{\isacharparenright}{\isacharparenright}\ {\isasymequiv}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ A\ {\isasymLongrightarrow}\ B\ x{\isacharparenright}}. - - This means that any operations within the rule calculus may be - subject to spontaneous \isa{{\isasymalpha}{\isasymbeta}{\isasymeta}}-HHF conversions. It is common - practice not to contract or expand unnecessarily. Some mechanisms - prefer an one form, others the opposite, so there is a potential - danger to produce some oscillation! - - Only few operations really work \emph{modulo} HHF conversion, but - expect a normal form: quantifiers \isa{{\isasymAnd}} before implications - \isa{{\isasymLongrightarrow}} at each level of nesting. - -\glossary{Hereditary Harrop Formula}{The set of propositions in HHF -format is defined inductively as \isa{H\ {\isacharequal}\ {\isacharparenleft}{\isasymAnd}x\isactrlsup {\isacharasterisk}{\isachardot}\ H\isactrlsup {\isacharasterisk}\ {\isasymLongrightarrow}\ A{\isacharparenright}}, for variables \isa{x} and atomic propositions \isa{A}. -Any proposition may be put into HHF form by normalizing with the rule -\isa{{\isacharparenleft}A\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ B\ x{\isacharparenright}{\isacharparenright}\ {\isasymequiv}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ A\ {\isasymLongrightarrow}\ B\ x{\isacharparenright}}. In Isabelle, the outermost -quantifier prefix is represented via \seeglossary{schematic -variables}, such that the top-level structure is merely that of a -\seeglossary{Horn Clause}}. - -\glossary{HHF}{See \seeglossary{Hereditary Harrop Formula}.} - - - \[ - \infer[\isa{{\isacharparenleft}assumption{\isacharparenright}}]{\isa{C{\isasymvartheta}}} - {\isa{{\isacharparenleft}{\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec H\ \isactrlvec x\ {\isasymLongrightarrow}\ A\ \isactrlvec x{\isacharparenright}\ {\isasymLongrightarrow}\ C} & \isa{A{\isasymvartheta}\ {\isacharequal}\ H\isactrlsub i{\isasymvartheta}}~~\text{(for some~\isa{i})}} - \] - - - \[ - \infer[\isa{{\isacharparenleft}compose{\isacharparenright}}]{\isa{\isactrlvec A{\isasymvartheta}\ {\isasymLongrightarrow}\ C{\isasymvartheta}}} - {\isa{\isactrlvec A\ {\isasymLongrightarrow}\ B} & \isa{B{\isacharprime}\ {\isasymLongrightarrow}\ C} & \isa{B{\isasymvartheta}\ {\isacharequal}\ B{\isacharprime}{\isasymvartheta}}} - \] - - - \[ - \infer[\isa{{\isacharparenleft}{\isasymAnd}{\isacharunderscore}lift{\isacharparenright}}]{\isa{{\isacharparenleft}{\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec A\ {\isacharparenleft}{\isacharquery}\isactrlvec a\ \isactrlvec x{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymAnd}\isactrlvec x{\isachardot}\ B\ {\isacharparenleft}{\isacharquery}\isactrlvec a\ \isactrlvec x{\isacharparenright}{\isacharparenright}}}{\isa{\isactrlvec A\ {\isacharquery}\isactrlvec a\ {\isasymLongrightarrow}\ B\ {\isacharquery}\isactrlvec a}} - \] - \[ - \infer[\isa{{\isacharparenleft}{\isasymLongrightarrow}{\isacharunderscore}lift{\isacharparenright}}]{\isa{{\isacharparenleft}\isactrlvec H\ {\isasymLongrightarrow}\ \isactrlvec A{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}\isactrlvec H\ {\isasymLongrightarrow}\ B{\isacharparenright}}}{\isa{\isactrlvec A\ {\isasymLongrightarrow}\ B}} - \] - - The \isa{resolve} scheme is now acquired from \isa{{\isasymAnd}{\isacharunderscore}lift}, - \isa{{\isasymLongrightarrow}{\isacharunderscore}lift}, and \isa{compose}. - - \[ - \infer[\isa{{\isacharparenleft}resolution{\isacharparenright}}] - {\isa{{\isacharparenleft}{\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec H\ \isactrlvec x\ {\isasymLongrightarrow}\ \isactrlvec A\ {\isacharparenleft}{\isacharquery}\isactrlvec a\ \isactrlvec x{\isacharparenright}{\isacharparenright}{\isasymvartheta}\ {\isasymLongrightarrow}\ C{\isasymvartheta}}} - {\begin{tabular}{l} - \isa{\isactrlvec A\ {\isacharquery}\isactrlvec a\ {\isasymLongrightarrow}\ B\ {\isacharquery}\isactrlvec a} \\ - \isa{{\isacharparenleft}{\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec H\ \isactrlvec x\ {\isasymLongrightarrow}\ B{\isacharprime}\ \isactrlvec x{\isacharparenright}\ {\isasymLongrightarrow}\ C} \\ - \isa{{\isacharparenleft}{\isasymlambda}\isactrlvec x{\isachardot}\ B\ {\isacharparenleft}{\isacharquery}\isactrlvec a\ \isactrlvec x{\isacharparenright}{\isacharparenright}{\isasymvartheta}\ {\isacharequal}\ B{\isacharprime}{\isasymvartheta}} \\ - \end{tabular}} - \] - - - FIXME \isa{elim{\isacharunderscore}resolution}, \isa{dest{\isacharunderscore}resolution}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagFIXME -{\isafoldFIXME}% -% -\isadelimFIXME -% -\endisadelimFIXME -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/document/prelim.tex --- a/doc-src/IsarImplementation/Thy/document/prelim.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,911 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{prelim}% -% -\isadelimtheory -\isanewline -\isanewline -\isanewline -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ prelim\ \isakeyword{imports}\ base\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Preliminaries% -} -\isamarkuptrue% -% -\isamarkupsection{Contexts \label{sec:context}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A logical context represents the background that is required for - formulating statements and composing proofs. It acts as a medium to - produce formal content, depending on earlier material (declarations, - results etc.). - - For example, derivations within the Isabelle/Pure logic can be - described as a judgment \isa{{\isasymGamma}\ {\isasymturnstile}\isactrlsub {\isasymTheta}\ {\isasymphi}}, which means that a - proposition \isa{{\isasymphi}} is derivable from hypotheses \isa{{\isasymGamma}} - within the theory \isa{{\isasymTheta}}. There are logical reasons for - keeping \isa{{\isasymTheta}} and \isa{{\isasymGamma}} separate: theories can be - liberal about supporting type constructors and schematic - polymorphism of constants and axioms, while the inner calculus of - \isa{{\isasymGamma}\ {\isasymturnstile}\ {\isasymphi}} is strictly limited to Simple Type Theory (with - fixed type variables in the assumptions). - - \medskip Contexts and derivations are linked by the following key - principles: - - \begin{itemize} - - \item Transfer: monotonicity of derivations admits results to be - transferred into a \emph{larger} context, i.e.\ \isa{{\isasymGamma}\ {\isasymturnstile}\isactrlsub {\isasymTheta}\ {\isasymphi}} implies \isa{{\isasymGamma}{\isacharprime}\ {\isasymturnstile}\isactrlsub {\isasymTheta}\isactrlsub {\isacharprime}\ {\isasymphi}} for contexts \isa{{\isasymTheta}{\isacharprime}\ {\isasymsupseteq}\ {\isasymTheta}} and \isa{{\isasymGamma}{\isacharprime}\ {\isasymsupseteq}\ {\isasymGamma}}. - - \item Export: discharge of hypotheses admits results to be exported - into a \emph{smaller} context, i.e.\ \isa{{\isasymGamma}{\isacharprime}\ {\isasymturnstile}\isactrlsub {\isasymTheta}\ {\isasymphi}} - implies \isa{{\isasymGamma}\ {\isasymturnstile}\isactrlsub {\isasymTheta}\ {\isasymDelta}\ {\isasymLongrightarrow}\ {\isasymphi}} where \isa{{\isasymGamma}{\isacharprime}\ {\isasymsupseteq}\ {\isasymGamma}} and - \isa{{\isasymDelta}\ {\isacharequal}\ {\isasymGamma}{\isacharprime}\ {\isacharminus}\ {\isasymGamma}}. Note that \isa{{\isasymTheta}} remains unchanged here, - only the \isa{{\isasymGamma}} part is affected. - - \end{itemize} - - \medskip By modeling the main characteristics of the primitive - \isa{{\isasymTheta}} and \isa{{\isasymGamma}} above, and abstracting over any - particular logical content, we arrive at the fundamental notions of - \emph{theory context} and \emph{proof context} in Isabelle/Isar. - These implement a certain policy to manage arbitrary \emph{context - data}. There is a strongly-typed mechanism to declare new kinds of - data at compile time. - - The internal bootstrap process of Isabelle/Pure eventually reaches a - stage where certain data slots provide the logical content of \isa{{\isasymTheta}} and \isa{{\isasymGamma}} sketched above, but this does not stop there! - Various additional data slots support all kinds of mechanisms that - are not necessarily part of the core logic. - - For example, there would be data for canonical introduction and - elimination rules for arbitrary operators (depending on the - object-logic and application), which enables users to perform - standard proof steps implicitly (cf.\ the \isa{rule} method - \cite{isabelle-isar-ref}). - - \medskip Thus Isabelle/Isar is able to bring forth more and more - concepts successively. In particular, an object-logic like - Isabelle/HOL continues the Isabelle/Pure setup by adding specific - components for automated reasoning (classical reasoner, tableau - prover, structured induction etc.) and derived specification - mechanisms (inductive predicates, recursive functions etc.). All of - this is ultimately based on the generic data management by theory - and proof contexts introduced here.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Theory context \label{sec:context-theory}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\glossary{Theory}{FIXME} - - A \emph{theory} is a data container with explicit named and unique - identifier. Theories are related by a (nominal) sub-theory - relation, which corresponds to the dependency graph of the original - construction; each theory is derived from a certain sub-graph of - ancestor theories. - - The \isa{merge} operation produces the least upper bound of two - theories, which actually degenerates into absorption of one theory - into the other (due to the nominal sub-theory relation). - - The \isa{begin} operation starts a new theory by importing - several parent theories and entering a special \isa{draft} mode, - which is sustained until the final \isa{end} operation. A draft - theory acts like a linear type, where updates invalidate earlier - versions. An invalidated draft is called ``stale''. - - The \isa{checkpoint} operation produces an intermediate stepping - stone that will survive the next update: both the original and the - changed theory remain valid and are related by the sub-theory - relation. Checkpointing essentially recovers purely functional - theory values, at the expense of some extra internal bookkeeping. - - The \isa{copy} operation produces an auxiliary version that has - the same data content, but is unrelated to the original: updates of - the copy do not affect the original, neither does the sub-theory - relation hold. - - \medskip The example in \figref{fig:ex-theory} below shows a theory - graph derived from \isa{Pure}, with theory \isa{Length} - importing \isa{Nat} and \isa{List}. The body of \isa{Length} consists of a sequence of updates, working mostly on - drafts. Intermediate checkpoints may occur as well, due to the - history mechanism provided by the Isar top-level, cf.\ - \secref{sec:isar-toplevel}. - - \begin{figure}[htb] - \begin{center} - \begin{tabular}{rcccl} - & & \isa{Pure} \\ - & & \isa{{\isasymdown}} \\ - & & \isa{FOL} \\ - & $\swarrow$ & & $\searrow$ & \\ - \isa{Nat} & & & & \isa{List} \\ - & $\searrow$ & & $\swarrow$ \\ - & & \isa{Length} \\ - & & \multicolumn{3}{l}{~~\hyperlink{keyword.imports}{\mbox{\isa{\isakeyword{imports}}}}} \\ - & & \multicolumn{3}{l}{~~\hyperlink{keyword.begin}{\mbox{\isa{\isakeyword{begin}}}}} \\ - & & $\vdots$~~ \\ - & & \isa{{\isasymbullet}}~~ \\ - & & $\vdots$~~ \\ - & & \isa{{\isasymbullet}}~~ \\ - & & $\vdots$~~ \\ - & & \multicolumn{3}{l}{~~\hyperlink{command.end}{\mbox{\isa{\isacommand{end}}}}} \\ - \end{tabular} - \caption{A theory definition depending on ancestors}\label{fig:ex-theory} - \end{center} - \end{figure} - - \medskip There is a separate notion of \emph{theory reference} for - maintaining a live link to an evolving theory context: updates on - drafts are propagated automatically. Dynamic updating stops after - an explicit \isa{end} only. - - Derived entities may store a theory reference in order to indicate - the context they belong to. This implicitly assumes monotonic - reasoning, because the referenced context may become larger without - further notice.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexmltype{theory}\verb|type theory| \\ - \indexml{Theory.subthy}\verb|Theory.subthy: theory * theory -> bool| \\ - \indexml{Theory.merge}\verb|Theory.merge: theory * theory -> theory| \\ - \indexml{Theory.checkpoint}\verb|Theory.checkpoint: theory -> theory| \\ - \indexml{Theory.copy}\verb|Theory.copy: theory -> theory| \\ - \end{mldecls} - \begin{mldecls} - \indexmltype{theory\_ref}\verb|type theory_ref| \\ - \indexml{Theory.deref}\verb|Theory.deref: theory_ref -> theory| \\ - \indexml{Theory.check\_thy}\verb|Theory.check_thy: theory -> theory_ref| \\ - \end{mldecls} - - \begin{description} - - \item \verb|theory| represents theory contexts. This is - essentially a linear type! Most operations destroy the original - version, which then becomes ``stale''. - - \item \verb|Theory.subthy|~\isa{{\isacharparenleft}thy\isactrlsub {\isadigit{1}}{\isacharcomma}\ thy\isactrlsub {\isadigit{2}}{\isacharparenright}} - compares theories according to the inherent graph structure of the - construction. This sub-theory relation is a nominal approximation - of inclusion (\isa{{\isasymsubseteq}}) of the corresponding content. - - \item \verb|Theory.merge|~\isa{{\isacharparenleft}thy\isactrlsub {\isadigit{1}}{\isacharcomma}\ thy\isactrlsub {\isadigit{2}}{\isacharparenright}} - absorbs one theory into the other. This fails for unrelated - theories! - - \item \verb|Theory.checkpoint|~\isa{thy} produces a safe - stepping stone in the linear development of \isa{thy}. The next - update will result in two related, valid theories. - - \item \verb|Theory.copy|~\isa{thy} produces a variant of \isa{thy} that holds a copy of the same data. The result is not - related to the original; the original is unchanched. - - \item \verb|theory_ref| represents a sliding reference to an - always valid theory; updates on the original are propagated - automatically. - - \item \verb|Theory.deref|~\isa{thy{\isacharunderscore}ref} turns a \verb|theory_ref| into an \verb|theory| value. As the referenced - theory evolves monotonically over time, later invocations of \verb|Theory.deref| may refer to a larger context. - - \item \verb|Theory.check_thy|~\isa{thy} produces a \verb|theory_ref| from a valid \verb|theory| value. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Proof context \label{sec:context-proof}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\glossary{Proof context}{The static context of a structured proof, - acts like a local ``theory'' of the current portion of Isar proof - text, generalizes the idea of local hypotheses \isa{{\isasymGamma}} in - judgments \isa{{\isasymGamma}\ {\isasymturnstile}\ {\isasymphi}} of natural deduction calculi. There is a - generic notion of introducing and discharging hypotheses. - Arbritrary auxiliary context data may be adjoined.} - - A proof context is a container for pure data with a back-reference - to the theory it belongs to. The \isa{init} operation creates a - proof context from a given theory. Modifications to draft theories - are propagated to the proof context as usual, but there is also an - explicit \isa{transfer} operation to force resynchronization - with more substantial updates to the underlying theory. The actual - context data does not require any special bookkeeping, thanks to the - lack of destructive features. - - Entities derived in a proof context need to record inherent logical - requirements explicitly, since there is no separate context - identification as for theories. For example, hypotheses used in - primitive derivations (cf.\ \secref{sec:thms}) are recorded - separately within the sequent \isa{{\isasymGamma}\ {\isasymturnstile}\ {\isasymphi}}, just to make double - sure. Results could still leak into an alien proof context do to - programming errors, but Isabelle/Isar includes some extra validity - checks in critical positions, notably at the end of a sub-proof. - - Proof contexts may be manipulated arbitrarily, although the common - discipline is to follow block structure as a mental model: a given - context is extended consecutively, and results are exported back - into the original context. Note that the Isar proof states model - block-structured reasoning explicitly, using a stack of proof - contexts internally, cf.\ \secref{sec:isar-proof-state}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexmltype{Proof.context}\verb|type Proof.context| \\ - \indexml{ProofContext.init}\verb|ProofContext.init: theory -> Proof.context| \\ - \indexml{ProofContext.theory\_of}\verb|ProofContext.theory_of: Proof.context -> theory| \\ - \indexml{ProofContext.transfer}\verb|ProofContext.transfer: theory -> Proof.context -> Proof.context| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Proof.context| represents proof contexts. Elements - of this type are essentially pure values, with a sliding reference - to the background theory. - - \item \verb|ProofContext.init|~\isa{thy} produces a proof context - derived from \isa{thy}, initializing all data. - - \item \verb|ProofContext.theory_of|~\isa{ctxt} selects the - background theory from \isa{ctxt}, dereferencing its internal - \verb|theory_ref|. - - \item \verb|ProofContext.transfer|~\isa{thy\ ctxt} promotes the - background theory of \isa{ctxt} to the super theory \isa{thy}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Generic contexts \label{sec:generic-context}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A generic context is the disjoint sum of either a theory or proof - context. Occasionally, this enables uniform treatment of generic - context data, typically extra-logical information. Operations on - generic contexts include the usual injections, partial selections, - and combinators for lifting operations on either component of the - disjoint sum. - - Moreover, there are total operations \isa{theory{\isacharunderscore}of} and \isa{proof{\isacharunderscore}of} to convert a generic context into either kind: a theory - can always be selected from the sum, while a proof context might - have to be constructed by an ad-hoc \isa{init} operation.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexmltype{Context.generic}\verb|type Context.generic| \\ - \indexml{Context.theory\_of}\verb|Context.theory_of: Context.generic -> theory| \\ - \indexml{Context.proof\_of}\verb|Context.proof_of: Context.generic -> Proof.context| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Context.generic| is the direct sum of \verb|theory| and \verb|Proof.context|, with the datatype - constructors \verb|Context.Theory| and \verb|Context.Proof|. - - \item \verb|Context.theory_of|~\isa{context} always produces a - theory from the generic \isa{context}, using \verb|ProofContext.theory_of| as required. - - \item \verb|Context.proof_of|~\isa{context} always produces a - proof context from the generic \isa{context}, using \verb|ProofContext.init| as required (note that this re-initializes the - context data with each invocation). - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Context data \label{sec:context-data}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The main purpose of theory and proof contexts is to manage arbitrary - data. New data types can be declared incrementally at compile time. - There are separate declaration mechanisms for any of the three kinds - of contexts: theory, proof, generic. - - \paragraph{Theory data} may refer to destructive entities, which are - maintained in direct correspondence to the linear evolution of - theory values, including explicit copies.\footnote{Most existing - instances of destructive theory data are merely historical relics - (e.g.\ the destructive theorem storage, and destructive hints for - the Simplifier and Classical rules).} A theory data declaration - needs to implement the following SML signature: - - \medskip - \begin{tabular}{ll} - \isa{{\isasymtype}\ T} & representing type \\ - \isa{{\isasymval}\ empty{\isacharcolon}\ T} & empty default value \\ - \isa{{\isasymval}\ copy{\isacharcolon}\ T\ {\isasymrightarrow}\ T} & refresh impure data \\ - \isa{{\isasymval}\ extend{\isacharcolon}\ T\ {\isasymrightarrow}\ T} & re-initialize on import \\ - \isa{{\isasymval}\ merge{\isacharcolon}\ T\ {\isasymtimes}\ T\ {\isasymrightarrow}\ T} & join on import \\ - \end{tabular} - \medskip - - \noindent The \isa{empty} value acts as initial default for - \emph{any} theory that does not declare actual data content; \isa{copy} maintains persistent integrity for impure data, it is just - the identity for pure values; \isa{extend} is acts like a - unitary version of \isa{merge}, both operations should also - include the functionality of \isa{copy} for impure data. - - \paragraph{Proof context data} is purely functional. A declaration - needs to implement the following SML signature: - - \medskip - \begin{tabular}{ll} - \isa{{\isasymtype}\ T} & representing type \\ - \isa{{\isasymval}\ init{\isacharcolon}\ theory\ {\isasymrightarrow}\ T} & produce initial value \\ - \end{tabular} - \medskip - - \noindent The \isa{init} operation is supposed to produce a pure - value from the given background theory. - - \paragraph{Generic data} provides a hybrid interface for both theory - and proof data. The declaration is essentially the same as for - (pure) theory data, without \isa{copy}. The \isa{init} - operation for proof contexts merely selects the current data value - from the background theory. - - \bigskip A data declaration of type \isa{T} results in the - following interface: - - \medskip - \begin{tabular}{ll} - \isa{init{\isacharcolon}\ theory\ {\isasymrightarrow}\ theory} \\ - \isa{get{\isacharcolon}\ context\ {\isasymrightarrow}\ T} \\ - \isa{put{\isacharcolon}\ T\ {\isasymrightarrow}\ context\ {\isasymrightarrow}\ context} \\ - \isa{map{\isacharcolon}\ {\isacharparenleft}T\ {\isasymrightarrow}\ T{\isacharparenright}\ {\isasymrightarrow}\ context\ {\isasymrightarrow}\ context} \\ - \end{tabular} - \medskip - - \noindent Here \isa{init} is only applicable to impure theory - data to install a fresh copy persistently (destructive update on - uninitialized has no permanent effect). The other operations provide - access for the particular kind of context (theory, proof, or generic - context). Note that this is a safe interface: there is no other way - to access the corresponding data slot of a context. By keeping - these operations private, a component may maintain abstract values - authentically, without other components interfering.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexmlfunctor{TheoryDataFun}\verb|functor TheoryDataFun| \\ - \indexmlfunctor{ProofDataFun}\verb|functor ProofDataFun| \\ - \indexmlfunctor{GenericDataFun}\verb|functor GenericDataFun| \\ - \end{mldecls} - - \begin{description} - - \item \verb|TheoryDataFun|\isa{{\isacharparenleft}spec{\isacharparenright}} declares data for - type \verb|theory| according to the specification provided as - argument structure. The resulting structure provides data init and - access operations as described above. - - \item \verb|ProofDataFun|\isa{{\isacharparenleft}spec{\isacharparenright}} is analogous to - \verb|TheoryDataFun| for type \verb|Proof.context|. - - \item \verb|GenericDataFun|\isa{{\isacharparenleft}spec{\isacharparenright}} is analogous to - \verb|TheoryDataFun| for type \verb|Context.generic|. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsection{Names \label{sec:names}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In principle, a name is just a string, but there are various - convention for encoding additional structure. For example, ``\isa{Foo{\isachardot}bar{\isachardot}baz}'' is considered as a qualified name consisting of - three basic name components. The individual constituents of a name - may have further substructure, e.g.\ the string - ``\verb,\,\verb,,'' encodes as a single symbol.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Strings of symbols% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\glossary{Symbol}{The smallest unit of text in Isabelle, subsumes - plain ASCII characters as well as an infinite collection of named - symbols (for greek, math etc.).} - - A \emph{symbol} constitutes the smallest textual unit in Isabelle - --- raw characters are normally not encountered at all. Isabelle - strings consist of a sequence of symbols, represented as a packed - string or a list of strings. Each symbol is in itself a small - string, which has either one of the following forms: - - \begin{enumerate} - - \item a single ASCII character ``\isa{c}'', for example - ``\verb,a,'', - - \item a regular symbol ``\verb,\,\verb,<,\isa{ident}\verb,>,'', - for example ``\verb,\,\verb,,'', - - \item a control symbol ``\verb,\,\verb,<^,\isa{ident}\verb,>,'', - for example ``\verb,\,\verb,<^bold>,'', - - \item a raw symbol ``\verb,\,\verb,<^raw:,\isa{text}\verb,>,'' - where \isa{text} constists of printable characters excluding - ``\verb,.,'' and ``\verb,>,'', for example - ``\verb,\,\verb,<^raw:$\sum_{i = 1}^n$>,'', - - \item a numbered raw control symbol ``\verb,\,\verb,<^raw,\isa{n}\verb,>, where \isa{n} consists of digits, for example - ``\verb,\,\verb,<^raw42>,''. - - \end{enumerate} - - \noindent The \isa{ident} syntax for symbol names is \isa{letter\ {\isacharparenleft}letter\ {\isacharbar}\ digit{\isacharparenright}\isactrlsup {\isacharasterisk}}, where \isa{letter\ {\isacharequal}\ A{\isachardot}{\isachardot}Za{\isachardot}{\isachardot}z} and \isa{digit\ {\isacharequal}\ {\isadigit{0}}{\isachardot}{\isachardot}{\isadigit{9}}}. There are infinitely many - regular symbols and control symbols, but a fixed collection of - standard symbols is treated specifically. For example, - ``\verb,\,\verb,,'' is classified as a letter, which means it - may occur within regular Isabelle identifiers. - - Since the character set underlying Isabelle symbols is 7-bit ASCII - and 8-bit characters are passed through transparently, Isabelle may - also process Unicode/UCS data in UTF-8 encoding. Unicode provides - its own collection of mathematical symbols, but there is no built-in - link to the standard collection of Isabelle. - - \medskip Output of Isabelle symbols depends on the print mode - (\secref{FIXME}). For example, the standard {\LaTeX} setup of the - Isabelle document preparation system would present - ``\verb,\,\verb,,'' as \isa{{\isasymalpha}}, and - ``\verb,\,\verb,<^bold>,\verb,\,\verb,,'' as \isa{\isactrlbold {\isasymalpha}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexmltype{Symbol.symbol}\verb|type Symbol.symbol| \\ - \indexml{Symbol.explode}\verb|Symbol.explode: string -> Symbol.symbol list| \\ - \indexml{Symbol.is\_letter}\verb|Symbol.is_letter: Symbol.symbol -> bool| \\ - \indexml{Symbol.is\_digit}\verb|Symbol.is_digit: Symbol.symbol -> bool| \\ - \indexml{Symbol.is\_quasi}\verb|Symbol.is_quasi: Symbol.symbol -> bool| \\ - \indexml{Symbol.is\_blank}\verb|Symbol.is_blank: Symbol.symbol -> bool| \\ - \end{mldecls} - \begin{mldecls} - \indexmltype{Symbol.sym}\verb|type Symbol.sym| \\ - \indexml{Symbol.decode}\verb|Symbol.decode: Symbol.symbol -> Symbol.sym| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Symbol.symbol| represents individual Isabelle - symbols; this is an alias for \verb|string|. - - \item \verb|Symbol.explode|~\isa{str} produces a symbol list - from the packed form. This function supercedes \verb|String.explode| for virtually all purposes of manipulating text in - Isabelle! - - \item \verb|Symbol.is_letter|, \verb|Symbol.is_digit|, \verb|Symbol.is_quasi|, \verb|Symbol.is_blank| classify standard - symbols according to fixed syntactic conventions of Isabelle, cf.\ - \cite{isabelle-isar-ref}. - - \item \verb|Symbol.sym| is a concrete datatype that represents - the different kinds of symbols explicitly, with constructors \verb|Symbol.Char|, \verb|Symbol.Sym|, \verb|Symbol.Ctrl|, \verb|Symbol.Raw|. - - \item \verb|Symbol.decode| converts the string representation of a - symbol into the datatype version. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Basic names \label{sec:basic-names}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A \emph{basic name} essentially consists of a single Isabelle - identifier. There are conventions to mark separate classes of basic - names, by attaching a suffix of underscores (\isa{{\isacharunderscore}}): one - underscore means \emph{internal name}, two underscores means - \emph{Skolem name}, three underscores means \emph{internal Skolem - name}. - - For example, the basic name \isa{foo} has the internal version - \isa{foo{\isacharunderscore}}, with Skolem versions \isa{foo{\isacharunderscore}{\isacharunderscore}} and \isa{foo{\isacharunderscore}{\isacharunderscore}{\isacharunderscore}}, respectively. - - These special versions provide copies of the basic name space, apart - from anything that normally appears in the user text. For example, - system generated variables in Isar proof contexts are usually marked - as internal, which prevents mysterious name references like \isa{xaa} to appear in the text. - - \medskip Manipulating binding scopes often requires on-the-fly - renamings. A \emph{name context} contains a collection of already - used names. The \isa{declare} operation adds names to the - context. - - The \isa{invents} operation derives a number of fresh names from - a given starting point. For example, the first three names derived - from \isa{a} are \isa{a}, \isa{b}, \isa{c}. - - The \isa{variants} operation produces fresh names by - incrementing tentative names as base-26 numbers (with digits \isa{a{\isachardot}{\isachardot}z}) until all clashes are resolved. For example, name \isa{foo} results in variants \isa{fooa}, \isa{foob}, \isa{fooc}, \dots, \isa{fooaa}, \isa{fooab} etc.; each renaming - step picks the next unused variant from this sequence.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Name.internal}\verb|Name.internal: string -> string| \\ - \indexml{Name.skolem}\verb|Name.skolem: string -> string| \\ - \end{mldecls} - \begin{mldecls} - \indexmltype{Name.context}\verb|type Name.context| \\ - \indexml{Name.context}\verb|Name.context: Name.context| \\ - \indexml{Name.declare}\verb|Name.declare: string -> Name.context -> Name.context| \\ - \indexml{Name.invents}\verb|Name.invents: Name.context -> string -> int -> string list| \\ - \indexml{Name.variants}\verb|Name.variants: string list -> Name.context -> string list * Name.context| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Name.internal|~\isa{name} produces an internal name - by adding one underscore. - - \item \verb|Name.skolem|~\isa{name} produces a Skolem name by - adding two underscores. - - \item \verb|Name.context| represents the context of already used - names; the initial value is \verb|Name.context|. - - \item \verb|Name.declare|~\isa{name} enters a used name into the - context. - - \item \verb|Name.invents|~\isa{context\ name\ n} produces \isa{n} fresh names derived from \isa{name}. - - \item \verb|Name.variants|~\isa{names\ context} produces fresh - varians of \isa{names}; the result is entered into the context. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Indexed names% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -An \emph{indexed name} (or \isa{indexname}) is a pair of a basic - name and a natural number. This representation allows efficient - renaming by incrementing the second component only. The canonical - way to rename two collections of indexnames apart from each other is - this: determine the maximum index \isa{maxidx} of the first - collection, then increment all indexes of the second collection by - \isa{maxidx\ {\isacharplus}\ {\isadigit{1}}}; the maximum index of an empty collection is - \isa{{\isacharminus}{\isadigit{1}}}. - - Occasionally, basic names and indexed names are injected into the - same pair type: the (improper) indexname \isa{{\isacharparenleft}x{\isacharcomma}\ {\isacharminus}{\isadigit{1}}{\isacharparenright}} is used - to encode basic names. - - \medskip Isabelle syntax observes the following rules for - representing an indexname \isa{{\isacharparenleft}x{\isacharcomma}\ i{\isacharparenright}} as a packed string: - - \begin{itemize} - - \item \isa{{\isacharquery}x} if \isa{x} does not end with a digit and \isa{i\ {\isacharequal}\ {\isadigit{0}}}, - - \item \isa{{\isacharquery}xi} if \isa{x} does not end with a digit, - - \item \isa{{\isacharquery}x{\isachardot}i} otherwise. - - \end{itemize} - - Indexnames may acquire large index numbers over time. Results are - normalized towards \isa{{\isadigit{0}}} at certain checkpoints, notably at - the end of a proof. This works by producing variants of the - corresponding basic name components. For example, the collection - \isa{{\isacharquery}x{\isadigit{1}}{\isacharcomma}\ {\isacharquery}x{\isadigit{7}}{\isacharcomma}\ {\isacharquery}x{\isadigit{4}}{\isadigit{2}}} becomes \isa{{\isacharquery}x{\isacharcomma}\ {\isacharquery}xa{\isacharcomma}\ {\isacharquery}xb}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexmltype{indexname}\verb|type indexname| \\ - \end{mldecls} - - \begin{description} - - \item \verb|indexname| represents indexed names. This is an - abbreviation for \verb|string * int|. The second component is - usually non-negative, except for situations where \isa{{\isacharparenleft}x{\isacharcomma}\ {\isacharminus}{\isadigit{1}}{\isacharparenright}} - is used to embed basic names into this type. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Qualified names and name spaces% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A \emph{qualified name} consists of a non-empty sequence of basic - name components. The packed representation uses a dot as separator, - as in ``\isa{A{\isachardot}b{\isachardot}c}''. The last component is called \emph{base} - name, the remaining prefix \emph{qualifier} (which may be empty). - The idea of qualified names is to encode nested structures by - recording the access paths as qualifiers. For example, an item - named ``\isa{A{\isachardot}b{\isachardot}c}'' may be understood as a local entity \isa{c}, within a local structure \isa{b}, within a global - structure \isa{A}. Typically, name space hierarchies consist of - 1--2 levels of qualification, but this need not be always so. - - The empty name is commonly used as an indication of unnamed - entities, whenever this makes any sense. The basic operations on - qualified names are smart enough to pass through such improper names - unchanged. - - \medskip A \isa{naming} policy tells how to turn a name - specification into a fully qualified internal name (by the \isa{full} operation), and how fully qualified names may be accessed - externally. For example, the default naming policy is to prefix an - implicit path: \isa{full\ x} produces \isa{path{\isachardot}x}, and the - standard accesses for \isa{path{\isachardot}x} include both \isa{x} and - \isa{path{\isachardot}x}. Normally, the naming is implicit in the theory or - proof context; there are separate versions of the corresponding. - - \medskip A \isa{name\ space} manages a collection of fully - internalized names, together with a mapping between external names - and internal names (in both directions). The corresponding \isa{intern} and \isa{extern} operations are mostly used for - parsing and printing only! The \isa{declare} operation augments - a name space according to the accesses determined by the naming - policy. - - \medskip As a general principle, there is a separate name space for - each kind of formal entity, e.g.\ logical constant, type - constructor, type class, theorem. It is usually clear from the - occurrence in concrete syntax (or from the scope) which kind of - entity a name refers to. For example, the very same name \isa{c} may be used uniformly for a constant, type constructor, and - type class. - - There are common schemes to name theorems systematically, according - to the name of the main logical entity involved, e.g.\ \isa{c{\isachardot}intro} for a canonical theorem related to constant \isa{c}. - This technique of mapping names from one space into another requires - some care in order to avoid conflicts. In particular, theorem names - derived from a type constructor or type class are better suffixed in - addition to the usual qualification, e.g.\ \isa{c{\isacharunderscore}type{\isachardot}intro} - and \isa{c{\isacharunderscore}class{\isachardot}intro} for theorems related to type \isa{c} - and class \isa{c}, respectively.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{NameSpace.base}\verb|NameSpace.base: string -> string| \\ - \indexml{NameSpace.qualifier}\verb|NameSpace.qualifier: string -> string| \\ - \indexml{NameSpace.append}\verb|NameSpace.append: string -> string -> string| \\ - \indexml{NameSpace.implode}\verb|NameSpace.implode: string list -> string| \\ - \indexml{NameSpace.explode}\verb|NameSpace.explode: string -> string list| \\ - \end{mldecls} - \begin{mldecls} - \indexmltype{NameSpace.naming}\verb|type NameSpace.naming| \\ - \indexml{NameSpace.default\_naming}\verb|NameSpace.default_naming: NameSpace.naming| \\ - \indexml{NameSpace.add\_path}\verb|NameSpace.add_path: string -> NameSpace.naming -> NameSpace.naming| \\ - \indexml{NameSpace.full\_name}\verb|NameSpace.full_name: NameSpace.naming -> binding -> string| \\ - \end{mldecls} - \begin{mldecls} - \indexmltype{NameSpace.T}\verb|type NameSpace.T| \\ - \indexml{NameSpace.empty}\verb|NameSpace.empty: NameSpace.T| \\ - \indexml{NameSpace.merge}\verb|NameSpace.merge: NameSpace.T * NameSpace.T -> NameSpace.T| \\ - \indexml{NameSpace.declare}\verb|NameSpace.declare: NameSpace.naming -> binding -> NameSpace.T -> string * NameSpace.T| \\ - \indexml{NameSpace.intern}\verb|NameSpace.intern: NameSpace.T -> string -> string| \\ - \indexml{NameSpace.extern}\verb|NameSpace.extern: NameSpace.T -> string -> string| \\ - \end{mldecls} - - \begin{description} - - \item \verb|NameSpace.base|~\isa{name} returns the base name of a - qualified name. - - \item \verb|NameSpace.qualifier|~\isa{name} returns the qualifier - of a qualified name. - - \item \verb|NameSpace.append|~\isa{name\isactrlisub {\isadigit{1}}\ name\isactrlisub {\isadigit{2}}} - appends two qualified names. - - \item \verb|NameSpace.implode|~\isa{name} and \verb|NameSpace.explode|~\isa{names} convert between the packed string - representation and the explicit list form of qualified names. - - \item \verb|NameSpace.naming| represents the abstract concept of - a naming policy. - - \item \verb|NameSpace.default_naming| is the default naming policy. - In a theory context, this is usually augmented by a path prefix - consisting of the theory name. - - \item \verb|NameSpace.add_path|~\isa{path\ naming} augments the - naming policy by extending its path component. - - \item \verb|NameSpace.full_name|\isa{naming\ binding} turns a name - binding (usually a basic name) into the fully qualified - internal name, according to the given naming policy. - - \item \verb|NameSpace.T| represents name spaces. - - \item \verb|NameSpace.empty| and \verb|NameSpace.merge|~\isa{{\isacharparenleft}space\isactrlisub {\isadigit{1}}{\isacharcomma}\ space\isactrlisub {\isadigit{2}}{\isacharparenright}} are the canonical operations for - maintaining name spaces according to theory data management - (\secref{sec:context-data}). - - \item \verb|NameSpace.declare|~\isa{naming\ bindings\ space} enters a - name binding as fully qualified internal name into the name space, - with external accesses determined by the naming policy. - - \item \verb|NameSpace.intern|~\isa{space\ name} internalizes a - (partially qualified) external name. - - This operation is mostly for parsing! Note that fully qualified - names stemming from declarations are produced via \verb|NameSpace.full_name| and \verb|NameSpace.declare| - (or their derivatives for \verb|theory| and - \verb|Proof.context|). - - \item \verb|NameSpace.extern|~\isa{space\ name} externalizes a - (fully qualified) internal name. - - This operation is mostly for printing! Note unqualified names are - produced via \verb|NameSpace.base|. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/document/proof.tex --- a/doc-src/IsarImplementation/Thy/document/proof.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,396 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{proof}% -% -\isadelimtheory -\isanewline -\isanewline -\isanewline -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ {\isachardoublequoteopen}proof{\isachardoublequoteclose}\ \isakeyword{imports}\ base\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Structured proofs% -} -\isamarkuptrue% -% -\isamarkupsection{Variables \label{sec:variables}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Any variable that is not explicitly bound by \isa{{\isasymlambda}}-abstraction - is considered as ``free''. Logically, free variables act like - outermost universal quantification at the sequent level: \isa{A\isactrlisub {\isadigit{1}}{\isacharparenleft}x{\isacharparenright}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ A\isactrlisub n{\isacharparenleft}x{\isacharparenright}\ {\isasymturnstile}\ B{\isacharparenleft}x{\isacharparenright}} means that the result - holds \emph{for all} values of \isa{x}. Free variables for - terms (not types) can be fully internalized into the logic: \isa{{\isasymturnstile}\ B{\isacharparenleft}x{\isacharparenright}} and \isa{{\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ B{\isacharparenleft}x{\isacharparenright}} are interchangeable, provided - that \isa{x} does not occur elsewhere in the context. - Inspecting \isa{{\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ B{\isacharparenleft}x{\isacharparenright}} more closely, we see that inside the - quantifier, \isa{x} is essentially ``arbitrary, but fixed'', - while from outside it appears as a place-holder for instantiation - (thanks to \isa{{\isasymAnd}} elimination). - - The Pure logic represents the idea of variables being either inside - or outside the current scope by providing separate syntactic - categories for \emph{fixed variables} (e.g.\ \isa{x}) vs.\ - \emph{schematic variables} (e.g.\ \isa{{\isacharquery}x}). Incidently, a - universal result \isa{{\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ B{\isacharparenleft}x{\isacharparenright}} has the HHF normal form \isa{{\isasymturnstile}\ B{\isacharparenleft}{\isacharquery}x{\isacharparenright}}, which represents its generality nicely without requiring - an explicit quantifier. The same principle works for type - variables: \isa{{\isasymturnstile}\ B{\isacharparenleft}{\isacharquery}{\isasymalpha}{\isacharparenright}} represents the idea of ``\isa{{\isasymturnstile}\ {\isasymforall}{\isasymalpha}{\isachardot}\ B{\isacharparenleft}{\isasymalpha}{\isacharparenright}}'' without demanding a truly polymorphic framework. - - \medskip Additional care is required to treat type variables in a - way that facilitates type-inference. In principle, term variables - depend on type variables, which means that type variables would have - to be declared first. For example, a raw type-theoretic framework - would demand the context to be constructed in stages as follows: - \isa{{\isasymGamma}\ {\isacharequal}\ {\isasymalpha}{\isacharcolon}\ type{\isacharcomma}\ x{\isacharcolon}\ {\isasymalpha}{\isacharcomma}\ a{\isacharcolon}\ A{\isacharparenleft}x\isactrlisub {\isasymalpha}{\isacharparenright}}. - - We allow a slightly less formalistic mode of operation: term - variables \isa{x} are fixed without specifying a type yet - (essentially \emph{all} potential occurrences of some instance - \isa{x\isactrlisub {\isasymtau}} are fixed); the first occurrence of \isa{x} - within a specific term assigns its most general type, which is then - maintained consistently in the context. The above example becomes - \isa{{\isasymGamma}\ {\isacharequal}\ x{\isacharcolon}\ term{\isacharcomma}\ {\isasymalpha}{\isacharcolon}\ type{\isacharcomma}\ A{\isacharparenleft}x\isactrlisub {\isasymalpha}{\isacharparenright}}, where type \isa{{\isasymalpha}} is fixed \emph{after} term \isa{x}, and the constraint - \isa{x\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}} is an implicit consequence of the occurrence of - \isa{x\isactrlisub {\isasymalpha}} in the subsequent proposition. - - This twist of dependencies is also accommodated by the reverse - operation of exporting results from a context: a type variable - \isa{{\isasymalpha}} is considered fixed as long as it occurs in some fixed - term variable of the context. For example, exporting \isa{x{\isacharcolon}\ term{\isacharcomma}\ {\isasymalpha}{\isacharcolon}\ type\ {\isasymturnstile}\ x\isactrlisub {\isasymalpha}\ {\isacharequal}\ x\isactrlisub {\isasymalpha}} produces in the first step - \isa{x{\isacharcolon}\ term\ {\isasymturnstile}\ x\isactrlisub {\isasymalpha}\ {\isacharequal}\ x\isactrlisub {\isasymalpha}} for fixed \isa{{\isasymalpha}}, - and only in the second step \isa{{\isasymturnstile}\ {\isacharquery}x\isactrlisub {\isacharquery}\isactrlisub {\isasymalpha}\ {\isacharequal}\ {\isacharquery}x\isactrlisub {\isacharquery}\isactrlisub {\isasymalpha}} for schematic \isa{{\isacharquery}x} and \isa{{\isacharquery}{\isasymalpha}}. - - \medskip The Isabelle/Isar proof context manages the gory details of - term vs.\ type variables, with high-level principles for moving the - frontier between fixed and schematic variables. - - The \isa{add{\isacharunderscore}fixes} operation explictly declares fixed - variables; the \isa{declare{\isacharunderscore}term} operation absorbs a term into - a context by fixing new type variables and adding syntactic - constraints. - - The \isa{export} operation is able to perform the main work of - generalizing term and type variables as sketched above, assuming - that fixing variables and terms have been declared properly. - - There \isa{import} operation makes a generalized fact a genuine - part of the context, by inventing fixed variables for the schematic - ones. The effect can be reversed by using \isa{export} later, - potentially with an extended context; the result is equivalent to - the original modulo renaming of schematic variables. - - The \isa{focus} operation provides a variant of \isa{import} - for nested propositions (with explicit quantification): \isa{{\isasymAnd}x\isactrlisub {\isadigit{1}}\ {\isasymdots}\ x\isactrlisub n{\isachardot}\ B{\isacharparenleft}x\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ x\isactrlisub n{\isacharparenright}} is - decomposed by inventing fixed variables \isa{x\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ x\isactrlisub n} for the body.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Variable.add\_fixes}\verb|Variable.add_fixes: |\isasep\isanewline% -\verb| string list -> Proof.context -> string list * Proof.context| \\ - \indexml{Variable.variant\_fixes}\verb|Variable.variant_fixes: |\isasep\isanewline% -\verb| string list -> Proof.context -> string list * Proof.context| \\ - \indexml{Variable.declare\_term}\verb|Variable.declare_term: term -> Proof.context -> Proof.context| \\ - \indexml{Variable.declare\_constraints}\verb|Variable.declare_constraints: term -> Proof.context -> Proof.context| \\ - \indexml{Variable.export}\verb|Variable.export: Proof.context -> Proof.context -> thm list -> thm list| \\ - \indexml{Variable.polymorphic}\verb|Variable.polymorphic: Proof.context -> term list -> term list| \\ - \indexml{Variable.import\_thms}\verb|Variable.import_thms: bool -> thm list -> Proof.context ->|\isasep\isanewline% -\verb| ((ctyp list * cterm list) * thm list) * Proof.context| \\ - \indexml{Variable.focus}\verb|Variable.focus: cterm -> Proof.context -> (cterm list * cterm) * Proof.context| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Variable.add_fixes|~\isa{xs\ ctxt} fixes term - variables \isa{xs}, returning the resulting internal names. By - default, the internal representation coincides with the external - one, which also means that the given variables must not be fixed - already. There is a different policy within a local proof body: the - given names are just hints for newly invented Skolem variables. - - \item \verb|Variable.variant_fixes| is similar to \verb|Variable.add_fixes|, but always produces fresh variants of the given - names. - - \item \verb|Variable.declare_term|~\isa{t\ ctxt} declares term - \isa{t} to belong to the context. This automatically fixes new - type variables, but not term variables. Syntactic constraints for - type and term variables are declared uniformly, though. - - \item \verb|Variable.declare_constraints|~\isa{t\ ctxt} declares - syntactic constraints from term \isa{t}, without making it part - of the context yet. - - \item \verb|Variable.export|~\isa{inner\ outer\ thms} generalizes - fixed type and term variables in \isa{thms} according to the - difference of the \isa{inner} and \isa{outer} context, - following the principles sketched above. - - \item \verb|Variable.polymorphic|~\isa{ctxt\ ts} generalizes type - variables in \isa{ts} as far as possible, even those occurring - in fixed term variables. The default policy of type-inference is to - fix newly introduced type variables, which is essentially reversed - with \verb|Variable.polymorphic|: here the given terms are detached - from the context as far as possible. - - \item \verb|Variable.import_thms|~\isa{open\ thms\ ctxt} invents fixed - type and term variables for the schematic ones occurring in \isa{thms}. The \isa{open} flag indicates whether the fixed names - should be accessible to the user, otherwise newly introduced names - are marked as ``internal'' (\secref{sec:names}). - - \item \verb|Variable.focus|~\isa{B} decomposes the outermost \isa{{\isasymAnd}} prefix of proposition \isa{B}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsection{Assumptions \label{sec:assumptions}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -An \emph{assumption} is a proposition that it is postulated in the - current context. Local conclusions may use assumptions as - additional facts, but this imposes implicit hypotheses that weaken - the overall statement. - - Assumptions are restricted to fixed non-schematic statements, i.e.\ - all generality needs to be expressed by explicit quantifiers. - Nevertheless, the result will be in HHF normal form with outermost - quantifiers stripped. For example, by assuming \isa{{\isasymAnd}x\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}{\isachardot}\ P\ x} we get \isa{{\isasymAnd}x\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}{\isachardot}\ P\ x\ {\isasymturnstile}\ P\ {\isacharquery}x} for schematic \isa{{\isacharquery}x} - of fixed type \isa{{\isasymalpha}}. Local derivations accumulate more and - more explicit references to hypotheses: \isa{A\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ A\isactrlisub n\ {\isasymturnstile}\ B} where \isa{A\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ A\isactrlisub n} needs to - be covered by the assumptions of the current context. - - \medskip The \isa{add{\isacharunderscore}assms} operation augments the context by - local assumptions, which are parameterized by an arbitrary \isa{export} rule (see below). - - The \isa{export} operation moves facts from a (larger) inner - context into a (smaller) outer context, by discharging the - difference of the assumptions as specified by the associated export - rules. Note that the discharged portion is determined by the - difference contexts, not the facts being exported! There is a - separate flag to indicate a goal context, where the result is meant - to refine an enclosing sub-goal of a structured proof state (cf.\ - \secref{sec:isar-proof-state}). - - \medskip The most basic export rule discharges assumptions directly - by means of the \isa{{\isasymLongrightarrow}} introduction rule: - \[ - \infer[(\isa{{\isasymLongrightarrow}{\isacharunderscore}intro})]{\isa{{\isasymGamma}\ {\isacharbackslash}\ A\ {\isasymturnstile}\ A\ {\isasymLongrightarrow}\ B}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B}} - \] - - The variant for goal refinements marks the newly introduced - premises, which causes the canonical Isar goal refinement scheme to - enforce unification with local premises within the goal: - \[ - \infer[(\isa{{\isacharhash}{\isasymLongrightarrow}{\isacharunderscore}intro})]{\isa{{\isasymGamma}\ {\isacharbackslash}\ A\ {\isasymturnstile}\ {\isacharhash}A\ {\isasymLongrightarrow}\ B}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B}} - \] - - \medskip Alternative versions of assumptions may perform arbitrary - transformations on export, as long as the corresponding portion of - hypotheses is removed from the given facts. For example, a local - definition works by fixing \isa{x} and assuming \isa{x\ {\isasymequiv}\ t}, - with the following export rule to reverse the effect: - \[ - \infer[(\isa{{\isasymequiv}{\isacharminus}expand})]{\isa{{\isasymGamma}\ {\isacharbackslash}\ x\ {\isasymequiv}\ t\ {\isasymturnstile}\ B\ t}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B\ x}} - \] - This works, because the assumption \isa{x\ {\isasymequiv}\ t} was introduced in - a context with \isa{x} being fresh, so \isa{x} does not - occur in \isa{{\isasymGamma}} here.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexmltype{Assumption.export}\verb|type Assumption.export| \\ - \indexml{Assumption.assume}\verb|Assumption.assume: cterm -> thm| \\ - \indexml{Assumption.add\_assms}\verb|Assumption.add_assms: Assumption.export ->|\isasep\isanewline% -\verb| cterm list -> Proof.context -> thm list * Proof.context| \\ - \indexml{Assumption.add\_assumes}\verb|Assumption.add_assumes: |\isasep\isanewline% -\verb| cterm list -> Proof.context -> thm list * Proof.context| \\ - \indexml{Assumption.export}\verb|Assumption.export: bool -> Proof.context -> Proof.context -> thm -> thm| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Assumption.export| represents arbitrary export - rules, which is any function of type \verb|bool -> cterm list -> thm -> thm|, - where the \verb|bool| indicates goal mode, and the \verb|cterm list| the collection of assumptions to be discharged - simultaneously. - - \item \verb|Assumption.assume|~\isa{A} turns proposition \isa{A} into a raw assumption \isa{A\ {\isasymturnstile}\ A{\isacharprime}}, where the conclusion - \isa{A{\isacharprime}} is in HHF normal form. - - \item \verb|Assumption.add_assms|~\isa{r\ As} augments the context - by assumptions \isa{As} with export rule \isa{r}. The - resulting facts are hypothetical theorems as produced by the raw - \verb|Assumption.assume|. - - \item \verb|Assumption.add_assumes|~\isa{As} is a special case of - \verb|Assumption.add_assms| where the export rule performs \isa{{\isasymLongrightarrow}{\isacharunderscore}intro} or \isa{{\isacharhash}{\isasymLongrightarrow}{\isacharunderscore}intro}, depending on goal mode. - - \item \verb|Assumption.export|~\isa{is{\isacharunderscore}goal\ inner\ outer\ thm} - exports result \isa{thm} from the the \isa{inner} context - back into the \isa{outer} one; \isa{is{\isacharunderscore}goal\ {\isacharequal}\ true} means - this is a goal context. The result is in HHF normal form. Note - that \verb|ProofContext.export| combines \verb|Variable.export| - and \verb|Assumption.export| in the canonical way. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsection{Results \label{sec:results}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Local results are established by monotonic reasoning from facts - within a context. This allows common combinations of theorems, - e.g.\ via \isa{{\isasymAnd}{\isacharslash}{\isasymLongrightarrow}} elimination, resolution rules, or equational - reasoning, see \secref{sec:thms}. Unaccounted context manipulations - should be avoided, notably raw \isa{{\isasymAnd}{\isacharslash}{\isasymLongrightarrow}} introduction or ad-hoc - references to free variables or assumptions not present in the proof - context. - - \medskip The \isa{SUBPROOF} combinator allows to structure a - tactical proof recursively by decomposing a selected sub-goal: - \isa{{\isacharparenleft}{\isasymAnd}x{\isachardot}\ A{\isacharparenleft}x{\isacharparenright}\ {\isasymLongrightarrow}\ B{\isacharparenleft}x{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymdots}} is turned into \isa{B{\isacharparenleft}x{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymdots}} - after fixing \isa{x} and assuming \isa{A{\isacharparenleft}x{\isacharparenright}}. This means - the tactic needs to solve the conclusion, but may use the premise as - a local fact, for locally fixed variables. - - The \isa{prove} operation provides an interface for structured - backwards reasoning under program control, with some explicit sanity - checks of the result. The goal context can be augmented by - additional fixed variables (cf.\ \secref{sec:variables}) and - assumptions (cf.\ \secref{sec:assumptions}), which will be available - as local facts during the proof and discharged into implications in - the result. Type and term variables are generalized as usual, - according to the context. - - The \isa{obtain} operation produces results by eliminating - existing facts by means of a given tactic. This acts like a dual - conclusion: the proof demonstrates that the context may be augmented - by certain fixed variables and assumptions. See also - \cite{isabelle-isar-ref} for the user-level \isa{{\isasymOBTAIN}} and - \isa{{\isasymGUESS}} elements. Final results, which may not refer to - the parameters in the conclusion, need to exported explicitly into - the original context.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{SUBPROOF}\verb|SUBPROOF: ({context: Proof.context, schematics: ctyp list * cterm list,|\isasep\isanewline% -\verb| params: cterm list, asms: cterm list, concl: cterm,|\isasep\isanewline% -\verb| prems: thm list} -> tactic) -> Proof.context -> int -> tactic| \\ - \end{mldecls} - \begin{mldecls} - \indexml{Goal.prove}\verb|Goal.prove: Proof.context -> string list -> term list -> term ->|\isasep\isanewline% -\verb| ({prems: thm list, context: Proof.context} -> tactic) -> thm| \\ - \indexml{Goal.prove\_multi}\verb|Goal.prove_multi: Proof.context -> string list -> term list -> term list ->|\isasep\isanewline% -\verb| ({prems: thm list, context: Proof.context} -> tactic) -> thm list| \\ - \end{mldecls} - \begin{mldecls} - \indexml{Obtain.result}\verb|Obtain.result: (Proof.context -> tactic) ->|\isasep\isanewline% -\verb| thm list -> Proof.context -> (cterm list * thm list) * Proof.context| \\ - \end{mldecls} - - \begin{description} - - \item \verb|SUBPROOF|~\isa{tac} decomposes the structure of a - particular sub-goal, producing an extended context and a reduced - goal, which needs to be solved by the given tactic. All schematic - parameters of the goal are imported into the context as fixed ones, - which may not be instantiated in the sub-proof. - - \item \verb|Goal.prove|~\isa{ctxt\ xs\ As\ C\ tac} states goal \isa{C} in the context augmented by fixed variables \isa{xs} and - assumptions \isa{As}, and applies tactic \isa{tac} to solve - it. The latter may depend on the local assumptions being presented - as facts. The result is in HHF normal form. - - \item \verb|Goal.prove_multi| is simular to \verb|Goal.prove|, but - states several conclusions simultaneously. The goal is encoded by - means of Pure conjunction; \verb|Goal.conjunction_tac| will turn this - into a collection of individual subgoals. - - \item \verb|Obtain.result|~\isa{tac\ thms\ ctxt} eliminates the - given facts using a tactic, which results in additional fixed - variables and assumptions in the context. Final results need to be - exported explicitly. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/document/session.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/document/tactic.tex --- a/doc-src/IsarImplementation/Thy/document/tactic.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,512 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{tactic}% -% -\isadelimtheory -\isanewline -\isanewline -\isanewline -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ tactic\ \isakeyword{imports}\ base\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Tactical reasoning% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Tactical reasoning works by refining the initial claim in a - backwards fashion, until a solved form is reached. A \isa{goal} - consists of several subgoals that need to be solved in order to - achieve the main statement; zero subgoals means that the proof may - be finished. A \isa{tactic} is a refinement operation that maps - a goal to a lazy sequence of potential successors. A \isa{tactical} is a combinator for composing tactics.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Goals \label{sec:tactical-goals}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle/Pure represents a goal\glossary{Tactical goal}{A theorem of - \seeglossary{Horn Clause} form stating that a number of subgoals - imply the main conclusion, which is marked as a protected - proposition.} as a theorem stating that the subgoals imply the main - goal: \isa{A\isactrlsub {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isasymdots}\ {\isasymLongrightarrow}\ A\isactrlsub n\ {\isasymLongrightarrow}\ C}. The outermost goal - structure is that of a Horn Clause\glossary{Horn Clause}{An iterated - implication \isa{A\isactrlsub {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isasymdots}\ {\isasymLongrightarrow}\ A\isactrlsub n\ {\isasymLongrightarrow}\ C}, without any - outermost quantifiers. Strictly speaking, propositions \isa{A\isactrlsub i} need to be atomic in Horn Clauses, but Isabelle admits - arbitrary substructure here (nested \isa{{\isasymLongrightarrow}} and \isa{{\isasymAnd}} - connectives).}: i.e.\ an iterated implication without any - quantifiers\footnote{Recall that outermost \isa{{\isasymAnd}x{\isachardot}\ {\isasymphi}{\isacharbrackleft}x{\isacharbrackright}} is - always represented via schematic variables in the body: \isa{{\isasymphi}{\isacharbrackleft}{\isacharquery}x{\isacharbrackright}}. These variables may get instantiated during the course of - reasoning.}. For \isa{n\ {\isacharequal}\ {\isadigit{0}}} a goal is called ``solved''. - - The structure of each subgoal \isa{A\isactrlsub i} is that of a general - Hereditary Harrop Formula \isa{{\isasymAnd}x\isactrlsub {\isadigit{1}}\ {\isasymdots}\ {\isasymAnd}x\isactrlsub k{\isachardot}\ H\isactrlsub {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isasymdots}\ {\isasymLongrightarrow}\ H\isactrlsub m\ {\isasymLongrightarrow}\ B} in - normal form. Here \isa{x\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ x\isactrlsub k} are goal parameters, i.e.\ - arbitrary-but-fixed entities of certain types, and \isa{H\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ H\isactrlsub m} are goal hypotheses, i.e.\ facts that may be assumed locally. - Together, this forms the goal context of the conclusion \isa{B} to - be established. The goal hypotheses may be again arbitrary - Hereditary Harrop Formulas, although the level of nesting rarely - exceeds 1--2 in practice. - - The main conclusion \isa{C} is internally marked as a protected - proposition\glossary{Protected proposition}{An arbitrarily - structured proposition \isa{C} which is forced to appear as - atomic by wrapping it into a propositional identity operator; - notation \isa{{\isacharhash}C}. Protecting a proposition prevents basic - inferences from entering into that structure for the time being.}, - which is represented explicitly by the notation \isa{{\isacharhash}C}. This - ensures that the decomposition into subgoals and main conclusion is - well-defined for arbitrarily structured claims. - - \medskip Basic goal management is performed via the following - Isabelle/Pure rules: - - \[ - \infer[\isa{{\isacharparenleft}init{\isacharparenright}}]{\isa{C\ {\isasymLongrightarrow}\ {\isacharhash}C}}{} \qquad - \infer[\isa{{\isacharparenleft}finish{\isacharparenright}}]{\isa{C}}{\isa{{\isacharhash}C}} - \] - - \medskip The following low-level variants admit general reasoning - with protected propositions: - - \[ - \infer[\isa{{\isacharparenleft}protect{\isacharparenright}}]{\isa{{\isacharhash}C}}{\isa{C}} \qquad - \infer[\isa{{\isacharparenleft}conclude{\isacharparenright}}]{\isa{A\isactrlsub {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isasymdots}\ {\isasymLongrightarrow}\ A\isactrlsub n\ {\isasymLongrightarrow}\ C}}{\isa{A\isactrlsub {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isasymdots}\ {\isasymLongrightarrow}\ A\isactrlsub n\ {\isasymLongrightarrow}\ {\isacharhash}C}} - \]% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Goal.init}\verb|Goal.init: cterm -> thm| \\ - \indexml{Goal.finish}\verb|Goal.finish: thm -> thm| \\ - \indexml{Goal.protect}\verb|Goal.protect: thm -> thm| \\ - \indexml{Goal.conclude}\verb|Goal.conclude: thm -> thm| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Goal.init|~\isa{C} initializes a tactical goal from - the well-formed proposition \isa{C}. - - \item \verb|Goal.finish|~\isa{thm} checks whether theorem - \isa{thm} is a solved goal (no subgoals), and concludes the - result by removing the goal protection. - - \item \verb|Goal.protect|~\isa{thm} protects the full statement - of theorem \isa{thm}. - - \item \verb|Goal.conclude|~\isa{thm} removes the goal - protection, even if there are pending subgoals. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsection{Tactics% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A \isa{tactic} is a function \isa{goal\ {\isasymrightarrow}\ goal\isactrlsup {\isacharasterisk}\isactrlsup {\isacharasterisk}} that - maps a given goal state (represented as a theorem, cf.\ - \secref{sec:tactical-goals}) to a lazy sequence of potential - successor states. The underlying sequence implementation is lazy - both in head and tail, and is purely functional in \emph{not} - supporting memoing.\footnote{The lack of memoing and the strict - nature of SML requires some care when working with low-level - sequence operations, to avoid duplicate or premature evaluation of - results.} - - An \emph{empty result sequence} means that the tactic has failed: in - a compound tactic expressions other tactics might be tried instead, - or the whole refinement step might fail outright, producing a - toplevel error message. When implementing tactics from scratch, one - should take care to observe the basic protocol of mapping regular - error conditions to an empty result; only serious faults should - emerge as exceptions. - - By enumerating \emph{multiple results}, a tactic can easily express - the potential outcome of an internal search process. There are also - combinators for building proof tools that involve search - systematically, see also \secref{sec:tacticals}. - - \medskip As explained in \secref{sec:tactical-goals}, a goal state - essentially consists of a list of subgoals that imply the main goal - (conclusion). Tactics may operate on all subgoals or on a - particularly specified subgoal, but must not change the main - conclusion (apart from instantiating schematic goal variables). - - Tactics with explicit \emph{subgoal addressing} are of the form - \isa{int\ {\isasymrightarrow}\ tactic} and may be applied to a particular subgoal - (counting from 1). If the subgoal number is out of range, the - tactic should fail with an empty result sequence, but must not raise - an exception! - - Operating on a particular subgoal means to replace it by an interval - of zero or more subgoals in the same place; other subgoals must not - be affected, apart from instantiating schematic variables ranging - over the whole goal state. - - A common pattern of composing tactics with subgoal addressing is to - try the first one, and then the second one only if the subgoal has - not been solved yet. Special care is required here to avoid bumping - into unrelated subgoals that happen to come after the original - subgoal. Assuming that there is only a single initial subgoal is a - very common error when implementing tactics! - - Tactics with internal subgoal addressing should expose the subgoal - index as \isa{int} argument in full generality; a hardwired - subgoal 1 inappropriate. - - \medskip The main well-formedness conditions for proper tactics are - summarized as follows. - - \begin{itemize} - - \item General tactic failure is indicated by an empty result, only - serious faults may produce an exception. - - \item The main conclusion must not be changed, apart from - instantiating schematic variables. - - \item A tactic operates either uniformly on all subgoals, or - specifically on a selected subgoal (without bumping into unrelated - subgoals). - - \item Range errors in subgoal addressing produce an empty result. - - \end{itemize} - - Some of these conditions are checked by higher-level goal - infrastructure (\secref{sec:results}); others are not checked - explicitly, and violating them merely results in ill-behaved tactics - experienced by the user (e.g.\ tactics that insist in being - applicable only to singleton goals, or disallow composition with - basic tacticals).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexmltype{tactic}\verb|type tactic = thm -> thm Seq.seq| \\ - \indexml{no\_tac}\verb|no_tac: tactic| \\ - \indexml{all\_tac}\verb|all_tac: tactic| \\ - \indexml{print\_tac}\verb|print_tac: string -> tactic| \\[1ex] - \indexml{PRIMITIVE}\verb|PRIMITIVE: (thm -> thm) -> tactic| \\[1ex] - \indexml{SUBGOAL}\verb|SUBGOAL: (term * int -> tactic) -> int -> tactic| \\ - \indexml{CSUBGOAL}\verb|CSUBGOAL: (cterm * int -> tactic) -> int -> tactic| \\ - \end{mldecls} - - \begin{description} - - \item \verb|tactic| represents tactics. The well-formedness - conditions described above need to be observed. See also \hyperlink{file.~~/src/Pure/General/seq.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}General{\isacharslash}seq{\isachardot}ML}}}} for the underlying implementation of - lazy sequences. - - \item \verb|int -> tactic| represents tactics with explicit - subgoal addressing, with well-formedness conditions as described - above. - - \item \verb|no_tac| is a tactic that always fails, returning the - empty sequence. - - \item \verb|all_tac| is a tactic that always succeeds, returning a - singleton sequence with unchanged goal state. - - \item \verb|print_tac|~\isa{message} is like \verb|all_tac|, but - prints a message together with the goal state on the tracing - channel. - - \item \verb|PRIMITIVE|~\isa{rule} turns a primitive inference rule - into a tactic with unique result. Exception \verb|THM| is considered - a regular tactic failure and produces an empty result; other - exceptions are passed through. - - \item \verb|SUBGOAL|~\isa{{\isacharparenleft}fn\ {\isacharparenleft}subgoal{\isacharcomma}\ i{\isacharparenright}\ {\isacharequal}{\isachargreater}\ tactic{\isacharparenright}} is the - most basic form to produce a tactic with subgoal addressing. The - given abstraction over the subgoal term and subgoal number allows to - peek at the relevant information of the full goal state. The - subgoal range is checked as required above. - - \item \verb|CSUBGOAL| is similar to \verb|SUBGOAL|, but passes the - subgoal as \verb|cterm| instead of raw \verb|term|. This - avoids expensive re-certification in situations where the subgoal is - used directly for primitive inferences. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Resolution and assumption tactics \label{sec:resolve-assume-tac}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\emph{Resolution} is the most basic mechanism for refining a - subgoal using a theorem as object-level rule. - \emph{Elim-resolution} is particularly suited for elimination rules: - it resolves with a rule, proves its first premise by assumption, and - finally deletes that assumption from any new subgoals. - \emph{Destruct-resolution} is like elim-resolution, but the given - destruction rules are first turned into canonical elimination - format. \emph{Forward-resolution} is like destruct-resolution, but - without deleting the selected assumption. The \isa{r{\isacharslash}e{\isacharslash}d{\isacharslash}f} - naming convention is maintained for several different kinds of - resolution rules and tactics. - - Assumption tactics close a subgoal by unifying some of its premises - against its conclusion. - - \medskip All the tactics in this section operate on a subgoal - designated by a positive integer. Other subgoals might be affected - indirectly, due to instantiation of schematic variables. - - There are various sources of non-determinism, the tactic result - sequence enumerates all possibilities of the following choices (if - applicable): - - \begin{enumerate} - - \item selecting one of the rules given as argument to the tactic; - - \item selecting a subgoal premise to eliminate, unifying it against - the first premise of the rule; - - \item unifying the conclusion of the subgoal to the conclusion of - the rule. - - \end{enumerate} - - Recall that higher-order unification may produce multiple results - that are enumerated here.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{resolve\_tac}\verb|resolve_tac: thm list -> int -> tactic| \\ - \indexml{eresolve\_tac}\verb|eresolve_tac: thm list -> int -> tactic| \\ - \indexml{dresolve\_tac}\verb|dresolve_tac: thm list -> int -> tactic| \\ - \indexml{forward\_tac}\verb|forward_tac: thm list -> int -> tactic| \\[1ex] - \indexml{assume\_tac}\verb|assume_tac: int -> tactic| \\ - \indexml{eq\_assume\_tac}\verb|eq_assume_tac: int -> tactic| \\[1ex] - \indexml{match\_tac}\verb|match_tac: thm list -> int -> tactic| \\ - \indexml{ematch\_tac}\verb|ematch_tac: thm list -> int -> tactic| \\ - \indexml{dmatch\_tac}\verb|dmatch_tac: thm list -> int -> tactic| \\ - \end{mldecls} - - \begin{description} - - \item \verb|resolve_tac|~\isa{thms\ i} refines the goal state - using the given theorems, which should normally be introduction - rules. The tactic resolves a rule's conclusion with subgoal \isa{i}, replacing it by the corresponding versions of the rule's - premises. - - \item \verb|eresolve_tac|~\isa{thms\ i} performs elim-resolution - with the given theorems, which should normally be elimination rules. - - \item \verb|dresolve_tac|~\isa{thms\ i} performs - destruct-resolution with the given theorems, which should normally - be destruction rules. This replaces an assumption by the result of - applying one of the rules. - - \item \verb|forward_tac| is like \verb|dresolve_tac| except that the - selected assumption is not deleted. It applies a rule to an - assumption, adding the result as a new assumption. - - \item \verb|assume_tac|~\isa{i} attempts to solve subgoal \isa{i} - by assumption (modulo higher-order unification). - - \item \verb|eq_assume_tac| is similar to \verb|assume_tac|, but checks - only for immediate \isa{{\isasymalpha}}-convertibility instead of using - unification. It succeeds (with a unique next state) if one of the - assumptions is equal to the subgoal's conclusion. Since it does not - instantiate variables, it cannot make other subgoals unprovable. - - \item \verb|match_tac|, \verb|ematch_tac|, and \verb|dmatch_tac| are - similar to \verb|resolve_tac|, \verb|eresolve_tac|, and \verb|dresolve_tac|, respectively, but do not instantiate schematic - variables in the goal state. - - Flexible subgoals are not updated at will, but are left alone. - Strictly speaking, matching means to treat the unknowns in the goal - state as constants; these tactics merely discard unifiers that would - update the goal state. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Explicit instantiation within a subgoal context% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The main resolution tactics (\secref{sec:resolve-assume-tac}) - use higher-order unification, which works well in many practical - situations despite its daunting theoretical properties. - Nonetheless, there are important problem classes where unguided - higher-order unification is not so useful. This typically involves - rules like universal elimination, existential introduction, or - equational substitution. Here the unification problem involves - fully flexible \isa{{\isacharquery}P\ {\isacharquery}x} schemes, which are hard to manage - without further hints. - - By providing a (small) rigid term for \isa{{\isacharquery}x} explicitly, the - remaining unification problem is to assign a (large) term to \isa{{\isacharquery}P}, according to the shape of the given subgoal. This is - sufficiently well-behaved in most practical situations. - - \medskip Isabelle provides separate versions of the standard \isa{r{\isacharslash}e{\isacharslash}d{\isacharslash}f} resolution tactics that allow to provide explicit - instantiations of unknowns of the given rule, wrt.\ terms that refer - to the implicit context of the selected subgoal. - - An instantiation consists of a list of pairs of the form \isa{{\isacharparenleft}{\isacharquery}x{\isacharcomma}\ t{\isacharparenright}}, where \isa{{\isacharquery}x} is a schematic variable occurring in - the given rule, and \isa{t} is a term from the current proof - context, augmented by the local goal parameters of the selected - subgoal; cf.\ the \isa{focus} operation described in - \secref{sec:variables}. - - Entering the syntactic context of a subgoal is a brittle operation, - because its exact form is somewhat accidental, and the choice of - bound variable names depends on the presence of other local and - global names. Explicit renaming of subgoal parameters prior to - explicit instantiation might help to achieve a bit more robustness. - - Type instantiations may be given as well, via pairs like \isa{{\isacharparenleft}{\isacharquery}{\isacharprime}a{\isacharcomma}\ {\isasymtau}{\isacharparenright}}. Type instantiations are distinguished from term - instantiations by the syntactic form of the schematic variable. - Types are instantiated before terms are. Since term instantiation - already performs type-inference as expected, explicit type - instantiations are seldom necessary.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{res\_inst\_tac}\verb|res_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic| \\ - \indexml{eres\_inst\_tac}\verb|eres_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic| \\ - \indexml{dres\_inst\_tac}\verb|dres_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic| \\ - \indexml{forw\_inst\_tac}\verb|forw_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic| \\[1ex] - \indexml{rename\_tac}\verb|rename_tac: string list -> int -> tactic| \\ - \end{mldecls} - - \begin{description} - - \item \verb|res_inst_tac|~\isa{ctxt\ insts\ thm\ i} instantiates the - rule \isa{thm} with the instantiations \isa{insts}, as described - above, and then performs resolution on subgoal \isa{i}. - - \item \verb|eres_inst_tac| is like \verb|res_inst_tac|, but performs - elim-resolution. - - \item \verb|dres_inst_tac| is like \verb|res_inst_tac|, but performs - destruct-resolution. - - \item \verb|forw_inst_tac| is like \verb|dres_inst_tac| except that - the selected assumption is not deleted. - - \item \verb|rename_tac|~\isa{names\ i} renames the innermost - parameters of subgoal \isa{i} according to the provided \isa{names} (which need to be distinct indentifiers). - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsection{Tacticals \label{sec:tacticals}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -FIXME - -\glossary{Tactical}{A functional combinator for building up complex -tactics from simpler ones. Typical tactical perform sequential -composition, disjunction (choice), iteration, or goal addressing. -Various search strategies may be expressed via tacticals.}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/integration.thy --- a/doc-src/IsarImplementation/Thy/integration.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,426 +0,0 @@ - -(* $Id$ *) - -theory integration imports base begin - -chapter {* System integration *} - -section {* Isar toplevel \label{sec:isar-toplevel} *} - -text {* The Isar toplevel may be considered the centeral hub of the - Isabelle/Isar system, where all key components and sub-systems are - integrated into a single read-eval-print loop of Isar commands. We - shall even incorporate the existing {\ML} toplevel of the compiler - and run-time system (cf.\ \secref{sec:ML-toplevel}). - - Isabelle/Isar departs from the original ``LCF system architecture'' - where {\ML} was really The Meta Language for defining theories and - conducting proofs. Instead, {\ML} now only serves as the - implementation language for the system (and user extensions), while - the specific Isar toplevel supports the concepts of theory and proof - development natively. This includes the graph structure of theories - and the block structure of proofs, support for unlimited undo, - facilities for tracing, debugging, timing, profiling etc. - - \medskip The toplevel maintains an implicit state, which is - transformed by a sequence of transitions -- either interactively or - in batch-mode. In interactive mode, Isar state transitions are - encapsulated as safe transactions, such that both failure and undo - are handled conveniently without destroying the underlying draft - theory (cf.~\secref{sec:context-theory}). In batch mode, - transitions operate in a linear (destructive) fashion, such that - error conditions abort the present attempt to construct a theory or - proof altogether. - - The toplevel state is a disjoint sum of empty @{text toplevel}, or - @{text theory}, or @{text proof}. On entering the main Isar loop we - start with an empty toplevel. A theory is commenced by giving a - @{text \} header; within a theory we may issue theory - commands such as @{text \}, or state a @{text - \} to be proven. Now we are within a proof state, with a - rich collection of Isar proof commands for structured proof - composition, or unstructured proof scripts. When the proof is - concluded we get back to the theory, which is then updated by - storing the resulting fact. Further theory declarations or theorem - statements with proofs may follow, until we eventually conclude the - theory development by issuing @{text \}. The resulting theory - is then stored within the theory database and we are back to the - empty toplevel. - - In addition to these proper state transformations, there are also - some diagnostic commands for peeking at the toplevel state without - modifying it (e.g.\ \isakeyword{thm}, \isakeyword{term}, - \isakeyword{print-cases}). -*} - -text %mlref {* - \begin{mldecls} - @{index_ML_type Toplevel.state} \\ - @{index_ML Toplevel.UNDEF: "exn"} \\ - @{index_ML Toplevel.is_toplevel: "Toplevel.state -> bool"} \\ - @{index_ML Toplevel.theory_of: "Toplevel.state -> theory"} \\ - @{index_ML Toplevel.proof_of: "Toplevel.state -> Proof.state"} \\ - @{index_ML Toplevel.debug: "bool ref"} \\ - @{index_ML Toplevel.timing: "bool ref"} \\ - @{index_ML Toplevel.profiling: "int ref"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML_type Toplevel.state} represents Isar toplevel states, - which are normally manipulated through the concept of toplevel - transitions only (\secref{sec:toplevel-transition}). Also note that - a raw toplevel state is subject to the same linearity restrictions - as a theory context (cf.~\secref{sec:context-theory}). - - \item @{ML Toplevel.UNDEF} is raised for undefined toplevel - operations. Many operations work only partially for certain cases, - since @{ML_type Toplevel.state} is a sum type. - - \item @{ML Toplevel.is_toplevel}~@{text "state"} checks for an empty - toplevel state. - - \item @{ML Toplevel.theory_of}~@{text "state"} selects the theory of - a theory or proof (!), otherwise raises @{ML Toplevel.UNDEF}. - - \item @{ML Toplevel.proof_of}~@{text "state"} selects the Isar proof - state if available, otherwise raises @{ML Toplevel.UNDEF}. - - \item @{ML "set Toplevel.debug"} makes the toplevel print further - details about internal error conditions, exceptions being raised - etc. - - \item @{ML "set Toplevel.timing"} makes the toplevel print timing - information for each Isar command being executed. - - \item @{ML Toplevel.profiling}~@{verbatim ":="}~@{text "n"} controls - low-level profiling of the underlying {\ML} runtime system. For - Poly/ML, @{text "n = 1"} means time and @{text "n = 2"} space - profiling. - - \end{description} -*} - - -subsection {* Toplevel transitions \label{sec:toplevel-transition} *} - -text {* - An Isar toplevel transition consists of a partial function on the - toplevel state, with additional information for diagnostics and - error reporting: there are fields for command name, source position, - optional source text, as well as flags for interactive-only commands - (which issue a warning in batch-mode), printing of result state, - etc. - - The operational part is represented as the sequential union of a - list of partial functions, which are tried in turn until the first - one succeeds. This acts like an outer case-expression for various - alternative state transitions. For example, \isakeyword{qed} acts - differently for a local proofs vs.\ the global ending of the main - proof. - - Toplevel transitions are composed via transition transformers. - Internally, Isar commands are put together from an empty transition - extended by name and source position (and optional source text). It - is then left to the individual command parser to turn the given - concrete syntax into a suitable transition transformer that adjoin - actual operations on a theory or proof state etc. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML Toplevel.print: "Toplevel.transition -> Toplevel.transition"} \\ - @{index_ML Toplevel.no_timing: "Toplevel.transition -> Toplevel.transition"} \\ - @{index_ML Toplevel.keep: "(Toplevel.state -> unit) -> - Toplevel.transition -> Toplevel.transition"} \\ - @{index_ML Toplevel.theory: "(theory -> theory) -> - Toplevel.transition -> Toplevel.transition"} \\ - @{index_ML Toplevel.theory_to_proof: "(theory -> Proof.state) -> - Toplevel.transition -> Toplevel.transition"} \\ - @{index_ML Toplevel.proof: "(Proof.state -> Proof.state) -> - Toplevel.transition -> Toplevel.transition"} \\ - @{index_ML Toplevel.proofs: "(Proof.state -> Proof.state Seq.seq) -> - Toplevel.transition -> Toplevel.transition"} \\ - @{index_ML Toplevel.end_proof: "(bool -> Proof.state -> Proof.context) -> - Toplevel.transition -> Toplevel.transition"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML Toplevel.print}~@{text "tr"} sets the print flag, which - causes the toplevel loop to echo the result state (in interactive - mode). - - \item @{ML Toplevel.no_timing}~@{text "tr"} indicates that the - transition should never show timing information, e.g.\ because it is - a diagnostic command. - - \item @{ML Toplevel.keep}~@{text "tr"} adjoins a diagnostic - function. - - \item @{ML Toplevel.theory}~@{text "tr"} adjoins a theory - transformer. - - \item @{ML Toplevel.theory_to_proof}~@{text "tr"} adjoins a global - goal function, which turns a theory into a proof state. The theory - may be changed before entering the proof; the generic Isar goal - setup includes an argument that specifies how to apply the proven - result to the theory, when the proof is finished. - - \item @{ML Toplevel.proof}~@{text "tr"} adjoins a deterministic - proof command, with a singleton result. - - \item @{ML Toplevel.proofs}~@{text "tr"} adjoins a general proof - command, with zero or more result states (represented as a lazy - list). - - \item @{ML Toplevel.end_proof}~@{text "tr"} adjoins a concluding - proof command, that returns the resulting theory, after storing the - resulting facts in the context etc. - - \end{description} -*} - - -subsection {* Toplevel control *} - -text {* - There are a few special control commands that modify the behavior - the toplevel itself, and only make sense in interactive mode. Under - normal circumstances, the user encounters these only implicitly as - part of the protocol between the Isabelle/Isar system and a - user-interface such as ProofGeneral. - - \begin{description} - - \item \isacommand{undo} follows the three-level hierarchy of empty - toplevel vs.\ theory vs.\ proof: undo within a proof reverts to the - previous proof context, undo after a proof reverts to the theory - before the initial goal statement, undo of a theory command reverts - to the previous theory value, undo of a theory header discontinues - the current theory development and removes it from the theory - database (\secref{sec:theory-database}). - - \item \isacommand{kill} aborts the current level of development: - kill in a proof context reverts to the theory before the initial - goal statement, kill in a theory context aborts the current theory - development, removing it from the database. - - \item \isacommand{exit} drops out of the Isar toplevel into the - underlying {\ML} toplevel (\secref{sec:ML-toplevel}). The Isar - toplevel state is preserved and may be continued later. - - \item \isacommand{quit} terminates the Isabelle/Isar process without - saving. - - \end{description} -*} - - -section {* ML toplevel \label{sec:ML-toplevel} *} - -text {* - The {\ML} toplevel provides a read-compile-eval-print loop for {\ML} - values, types, structures, and functors. {\ML} declarations operate - on the global system state, which consists of the compiler - environment plus the values of {\ML} reference variables. There is - no clean way to undo {\ML} declarations, except for reverting to a - previously saved state of the whole Isabelle process. {\ML} input - is either read interactively from a TTY, or from a string (usually - within a theory text), or from a source file (usually loaded from a - theory). - - Whenever the {\ML} toplevel is active, the current Isabelle theory - context is passed as an internal reference variable. Thus {\ML} - code may access the theory context during compilation, it may even - change the value of a theory being under construction --- while - observing the usual linearity restrictions - (cf.~\secref{sec:context-theory}). -*} - -text %mlref {* - \begin{mldecls} - @{index_ML the_context: "unit -> theory"} \\ - @{index_ML "Context.>> ": "(Context.generic -> Context.generic) -> unit"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML "the_context ()"} refers to the theory context of the - {\ML} toplevel --- at compile time! {\ML} code needs to take care - to refer to @{ML "the_context ()"} correctly. Recall that - evaluation of a function body is delayed until actual runtime. - Moreover, persistent {\ML} toplevel bindings to an unfinished theory - should be avoided: code should either project out the desired - information immediately, or produce an explicit @{ML_type - theory_ref} (cf.\ \secref{sec:context-theory}). - - \item @{ML "Context.>>"}~@{text f} applies context transformation - @{text f} to the implicit context of the {\ML} toplevel. - - \end{description} - - It is very important to note that the above functions are really - restricted to the compile time, even though the {\ML} compiler is - invoked at runtime! The majority of {\ML} code uses explicit - functional arguments of a theory or proof context instead. Thus it - may be invoked for an arbitrary context later on, without having to - worry about any operational details. - - \bigskip - - \begin{mldecls} - @{index_ML Isar.main: "unit -> unit"} \\ - @{index_ML Isar.loop: "unit -> unit"} \\ - @{index_ML Isar.state: "unit -> Toplevel.state"} \\ - @{index_ML Isar.exn: "unit -> (exn * string) option"} \\ - @{index_ML Isar.context: "unit -> Proof.context"} \\ - @{index_ML Isar.goal: "unit -> thm"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML "Isar.main ()"} invokes the Isar toplevel from {\ML}, - initializing an empty toplevel state. - - \item @{ML "Isar.loop ()"} continues the Isar toplevel with the - current state, after having dropped out of the Isar toplevel loop. - - \item @{ML "Isar.state ()"} and @{ML "Isar.exn ()"} get current - toplevel state and error condition, respectively. This only works - after having dropped out of the Isar toplevel loop. - - \item @{ML "Isar.context ()"} produces the proof context from @{ML - "Isar.state ()"}, analogous to @{ML Context.proof_of} - (\secref{sec:generic-context}). - - \item @{ML "Isar.goal ()"} picks the tactical goal from @{ML - "Isar.state ()"}, represented as a theorem according to - \secref{sec:tactical-goals}. - - \end{description} -*} - - -section {* Theory database \label{sec:theory-database} *} - -text {* - The theory database maintains a collection of theories, together - with some administrative information about their original sources, - which are held in an external store (i.e.\ some directory within the - regular file system). - - The theory database is organized as a directed acyclic graph; - entries are referenced by theory name. Although some additional - interfaces allow to include a directory specification as well, this - is only a hint to the underlying theory loader. The internal theory - name space is flat! - - Theory @{text A} is associated with the main theory file @{text - A}\verb,.thy,, which needs to be accessible through the theory - loader path. Any number of additional {\ML} source files may be - associated with each theory, by declaring these dependencies in the - theory header as @{text \}, and loading them consecutively - within the theory context. The system keeps track of incoming {\ML} - sources and associates them with the current theory. The file - @{text A}\verb,.ML, is loaded after a theory has been concluded, in - order to support legacy proof {\ML} proof scripts. - - The basic internal actions of the theory database are @{text - "update"}, @{text "outdate"}, and @{text "remove"}: - - \begin{itemize} - - \item @{text "update A"} introduces a link of @{text "A"} with a - @{text "theory"} value of the same name; it asserts that the theory - sources are now consistent with that value; - - \item @{text "outdate A"} invalidates the link of a theory database - entry to its sources, but retains the present theory value; - - \item @{text "remove A"} deletes entry @{text "A"} from the theory - database. - - \end{itemize} - - These actions are propagated to sub- or super-graphs of a theory - entry as expected, in order to preserve global consistency of the - state of all loaded theories with the sources of the external store. - This implies certain causalities between actions: @{text "update"} - or @{text "outdate"} of an entry will @{text "outdate"} all - descendants; @{text "remove"} will @{text "remove"} all descendants. - - \medskip There are separate user-level interfaces to operate on the - theory database directly or indirectly. The primitive actions then - just happen automatically while working with the system. In - particular, processing a theory header @{text "\ A - \ B\<^sub>1 \ B\<^sub>n \"} ensures that the - sub-graph of the collective imports @{text "B\<^sub>1 \ B\<^sub>n"} - is up-to-date, too. Earlier theories are reloaded as required, with - @{text update} actions proceeding in topological order according to - theory dependencies. There may be also a wave of implied @{text - outdate} actions for derived theory nodes until a stable situation - is achieved eventually. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML theory: "string -> theory"} \\ - @{index_ML use_thy: "string -> unit"} \\ - @{index_ML use_thys: "string list -> unit"} \\ - @{index_ML ThyInfo.touch_thy: "string -> unit"} \\ - @{index_ML ThyInfo.remove_thy: "string -> unit"} \\[1ex] - @{index_ML ThyInfo.begin_theory}@{verbatim ": ... -> bool -> theory"} \\ - @{index_ML ThyInfo.end_theory: "theory -> unit"} \\ - @{index_ML ThyInfo.register_theory: "theory -> unit"} \\[1ex] - @{verbatim "datatype action = Update | Outdate | Remove"} \\ - @{index_ML ThyInfo.add_hook: "(ThyInfo.action -> string -> unit) -> unit"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML theory}~@{text A} retrieves the theory value presently - associated with name @{text A}. Note that the result might be - outdated. - - \item @{ML use_thy}~@{text A} ensures that theory @{text A} is fully - up-to-date wrt.\ the external file store, reloading outdated - ancestors as required. - - \item @{ML use_thys} is similar to @{ML use_thy}, but handles - several theories simultaneously. Thus it acts like processing the - import header of a theory, without performing the merge of the - result, though. - - \item @{ML ThyInfo.touch_thy}~@{text A} performs and @{text outdate} action - on theory @{text A} and all descendants. - - \item @{ML ThyInfo.remove_thy}~@{text A} deletes theory @{text A} and all - descendants from the theory database. - - \item @{ML ThyInfo.begin_theory} is the basic operation behind a - @{text \} header declaration. This is {\ML} functions is - normally not invoked directly. - - \item @{ML ThyInfo.end_theory} concludes the loading of a theory - proper and stores the result in the theory database. - - \item @{ML ThyInfo.register_theory}~@{text "text thy"} registers an - existing theory value with the theory loader database. There is no - management of associated sources. - - \item @{ML "ThyInfo.add_hook"}~@{text f} registers function @{text - f} as a hook for theory database actions. The function will be - invoked with the action and theory name being involved; thus derived - actions may be performed in associated system components, e.g.\ - maintaining the state of an editor for the theory sources. - - The kind and order of actions occurring in practice depends both on - user interactions and the internal process of resolving theory - imports. Hooks should not rely on a particular policy here! Any - exceptions raised by the hook are ignored. - - \end{description} -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/isar.thy --- a/doc-src/IsarImplementation/Thy/isar.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,41 +0,0 @@ - -(* $Id$ *) - -theory isar imports base begin - -chapter {* Isar proof texts *} - -section {* Proof context *} - -text FIXME - - -section {* Proof state \label{sec:isar-proof-state} *} - -text {* - FIXME - -\glossary{Proof state}{The whole configuration of a structured proof, -consisting of a \seeglossary{proof context} and an optional -\seeglossary{structured goal}. Internally, an Isar proof state is -organized as a stack to accomodate block structure of proof texts. -For historical reasons, a low-level \seeglossary{tactical goal} is -occasionally called ``proof state'' as well.} - -\glossary{Structured goal}{FIXME} - -\glossary{Goal}{See \seeglossary{tactical goal} or \seeglossary{structured goal}. \norefpage} - - -*} - -section {* Proof methods *} - -text FIXME - -section {* Attributes *} - -text "FIXME ?!" - - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/locale.thy --- a/doc-src/IsarImplementation/Thy/locale.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,26 +0,0 @@ - -(* $Id$ *) - -theory "locale" imports base begin - -chapter {* Structured specifications *} - -section {* Specification elements *} - -text FIXME - - -section {* Type-inference *} - -text FIXME - - -section {* Local theories *} - -text {* - FIXME - - \glossary{Local theory}{FIXME} -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/logic.thy --- a/doc-src/IsarImplementation/Thy/logic.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,851 +0,0 @@ -theory logic imports base begin - -chapter {* Primitive logic \label{ch:logic} *} - -text {* - The logical foundations of Isabelle/Isar are that of the Pure logic, - which has been introduced as a natural-deduction framework in - \cite{paulson700}. This is essentially the same logic as ``@{text - "\HOL"}'' in the more abstract setting of Pure Type Systems (PTS) - \cite{Barendregt-Geuvers:2001}, although there are some key - differences in the specific treatment of simple types in - Isabelle/Pure. - - Following type-theoretic parlance, the Pure logic consists of three - levels of @{text "\"}-calculus with corresponding arrows, @{text - "\"} for syntactic function space (terms depending on terms), @{text - "\"} for universal quantification (proofs depending on terms), and - @{text "\"} for implication (proofs depending on proofs). - - Derivations are relative to a logical theory, which declares type - constructors, constants, and axioms. Theory declarations support - schematic polymorphism, which is strictly speaking outside the - logic.\footnote{This is the deeper logical reason, why the theory - context @{text "\"} is separate from the proof context @{text "\"} - of the core calculus.} -*} - - -section {* Types \label{sec:types} *} - -text {* - The language of types is an uninterpreted order-sorted first-order - algebra; types are qualified by ordered type classes. - - \medskip A \emph{type class} is an abstract syntactic entity - declared in the theory context. The \emph{subclass relation} @{text - "c\<^isub>1 \ c\<^isub>2"} is specified by stating an acyclic - generating relation; the transitive closure is maintained - internally. The resulting relation is an ordering: reflexive, - transitive, and antisymmetric. - - A \emph{sort} is a list of type classes written as @{text "s = - {c\<^isub>1, \, c\<^isub>m}"}, which represents symbolic - intersection. Notationally, the curly braces are omitted for - singleton intersections, i.e.\ any class @{text "c"} may be read as - a sort @{text "{c}"}. The ordering on type classes is extended to - sorts according to the meaning of intersections: @{text - "{c\<^isub>1, \ c\<^isub>m} \ {d\<^isub>1, \, d\<^isub>n}"} iff - @{text "\j. \i. c\<^isub>i \ d\<^isub>j"}. The empty intersection - @{text "{}"} refers to the universal sort, which is the largest - element wrt.\ the sort order. The intersections of all (finitely - many) classes declared in the current theory are the minimal - elements wrt.\ the sort order. - - \medskip A \emph{fixed type variable} is a pair of a basic name - (starting with a @{text "'"} character) and a sort constraint, e.g.\ - @{text "('a, s)"} which is usually printed as @{text "\\<^isub>s"}. - A \emph{schematic type variable} is a pair of an indexname and a - sort constraint, e.g.\ @{text "(('a, 0), s)"} which is usually - printed as @{text "?\\<^isub>s"}. - - Note that \emph{all} syntactic components contribute to the identity - of type variables, including the sort constraint. The core logic - handles type variables with the same name but different sorts as - different, although some outer layers of the system make it hard to - produce anything like this. - - A \emph{type constructor} @{text "\"} is a @{text "k"}-ary operator - on types declared in the theory. Type constructor application is - written postfix as @{text "(\\<^isub>1, \, \\<^isub>k)\"}. For - @{text "k = 0"} the argument tuple is omitted, e.g.\ @{text "prop"} - instead of @{text "()prop"}. For @{text "k = 1"} the parentheses - are omitted, e.g.\ @{text "\ list"} instead of @{text "(\)list"}. - Further notation is provided for specific constructors, notably the - right-associative infix @{text "\ \ \"} instead of @{text "(\, - \)fun"}. - - A \emph{type} is defined inductively over type variables and type - constructors as follows: @{text "\ = \\<^isub>s | ?\\<^isub>s | - (\\<^sub>1, \, \\<^sub>k)\"}. - - A \emph{type abbreviation} is a syntactic definition @{text - "(\<^vec>\)\ = \"} of an arbitrary type expression @{text "\"} over - variables @{text "\<^vec>\"}. Type abbreviations appear as type - constructors in the syntax, but are expanded before entering the - logical core. - - A \emph{type arity} declares the image behavior of a type - constructor wrt.\ the algebra of sorts: @{text "\ :: (s\<^isub>1, \, - s\<^isub>k)s"} means that @{text "(\\<^isub>1, \, \\<^isub>k)\"} is - of sort @{text "s"} if every argument type @{text "\\<^isub>i"} is - of sort @{text "s\<^isub>i"}. Arity declarations are implicitly - completed, i.e.\ @{text "\ :: (\<^vec>s)c"} entails @{text "\ :: - (\<^vec>s)c'"} for any @{text "c' \ c"}. - - \medskip The sort algebra is always maintained as \emph{coregular}, - which means that type arities are consistent with the subclass - relation: for any type constructor @{text "\"}, and classes @{text - "c\<^isub>1 \ c\<^isub>2"}, and arities @{text "\ :: - (\<^vec>s\<^isub>1)c\<^isub>1"} and @{text "\ :: - (\<^vec>s\<^isub>2)c\<^isub>2"} holds @{text "\<^vec>s\<^isub>1 \ - \<^vec>s\<^isub>2"} component-wise. - - The key property of a coregular order-sorted algebra is that sort - constraints can be solved in a most general fashion: for each type - constructor @{text "\"} and sort @{text "s"} there is a most general - vector of argument sorts @{text "(s\<^isub>1, \, s\<^isub>k)"} such - that a type scheme @{text "(\\<^bsub>s\<^isub>1\<^esub>, \, - \\<^bsub>s\<^isub>k\<^esub>)\"} is of sort @{text "s"}. - Consequently, type unification has most general solutions (modulo - equivalence of sorts), so type-inference produces primary types as - expected \cite{nipkow-prehofer}. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML_type class} \\ - @{index_ML_type sort} \\ - @{index_ML_type arity} \\ - @{index_ML_type typ} \\ - @{index_ML map_atyps: "(typ -> typ) -> typ -> typ"} \\ - @{index_ML fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\ - \end{mldecls} - \begin{mldecls} - @{index_ML Sign.subsort: "theory -> sort * sort -> bool"} \\ - @{index_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\ - @{index_ML Sign.add_types: "(string * int * mixfix) list -> theory -> theory"} \\ - @{index_ML Sign.add_tyabbrs_i: " - (string * string list * typ * mixfix) list -> theory -> theory"} \\ - @{index_ML Sign.primitive_class: "string * class list -> theory -> theory"} \\ - @{index_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\ - @{index_ML Sign.primitive_arity: "arity -> theory -> theory"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML_type class} represents type classes; this is an alias for - @{ML_type string}. - - \item @{ML_type sort} represents sorts; this is an alias for - @{ML_type "class list"}. - - \item @{ML_type arity} represents type arities; this is an alias for - triples of the form @{text "(\, \<^vec>s, s)"} for @{text "\ :: - (\<^vec>s)s"} described above. - - \item @{ML_type typ} represents types; this is a datatype with - constructors @{ML TFree}, @{ML TVar}, @{ML Type}. - - \item @{ML map_atyps}~@{text "f \"} applies the mapping @{text "f"} - to all atomic types (@{ML TFree}, @{ML TVar}) occurring in @{text - "\"}. - - \item @{ML fold_atyps}~@{text "f \"} iterates the operation @{text - "f"} over all occurrences of atomic types (@{ML TFree}, @{ML TVar}) - in @{text "\"}; the type structure is traversed from left to right. - - \item @{ML Sign.subsort}~@{text "thy (s\<^isub>1, s\<^isub>2)"} - tests the subsort relation @{text "s\<^isub>1 \ s\<^isub>2"}. - - \item @{ML Sign.of_sort}~@{text "thy (\, s)"} tests whether type - @{text "\"} is of sort @{text "s"}. - - \item @{ML Sign.add_types}~@{text "[(\, k, mx), \]"} declares a new - type constructors @{text "\"} with @{text "k"} arguments and - optional mixfix syntax. - - \item @{ML Sign.add_tyabbrs_i}~@{text "[(\, \<^vec>\, \, mx), \]"} - defines a new type abbreviation @{text "(\<^vec>\)\ = \"} with - optional mixfix syntax. - - \item @{ML Sign.primitive_class}~@{text "(c, [c\<^isub>1, \, - c\<^isub>n])"} declares a new class @{text "c"}, together with class - relations @{text "c \ c\<^isub>i"}, for @{text "i = 1, \, n"}. - - \item @{ML Sign.primitive_classrel}~@{text "(c\<^isub>1, - c\<^isub>2)"} declares the class relation @{text "c\<^isub>1 \ - c\<^isub>2"}. - - \item @{ML Sign.primitive_arity}~@{text "(\, \<^vec>s, s)"} declares - the arity @{text "\ :: (\<^vec>s)s"}. - - \end{description} -*} - - - -section {* Terms \label{sec:terms} *} - -text {* - \glossary{Term}{FIXME} - - The language of terms is that of simply-typed @{text "\"}-calculus - with de-Bruijn indices for bound variables (cf.\ \cite{debruijn72} - or \cite{paulson-ml2}), with the types being determined determined - by the corresponding binders. In contrast, free variables and - constants are have an explicit name and type in each occurrence. - - \medskip A \emph{bound variable} is a natural number @{text "b"}, - which accounts for the number of intermediate binders between the - variable occurrence in the body and its binding position. For - example, the de-Bruijn term @{text - "\\<^bsub>nat\<^esub>. \\<^bsub>nat\<^esub>. 1 + 0"} would - correspond to @{text - "\x\<^bsub>nat\<^esub>. \y\<^bsub>nat\<^esub>. x + y"} in a named - representation. Note that a bound variable may be represented by - different de-Bruijn indices at different occurrences, depending on - the nesting of abstractions. - - A \emph{loose variable} is a bound variable that is outside the - scope of local binders. The types (and names) for loose variables - can be managed as a separate context, that is maintained as a stack - of hypothetical binders. The core logic operates on closed terms, - without any loose variables. - - A \emph{fixed variable} is a pair of a basic name and a type, e.g.\ - @{text "(x, \)"} which is usually printed @{text "x\<^isub>\"}. A - \emph{schematic variable} is a pair of an indexname and a type, - e.g.\ @{text "((x, 0), \)"} which is usually printed as @{text - "?x\<^isub>\"}. - - \medskip A \emph{constant} is a pair of a basic name and a type, - e.g.\ @{text "(c, \)"} which is usually printed as @{text - "c\<^isub>\"}. Constants are declared in the context as polymorphic - families @{text "c :: \"}, meaning that all substitution instances - @{text "c\<^isub>\"} for @{text "\ = \\"} are valid. - - The vector of \emph{type arguments} of constant @{text "c\<^isub>\"} - wrt.\ the declaration @{text "c :: \"} is defined as the codomain of - the matcher @{text "\ = {?\\<^isub>1 \ \\<^isub>1, \, - ?\\<^isub>n \ \\<^isub>n}"} presented in canonical order @{text - "(\\<^isub>1, \, \\<^isub>n)"}. Within a given theory context, - there is a one-to-one correspondence between any constant @{text - "c\<^isub>\"} and the application @{text "c(\\<^isub>1, \, - \\<^isub>n)"} of its type arguments. For example, with @{text "plus - :: \ \ \ \ \"}, the instance @{text "plus\<^bsub>nat \ nat \ - nat\<^esub>"} corresponds to @{text "plus(nat)"}. - - Constant declarations @{text "c :: \"} may contain sort constraints - for type variables in @{text "\"}. These are observed by - type-inference as expected, but \emph{ignored} by the core logic. - This means the primitive logic is able to reason with instances of - polymorphic constants that the user-level type-checker would reject - due to violation of type class restrictions. - - \medskip An \emph{atomic} term is either a variable or constant. A - \emph{term} is defined inductively over atomic terms, with - abstraction and application as follows: @{text "t = b | x\<^isub>\ | - ?x\<^isub>\ | c\<^isub>\ | \\<^isub>\. t | t\<^isub>1 t\<^isub>2"}. - Parsing and printing takes care of converting between an external - representation with named bound variables. Subsequently, we shall - use the latter notation instead of internal de-Bruijn - representation. - - The inductive relation @{text "t :: \"} assigns a (unique) type to a - term according to the structure of atomic terms, abstractions, and - applicatins: - \[ - \infer{@{text "a\<^isub>\ :: \"}}{} - \qquad - \infer{@{text "(\x\<^sub>\. t) :: \ \ \"}}{@{text "t :: \"}} - \qquad - \infer{@{text "t u :: \"}}{@{text "t :: \ \ \"} & @{text "u :: \"}} - \] - A \emph{well-typed term} is a term that can be typed according to these rules. - - Typing information can be omitted: type-inference is able to - reconstruct the most general type of a raw term, while assigning - most general types to all of its variables and constants. - Type-inference depends on a context of type constraints for fixed - variables, and declarations for polymorphic constants. - - The identity of atomic terms consists both of the name and the type - component. This means that different variables @{text - "x\<^bsub>\\<^isub>1\<^esub>"} and @{text - "x\<^bsub>\\<^isub>2\<^esub>"} may become the same after type - instantiation. Some outer layers of the system make it hard to - produce variables of the same name, but different types. In - contrast, mixed instances of polymorphic constants occur frequently. - - \medskip The \emph{hidden polymorphism} of a term @{text "t :: \"} - is the set of type variables occurring in @{text "t"}, but not in - @{text "\"}. This means that the term implicitly depends on type - arguments that are not accounted in the result type, i.e.\ there are - different type instances @{text "t\ :: \"} and @{text - "t\' :: \"} with the same type. This slightly - pathological situation notoriously demands additional care. - - \medskip A \emph{term abbreviation} is a syntactic definition @{text - "c\<^isub>\ \ t"} of a closed term @{text "t"} of type @{text "\"}, - without any hidden polymorphism. A term abbreviation looks like a - constant in the syntax, but is expanded before entering the logical - core. Abbreviations are usually reverted when printing terms, using - @{text "t \ c\<^isub>\"} as rules for higher-order rewriting. - - \medskip Canonical operations on @{text "\"}-terms include @{text - "\\\"}-conversion: @{text "\"}-conversion refers to capture-free - renaming of bound variables; @{text "\"}-conversion contracts an - abstraction applied to an argument term, substituting the argument - in the body: @{text "(\x. b)a"} becomes @{text "b[a/x]"}; @{text - "\"}-conversion contracts vacuous application-abstraction: @{text - "\x. f x"} becomes @{text "f"}, provided that the bound variable - does not occur in @{text "f"}. - - Terms are normally treated modulo @{text "\"}-conversion, which is - implicit in the de-Bruijn representation. Names for bound variables - in abstractions are maintained separately as (meaningless) comments, - mostly for parsing and printing. Full @{text "\\\"}-conversion is - commonplace in various standard operations (\secref{sec:obj-rules}) - that are based on higher-order unification and matching. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML_type term} \\ - @{index_ML "op aconv": "term * term -> bool"} \\ - @{index_ML map_types: "(typ -> typ) -> term -> term"} \\ - @{index_ML fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\ - @{index_ML map_aterms: "(term -> term) -> term -> term"} \\ - @{index_ML fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\ - \end{mldecls} - \begin{mldecls} - @{index_ML fastype_of: "term -> typ"} \\ - @{index_ML lambda: "term -> term -> term"} \\ - @{index_ML betapply: "term * term -> term"} \\ - @{index_ML Sign.declare_const: "Properties.T -> (binding * typ) * mixfix -> - theory -> term * theory"} \\ - @{index_ML Sign.add_abbrev: "string -> Properties.T -> binding * term -> - theory -> (term * term) * theory"} \\ - @{index_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\ - @{index_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML_type term} represents de-Bruijn terms, with comments in - abstractions, and explicitly named free variables and constants; - this is a datatype with constructors @{ML Bound}, @{ML Free}, @{ML - Var}, @{ML Const}, @{ML Abs}, @{ML "op $"}. - - \item @{text "t"}~@{ML aconv}~@{text "u"} checks @{text - "\"}-equivalence of two terms. This is the basic equality relation - on type @{ML_type term}; raw datatype equality should only be used - for operations related to parsing or printing! - - \item @{ML map_types}~@{text "f t"} applies the mapping @{text - "f"} to all types occurring in @{text "t"}. - - \item @{ML fold_types}~@{text "f t"} iterates the operation @{text - "f"} over all occurrences of types in @{text "t"}; the term - structure is traversed from left to right. - - \item @{ML map_aterms}~@{text "f t"} applies the mapping @{text "f"} - to all atomic terms (@{ML Bound}, @{ML Free}, @{ML Var}, @{ML - Const}) occurring in @{text "t"}. - - \item @{ML fold_aterms}~@{text "f t"} iterates the operation @{text - "f"} over all occurrences of atomic terms (@{ML Bound}, @{ML Free}, - @{ML Var}, @{ML Const}) in @{text "t"}; the term structure is - traversed from left to right. - - \item @{ML fastype_of}~@{text "t"} determines the type of a - well-typed term. This operation is relatively slow, despite the - omission of any sanity checks. - - \item @{ML lambda}~@{text "a b"} produces an abstraction @{text - "\a. b"}, where occurrences of the atomic term @{text "a"} in the - body @{text "b"} are replaced by bound variables. - - \item @{ML betapply}~@{text "(t, u)"} produces an application @{text - "t u"}, with topmost @{text "\"}-conversion if @{text "t"} is an - abstraction. - - \item @{ML Sign.declare_const}~@{text "properties ((c, \), mx)"} - declares a new constant @{text "c :: \"} with optional mixfix - syntax. - - \item @{ML Sign.add_abbrev}~@{text "print_mode properties (c, t)"} - introduces a new term abbreviation @{text "c \ t"}. - - \item @{ML Sign.const_typargs}~@{text "thy (c, \)"} and @{ML - Sign.const_instance}~@{text "thy (c, [\\<^isub>1, \, \\<^isub>n])"} - convert between two representations of polymorphic constants: full - type instance vs.\ compact type arguments form. - - \end{description} -*} - - -section {* Theorems \label{sec:thms} *} - -text {* - \glossary{Proposition}{FIXME A \seeglossary{term} of - \seeglossary{type} @{text "prop"}. Internally, there is nothing - special about propositions apart from their type, but the concrete - syntax enforces a clear distinction. Propositions are structured - via implication @{text "A \ B"} or universal quantification @{text - "\x. B x"} --- anything else is considered atomic. The canonical - form for propositions is that of a \seeglossary{Hereditary Harrop - Formula}. FIXME} - - \glossary{Theorem}{A proven proposition within a certain theory and - proof context, formally @{text "\ \\<^sub>\ \"}; both contexts are - rarely spelled out explicitly. Theorems are usually normalized - according to the \seeglossary{HHF} format. FIXME} - - \glossary{Fact}{Sometimes used interchangeably for - \seeglossary{theorem}. Strictly speaking, a list of theorems, - essentially an extra-logical conjunction. Facts emerge either as - local assumptions, or as results of local goal statements --- both - may be simultaneous, hence the list representation. FIXME} - - \glossary{Schematic variable}{FIXME} - - \glossary{Fixed variable}{A variable that is bound within a certain - proof context; an arbitrary-but-fixed entity within a portion of - proof text. FIXME} - - \glossary{Free variable}{Synonymous for \seeglossary{fixed - variable}. FIXME} - - \glossary{Bound variable}{FIXME} - - \glossary{Variable}{See \seeglossary{schematic variable}, - \seeglossary{fixed variable}, \seeglossary{bound variable}, or - \seeglossary{type variable}. The distinguishing feature of - different variables is their binding scope. FIXME} - - A \emph{proposition} is a well-typed term of type @{text "prop"}, a - \emph{theorem} is a proven proposition (depending on a context of - hypotheses and the background theory). Primitive inferences include - plain natural deduction rules for the primary connectives @{text - "\"} and @{text "\"} of the framework. There is also a builtin - notion of equality/equivalence @{text "\"}. -*} - -subsection {* Primitive connectives and rules \label{sec:prim-rules} *} - -text {* - The theory @{text "Pure"} contains constant declarations for the - primitive connectives @{text "\"}, @{text "\"}, and @{text "\"} of - the logical framework, see \figref{fig:pure-connectives}. The - derivability judgment @{text "A\<^isub>1, \, A\<^isub>n \ B"} is - defined inductively by the primitive inferences given in - \figref{fig:prim-rules}, with the global restriction that the - hypotheses must \emph{not} contain any schematic variables. The - builtin equality is conceptually axiomatized as shown in - \figref{fig:pure-equality}, although the implementation works - directly with derived inferences. - - \begin{figure}[htb] - \begin{center} - \begin{tabular}{ll} - @{text "all :: (\ \ prop) \ prop"} & universal quantification (binder @{text "\"}) \\ - @{text "\ :: prop \ prop \ prop"} & implication (right associative infix) \\ - @{text "\ :: \ \ \ \ prop"} & equality relation (infix) \\ - \end{tabular} - \caption{Primitive connectives of Pure}\label{fig:pure-connectives} - \end{center} - \end{figure} - - \begin{figure}[htb] - \begin{center} - \[ - \infer[@{text "(axiom)"}]{@{text "\ A"}}{@{text "A \ \"}} - \qquad - \infer[@{text "(assume)"}]{@{text "A \ A"}}{} - \] - \[ - \infer[@{text "(\_intro)"}]{@{text "\ \ \x. b[x]"}}{@{text "\ \ b[x]"} & @{text "x \ \"}} - \qquad - \infer[@{text "(\_elim)"}]{@{text "\ \ b[a]"}}{@{text "\ \ \x. b[x]"}} - \] - \[ - \infer[@{text "(\_intro)"}]{@{text "\ - A \ A \ B"}}{@{text "\ \ B"}} - \qquad - \infer[@{text "(\_elim)"}]{@{text "\\<^sub>1 \ \\<^sub>2 \ B"}}{@{text "\\<^sub>1 \ A \ B"} & @{text "\\<^sub>2 \ A"}} - \] - \caption{Primitive inferences of Pure}\label{fig:prim-rules} - \end{center} - \end{figure} - - \begin{figure}[htb] - \begin{center} - \begin{tabular}{ll} - @{text "\ (\x. b[x]) a \ b[a]"} & @{text "\"}-conversion \\ - @{text "\ x \ x"} & reflexivity \\ - @{text "\ x \ y \ P x \ P y"} & substitution \\ - @{text "\ (\x. f x \ g x) \ f \ g"} & extensionality \\ - @{text "\ (A \ B) \ (B \ A) \ A \ B"} & logical equivalence \\ - \end{tabular} - \caption{Conceptual axiomatization of Pure equality}\label{fig:pure-equality} - \end{center} - \end{figure} - - The introduction and elimination rules for @{text "\"} and @{text - "\"} are analogous to formation of dependently typed @{text - "\"}-terms representing the underlying proof objects. Proof terms - are irrelevant in the Pure logic, though; they cannot occur within - propositions. The system provides a runtime option to record - explicit proof terms for primitive inferences. Thus all three - levels of @{text "\"}-calculus become explicit: @{text "\"} for - terms, and @{text "\/\"} for proofs (cf.\ - \cite{Berghofer-Nipkow:2000:TPHOL}). - - Observe that locally fixed parameters (as in @{text "\_intro"}) need - not be recorded in the hypotheses, because the simple syntactic - types of Pure are always inhabitable. ``Assumptions'' @{text "x :: - \"} for type-membership are only present as long as some @{text - "x\<^isub>\"} occurs in the statement body.\footnote{This is the key - difference to ``@{text "\HOL"}'' in the PTS framework - \cite{Barendregt-Geuvers:2001}, where hypotheses @{text "x : A"} are - treated uniformly for propositions and types.} - - \medskip The axiomatization of a theory is implicitly closed by - forming all instances of type and term variables: @{text "\ - A\"} holds for any substitution instance of an axiom - @{text "\ A"}. By pushing substitutions through derivations - inductively, we also get admissible @{text "generalize"} and @{text - "instance"} rules as shown in \figref{fig:subst-rules}. - - \begin{figure}[htb] - \begin{center} - \[ - \infer{@{text "\ \ B[?\]"}}{@{text "\ \ B[\]"} & @{text "\ \ \"}} - \quad - \infer[\quad@{text "(generalize)"}]{@{text "\ \ B[?x]"}}{@{text "\ \ B[x]"} & @{text "x \ \"}} - \] - \[ - \infer{@{text "\ \ B[\]"}}{@{text "\ \ B[?\]"}} - \quad - \infer[\quad@{text "(instantiate)"}]{@{text "\ \ B[t]"}}{@{text "\ \ B[?x]"}} - \] - \caption{Admissible substitution rules}\label{fig:subst-rules} - \end{center} - \end{figure} - - Note that @{text "instantiate"} does not require an explicit - side-condition, because @{text "\"} may never contain schematic - variables. - - In principle, variables could be substituted in hypotheses as well, - but this would disrupt the monotonicity of reasoning: deriving - @{text "\\ \ B\"} from @{text "\ \ B"} is - correct, but @{text "\\ \ \"} does not necessarily hold: - the result belongs to a different proof context. - - \medskip An \emph{oracle} is a function that produces axioms on the - fly. Logically, this is an instance of the @{text "axiom"} rule - (\figref{fig:prim-rules}), but there is an operational difference. - The system always records oracle invocations within derivations of - theorems. Tracing plain axioms (and named theorems) is optional. - - Axiomatizations should be limited to the bare minimum, typically as - part of the initial logical basis of an object-logic formalization. - Later on, theories are usually developed in a strictly definitional - fashion, by stating only certain equalities over new constants. - - A \emph{simple definition} consists of a constant declaration @{text - "c :: \"} together with an axiom @{text "\ c \ t"}, where @{text "t - :: \"} is a closed term without any hidden polymorphism. The RHS - may depend on further defined constants, but not @{text "c"} itself. - Definitions of functions may be presented as @{text "c \<^vec>x \ - t"} instead of the puristic @{text "c \ \\<^vec>x. t"}. - - An \emph{overloaded definition} consists of a collection of axioms - for the same constant, with zero or one equations @{text - "c((\<^vec>\)\) \ t"} for each type constructor @{text "\"} (for - distinct variables @{text "\<^vec>\"}). The RHS may mention - previously defined constants as above, or arbitrary constants @{text - "d(\\<^isub>i)"} for some @{text "\\<^isub>i"} projected from @{text - "\<^vec>\"}. Thus overloaded definitions essentially work by - primitive recursion over the syntactic structure of a single type - argument. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML_type ctyp} \\ - @{index_ML_type cterm} \\ - @{index_ML Thm.ctyp_of: "theory -> typ -> ctyp"} \\ - @{index_ML Thm.cterm_of: "theory -> term -> cterm"} \\ - \end{mldecls} - \begin{mldecls} - @{index_ML_type thm} \\ - @{index_ML proofs: "int ref"} \\ - @{index_ML Thm.assume: "cterm -> thm"} \\ - @{index_ML Thm.forall_intr: "cterm -> thm -> thm"} \\ - @{index_ML Thm.forall_elim: "cterm -> thm -> thm"} \\ - @{index_ML Thm.implies_intr: "cterm -> thm -> thm"} \\ - @{index_ML Thm.implies_elim: "thm -> thm -> thm"} \\ - @{index_ML Thm.generalize: "string list * string list -> int -> thm -> thm"} \\ - @{index_ML Thm.instantiate: "(ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm"} \\ - @{index_ML Thm.axiom: "theory -> string -> thm"} \\ - @{index_ML Thm.add_oracle: "bstring * ('a -> cterm) -> theory - -> (string * ('a -> thm)) * theory"} \\ - \end{mldecls} - \begin{mldecls} - @{index_ML Theory.add_axioms_i: "(binding * term) list -> theory -> theory"} \\ - @{index_ML Theory.add_deps: "string -> string * typ -> (string * typ) list -> theory -> theory"} \\ - @{index_ML Theory.add_defs_i: "bool -> bool -> (binding * term) list -> theory -> theory"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML_type ctyp} and @{ML_type cterm} represent certified types - and terms, respectively. These are abstract datatypes that - guarantee that its values have passed the full well-formedness (and - well-typedness) checks, relative to the declarations of type - constructors, constants etc. in the theory. - - \item @{ML ctyp_of}~@{text "thy \"} and @{ML cterm_of}~@{text "thy - t"} explicitly checks types and terms, respectively. This also - involves some basic normalizations, such expansion of type and term - abbreviations from the theory context. - - Re-certification is relatively slow and should be avoided in tight - reasoning loops. There are separate operations to decompose - certified entities (including actual theorems). - - \item @{ML_type thm} represents proven propositions. This is an - abstract datatype that guarantees that its values have been - constructed by basic principles of the @{ML_struct Thm} module. - Every @{ML thm} value contains a sliding back-reference to the - enclosing theory, cf.\ \secref{sec:context-theory}. - - \item @{ML proofs} determines the detail of proof recording within - @{ML_type thm} values: @{ML 0} records only oracles, @{ML 1} records - oracles, axioms and named theorems, @{ML 2} records full proof - terms. - - \item @{ML Thm.assume}, @{ML Thm.forall_intr}, @{ML - Thm.forall_elim}, @{ML Thm.implies_intr}, and @{ML Thm.implies_elim} - correspond to the primitive inferences of \figref{fig:prim-rules}. - - \item @{ML Thm.generalize}~@{text "(\<^vec>\, \<^vec>x)"} - corresponds to the @{text "generalize"} rules of - \figref{fig:subst-rules}. Here collections of type and term - variables are generalized simultaneously, specified by the given - basic names. - - \item @{ML Thm.instantiate}~@{text "(\<^vec>\\<^isub>s, - \<^vec>x\<^isub>\)"} corresponds to the @{text "instantiate"} rules - of \figref{fig:subst-rules}. Type variables are substituted before - term variables. Note that the types in @{text "\<^vec>x\<^isub>\"} - refer to the instantiated versions. - - \item @{ML Thm.axiom}~@{text "thy name"} retrieves a named - axiom, cf.\ @{text "axiom"} in \figref{fig:prim-rules}. - - \item @{ML Thm.add_oracle}~@{text "(name, oracle)"} produces a named - oracle rule, essentially generating arbitrary axioms on the fly, - cf.\ @{text "axiom"} in \figref{fig:prim-rules}. - - \item @{ML Theory.add_axioms_i}~@{text "[(name, A), \]"} declares - arbitrary propositions as axioms. - - \item @{ML Theory.add_deps}~@{text "name c\<^isub>\ - \<^vec>d\<^isub>\"} declares dependencies of a named specification - for constant @{text "c\<^isub>\"}, relative to existing - specifications for constants @{text "\<^vec>d\<^isub>\"}. - - \item @{ML Theory.add_defs_i}~@{text "unchecked overloaded [(name, c - \<^vec>x \ t), \]"} states a definitional axiom for an existing - constant @{text "c"}. Dependencies are recorded (cf.\ @{ML - Theory.add_deps}), unless the @{text "unchecked"} option is set. - - \end{description} -*} - - -subsection {* Auxiliary definitions *} - -text {* - Theory @{text "Pure"} provides a few auxiliary definitions, see - \figref{fig:pure-aux}. These special constants are normally not - exposed to the user, but appear in internal encodings. - - \begin{figure}[htb] - \begin{center} - \begin{tabular}{ll} - @{text "conjunction :: prop \ prop \ prop"} & (infix @{text "&"}) \\ - @{text "\ A & B \ (\C. (A \ B \ C) \ C)"} \\[1ex] - @{text "prop :: prop \ prop"} & (prefix @{text "#"}, suppressed) \\ - @{text "#A \ A"} \\[1ex] - @{text "term :: \ \ prop"} & (prefix @{text "TERM"}) \\ - @{text "term x \ (\A. A \ A)"} \\[1ex] - @{text "TYPE :: \ itself"} & (prefix @{text "TYPE"}) \\ - @{text "(unspecified)"} \\ - \end{tabular} - \caption{Definitions of auxiliary connectives}\label{fig:pure-aux} - \end{center} - \end{figure} - - Derived conjunction rules include introduction @{text "A \ B \ A & - B"}, and destructions @{text "A & B \ A"} and @{text "A & B \ B"}. - Conjunction allows to treat simultaneous assumptions and conclusions - uniformly. For example, multiple claims are intermediately - represented as explicit conjunction, but this is refined into - separate sub-goals before the user continues the proof; the final - result is projected into a list of theorems (cf.\ - \secref{sec:tactical-goals}). - - The @{text "prop"} marker (@{text "#"}) makes arbitrarily complex - propositions appear as atomic, without changing the meaning: @{text - "\ \ A"} and @{text "\ \ #A"} are interchangeable. See - \secref{sec:tactical-goals} for specific operations. - - The @{text "term"} marker turns any well-typed term into a derivable - proposition: @{text "\ TERM t"} holds unconditionally. Although - this is logically vacuous, it allows to treat terms and proofs - uniformly, similar to a type-theoretic framework. - - The @{text "TYPE"} constructor is the canonical representative of - the unspecified type @{text "\ itself"}; it essentially injects the - language of types into that of terms. There is specific notation - @{text "TYPE(\)"} for @{text "TYPE\<^bsub>\ - itself\<^esub>"}. - Although being devoid of any particular meaning, the @{text - "TYPE(\)"} accounts for the type @{text "\"} within the term - language. In particular, @{text "TYPE(\)"} may be used as formal - argument in primitive definitions, in order to circumvent hidden - polymorphism (cf.\ \secref{sec:terms}). For example, @{text "c - TYPE(\) \ A[\]"} defines @{text "c :: \ itself \ prop"} in terms of - a proposition @{text "A"} that depends on an additional type - argument, which is essentially a predicate on types. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML Conjunction.intr: "thm -> thm -> thm"} \\ - @{index_ML Conjunction.elim: "thm -> thm * thm"} \\ - @{index_ML Drule.mk_term: "cterm -> thm"} \\ - @{index_ML Drule.dest_term: "thm -> cterm"} \\ - @{index_ML Logic.mk_type: "typ -> term"} \\ - @{index_ML Logic.dest_type: "term -> typ"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML Conjunction.intr} derives @{text "A & B"} from @{text - "A"} and @{text "B"}. - - \item @{ML Conjunction.elim} derives @{text "A"} and @{text "B"} - from @{text "A & B"}. - - \item @{ML Drule.mk_term} derives @{text "TERM t"}. - - \item @{ML Drule.dest_term} recovers term @{text "t"} from @{text - "TERM t"}. - - \item @{ML Logic.mk_type}~@{text "\"} produces the term @{text - "TYPE(\)"}. - - \item @{ML Logic.dest_type}~@{text "TYPE(\)"} recovers the type - @{text "\"}. - - \end{description} -*} - - -section {* Object-level rules \label{sec:obj-rules} *} - -text %FIXME {* - -FIXME - - A \emph{rule} is any Pure theorem in HHF normal form; there is a - separate calculus for rule composition, which is modeled after - Gentzen's Natural Deduction \cite{Gentzen:1935}, but allows - rules to be nested arbitrarily, similar to \cite{extensions91}. - - Normally, all theorems accessible to the user are proper rules. - Low-level inferences are occasional required internally, but the - result should be always presented in canonical form. The higher - interfaces of Isabelle/Isar will always produce proper rules. It is - important to maintain this invariant in add-on applications! - - There are two main principles of rule composition: @{text - "resolution"} (i.e.\ backchaining of rules) and @{text - "by-assumption"} (i.e.\ closing a branch); both principles are - combined in the variants of @{text "elim-resolution"} and @{text - "dest-resolution"}. Raw @{text "composition"} is occasionally - useful as well, also it is strictly speaking outside of the proper - rule calculus. - - Rules are treated modulo general higher-order unification, which is - unification modulo the equational theory of @{text "\\\"}-conversion - on @{text "\"}-terms. Moreover, propositions are understood modulo - the (derived) equivalence @{text "(A \ (\x. B x)) \ (\x. A \ B x)"}. - - This means that any operations within the rule calculus may be - subject to spontaneous @{text "\\\"}-HHF conversions. It is common - practice not to contract or expand unnecessarily. Some mechanisms - prefer an one form, others the opposite, so there is a potential - danger to produce some oscillation! - - Only few operations really work \emph{modulo} HHF conversion, but - expect a normal form: quantifiers @{text "\"} before implications - @{text "\"} at each level of nesting. - -\glossary{Hereditary Harrop Formula}{The set of propositions in HHF -format is defined inductively as @{text "H = (\x\<^sup>*. H\<^sup>* \ -A)"}, for variables @{text "x"} and atomic propositions @{text "A"}. -Any proposition may be put into HHF form by normalizing with the rule -@{text "(A \ (\x. B x)) \ (\x. A \ B x)"}. In Isabelle, the outermost -quantifier prefix is represented via \seeglossary{schematic -variables}, such that the top-level structure is merely that of a -\seeglossary{Horn Clause}}. - -\glossary{HHF}{See \seeglossary{Hereditary Harrop Formula}.} - - - \[ - \infer[@{text "(assumption)"}]{@{text "C\"}} - {@{text "(\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x) \ C"} & @{text "A\ = H\<^sub>i\"}~~\text{(for some~@{text i})}} - \] - - - \[ - \infer[@{text "(compose)"}]{@{text "\<^vec>A\ \ C\"}} - {@{text "\<^vec>A \ B"} & @{text "B' \ C"} & @{text "B\ = B'\"}} - \] - - - \[ - \infer[@{text "(\_lift)"}]{@{text "(\\<^vec>x. \<^vec>A (?\<^vec>a \<^vec>x)) \ (\\<^vec>x. B (?\<^vec>a \<^vec>x))"}}{@{text "\<^vec>A ?\<^vec>a \ B ?\<^vec>a"}} - \] - \[ - \infer[@{text "(\_lift)"}]{@{text "(\<^vec>H \ \<^vec>A) \ (\<^vec>H \ B)"}}{@{text "\<^vec>A \ B"}} - \] - - The @{text resolve} scheme is now acquired from @{text "\_lift"}, - @{text "\_lift"}, and @{text compose}. - - \[ - \infer[@{text "(resolution)"}] - {@{text "(\\<^vec>x. \<^vec>H \<^vec>x \ \<^vec>A (?\<^vec>a \<^vec>x))\ \ C\"}} - {\begin{tabular}{l} - @{text "\<^vec>A ?\<^vec>a \ B ?\<^vec>a"} \\ - @{text "(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C"} \\ - @{text "(\\<^vec>x. B (?\<^vec>a \<^vec>x))\ = B'\"} \\ - \end{tabular}} - \] - - - FIXME @{text "elim_resolution"}, @{text "dest_resolution"} -*} - - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/prelim.thy --- a/doc-src/IsarImplementation/Thy/prelim.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,779 +0,0 @@ - -(* $Id$ *) - -theory prelim imports base begin - -chapter {* Preliminaries *} - -section {* Contexts \label{sec:context} *} - -text {* - A logical context represents the background that is required for - formulating statements and composing proofs. It acts as a medium to - produce formal content, depending on earlier material (declarations, - results etc.). - - For example, derivations within the Isabelle/Pure logic can be - described as a judgment @{text "\ \\<^sub>\ \"}, which means that a - proposition @{text "\"} is derivable from hypotheses @{text "\"} - within the theory @{text "\"}. There are logical reasons for - keeping @{text "\"} and @{text "\"} separate: theories can be - liberal about supporting type constructors and schematic - polymorphism of constants and axioms, while the inner calculus of - @{text "\ \ \"} is strictly limited to Simple Type Theory (with - fixed type variables in the assumptions). - - \medskip Contexts and derivations are linked by the following key - principles: - - \begin{itemize} - - \item Transfer: monotonicity of derivations admits results to be - transferred into a \emph{larger} context, i.e.\ @{text "\ \\<^sub>\ - \"} implies @{text "\' \\<^sub>\\<^sub>' \"} for contexts @{text "\' - \ \"} and @{text "\' \ \"}. - - \item Export: discharge of hypotheses admits results to be exported - into a \emph{smaller} context, i.e.\ @{text "\' \\<^sub>\ \"} - implies @{text "\ \\<^sub>\ \ \ \"} where @{text "\' \ \"} and - @{text "\ = \' - \"}. Note that @{text "\"} remains unchanged here, - only the @{text "\"} part is affected. - - \end{itemize} - - \medskip By modeling the main characteristics of the primitive - @{text "\"} and @{text "\"} above, and abstracting over any - particular logical content, we arrive at the fundamental notions of - \emph{theory context} and \emph{proof context} in Isabelle/Isar. - These implement a certain policy to manage arbitrary \emph{context - data}. There is a strongly-typed mechanism to declare new kinds of - data at compile time. - - The internal bootstrap process of Isabelle/Pure eventually reaches a - stage where certain data slots provide the logical content of @{text - "\"} and @{text "\"} sketched above, but this does not stop there! - Various additional data slots support all kinds of mechanisms that - are not necessarily part of the core logic. - - For example, there would be data for canonical introduction and - elimination rules for arbitrary operators (depending on the - object-logic and application), which enables users to perform - standard proof steps implicitly (cf.\ the @{text "rule"} method - \cite{isabelle-isar-ref}). - - \medskip Thus Isabelle/Isar is able to bring forth more and more - concepts successively. In particular, an object-logic like - Isabelle/HOL continues the Isabelle/Pure setup by adding specific - components for automated reasoning (classical reasoner, tableau - prover, structured induction etc.) and derived specification - mechanisms (inductive predicates, recursive functions etc.). All of - this is ultimately based on the generic data management by theory - and proof contexts introduced here. -*} - - -subsection {* Theory context \label{sec:context-theory} *} - -text {* - \glossary{Theory}{FIXME} - - A \emph{theory} is a data container with explicit named and unique - identifier. Theories are related by a (nominal) sub-theory - relation, which corresponds to the dependency graph of the original - construction; each theory is derived from a certain sub-graph of - ancestor theories. - - The @{text "merge"} operation produces the least upper bound of two - theories, which actually degenerates into absorption of one theory - into the other (due to the nominal sub-theory relation). - - The @{text "begin"} operation starts a new theory by importing - several parent theories and entering a special @{text "draft"} mode, - which is sustained until the final @{text "end"} operation. A draft - theory acts like a linear type, where updates invalidate earlier - versions. An invalidated draft is called ``stale''. - - The @{text "checkpoint"} operation produces an intermediate stepping - stone that will survive the next update: both the original and the - changed theory remain valid and are related by the sub-theory - relation. Checkpointing essentially recovers purely functional - theory values, at the expense of some extra internal bookkeeping. - - The @{text "copy"} operation produces an auxiliary version that has - the same data content, but is unrelated to the original: updates of - the copy do not affect the original, neither does the sub-theory - relation hold. - - \medskip The example in \figref{fig:ex-theory} below shows a theory - graph derived from @{text "Pure"}, with theory @{text "Length"} - importing @{text "Nat"} and @{text "List"}. The body of @{text - "Length"} consists of a sequence of updates, working mostly on - drafts. Intermediate checkpoints may occur as well, due to the - history mechanism provided by the Isar top-level, cf.\ - \secref{sec:isar-toplevel}. - - \begin{figure}[htb] - \begin{center} - \begin{tabular}{rcccl} - & & @{text "Pure"} \\ - & & @{text "\"} \\ - & & @{text "FOL"} \\ - & $\swarrow$ & & $\searrow$ & \\ - @{text "Nat"} & & & & @{text "List"} \\ - & $\searrow$ & & $\swarrow$ \\ - & & @{text "Length"} \\ - & & \multicolumn{3}{l}{~~@{keyword "imports"}} \\ - & & \multicolumn{3}{l}{~~@{keyword "begin"}} \\ - & & $\vdots$~~ \\ - & & @{text "\"}~~ \\ - & & $\vdots$~~ \\ - & & @{text "\"}~~ \\ - & & $\vdots$~~ \\ - & & \multicolumn{3}{l}{~~@{command "end"}} \\ - \end{tabular} - \caption{A theory definition depending on ancestors}\label{fig:ex-theory} - \end{center} - \end{figure} - - \medskip There is a separate notion of \emph{theory reference} for - maintaining a live link to an evolving theory context: updates on - drafts are propagated automatically. Dynamic updating stops after - an explicit @{text "end"} only. - - Derived entities may store a theory reference in order to indicate - the context they belong to. This implicitly assumes monotonic - reasoning, because the referenced context may become larger without - further notice. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML_type theory} \\ - @{index_ML Theory.subthy: "theory * theory -> bool"} \\ - @{index_ML Theory.merge: "theory * theory -> theory"} \\ - @{index_ML Theory.checkpoint: "theory -> theory"} \\ - @{index_ML Theory.copy: "theory -> theory"} \\ - \end{mldecls} - \begin{mldecls} - @{index_ML_type theory_ref} \\ - @{index_ML Theory.deref: "theory_ref -> theory"} \\ - @{index_ML Theory.check_thy: "theory -> theory_ref"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML_type theory} represents theory contexts. This is - essentially a linear type! Most operations destroy the original - version, which then becomes ``stale''. - - \item @{ML "Theory.subthy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} - compares theories according to the inherent graph structure of the - construction. This sub-theory relation is a nominal approximation - of inclusion (@{text "\"}) of the corresponding content. - - \item @{ML "Theory.merge"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} - absorbs one theory into the other. This fails for unrelated - theories! - - \item @{ML "Theory.checkpoint"}~@{text "thy"} produces a safe - stepping stone in the linear development of @{text "thy"}. The next - update will result in two related, valid theories. - - \item @{ML "Theory.copy"}~@{text "thy"} produces a variant of @{text - "thy"} that holds a copy of the same data. The result is not - related to the original; the original is unchanched. - - \item @{ML_type theory_ref} represents a sliding reference to an - always valid theory; updates on the original are propagated - automatically. - - \item @{ML "Theory.deref"}~@{text "thy_ref"} turns a @{ML_type - "theory_ref"} into an @{ML_type "theory"} value. As the referenced - theory evolves monotonically over time, later invocations of @{ML - "Theory.deref"} may refer to a larger context. - - \item @{ML "Theory.check_thy"}~@{text "thy"} produces a @{ML_type - "theory_ref"} from a valid @{ML_type "theory"} value. - - \end{description} -*} - - -subsection {* Proof context \label{sec:context-proof} *} - -text {* - \glossary{Proof context}{The static context of a structured proof, - acts like a local ``theory'' of the current portion of Isar proof - text, generalizes the idea of local hypotheses @{text "\"} in - judgments @{text "\ \ \"} of natural deduction calculi. There is a - generic notion of introducing and discharging hypotheses. - Arbritrary auxiliary context data may be adjoined.} - - A proof context is a container for pure data with a back-reference - to the theory it belongs to. The @{text "init"} operation creates a - proof context from a given theory. Modifications to draft theories - are propagated to the proof context as usual, but there is also an - explicit @{text "transfer"} operation to force resynchronization - with more substantial updates to the underlying theory. The actual - context data does not require any special bookkeeping, thanks to the - lack of destructive features. - - Entities derived in a proof context need to record inherent logical - requirements explicitly, since there is no separate context - identification as for theories. For example, hypotheses used in - primitive derivations (cf.\ \secref{sec:thms}) are recorded - separately within the sequent @{text "\ \ \"}, just to make double - sure. Results could still leak into an alien proof context do to - programming errors, but Isabelle/Isar includes some extra validity - checks in critical positions, notably at the end of a sub-proof. - - Proof contexts may be manipulated arbitrarily, although the common - discipline is to follow block structure as a mental model: a given - context is extended consecutively, and results are exported back - into the original context. Note that the Isar proof states model - block-structured reasoning explicitly, using a stack of proof - contexts internally, cf.\ \secref{sec:isar-proof-state}. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML_type Proof.context} \\ - @{index_ML ProofContext.init: "theory -> Proof.context"} \\ - @{index_ML ProofContext.theory_of: "Proof.context -> theory"} \\ - @{index_ML ProofContext.transfer: "theory -> Proof.context -> Proof.context"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML_type Proof.context} represents proof contexts. Elements - of this type are essentially pure values, with a sliding reference - to the background theory. - - \item @{ML ProofContext.init}~@{text "thy"} produces a proof context - derived from @{text "thy"}, initializing all data. - - \item @{ML ProofContext.theory_of}~@{text "ctxt"} selects the - background theory from @{text "ctxt"}, dereferencing its internal - @{ML_type theory_ref}. - - \item @{ML ProofContext.transfer}~@{text "thy ctxt"} promotes the - background theory of @{text "ctxt"} to the super theory @{text - "thy"}. - - \end{description} -*} - - -subsection {* Generic contexts \label{sec:generic-context} *} - -text {* - A generic context is the disjoint sum of either a theory or proof - context. Occasionally, this enables uniform treatment of generic - context data, typically extra-logical information. Operations on - generic contexts include the usual injections, partial selections, - and combinators for lifting operations on either component of the - disjoint sum. - - Moreover, there are total operations @{text "theory_of"} and @{text - "proof_of"} to convert a generic context into either kind: a theory - can always be selected from the sum, while a proof context might - have to be constructed by an ad-hoc @{text "init"} operation. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML_type Context.generic} \\ - @{index_ML Context.theory_of: "Context.generic -> theory"} \\ - @{index_ML Context.proof_of: "Context.generic -> Proof.context"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML_type Context.generic} is the direct sum of @{ML_type - "theory"} and @{ML_type "Proof.context"}, with the datatype - constructors @{ML "Context.Theory"} and @{ML "Context.Proof"}. - - \item @{ML Context.theory_of}~@{text "context"} always produces a - theory from the generic @{text "context"}, using @{ML - "ProofContext.theory_of"} as required. - - \item @{ML Context.proof_of}~@{text "context"} always produces a - proof context from the generic @{text "context"}, using @{ML - "ProofContext.init"} as required (note that this re-initializes the - context data with each invocation). - - \end{description} -*} - - -subsection {* Context data \label{sec:context-data} *} - -text {* - The main purpose of theory and proof contexts is to manage arbitrary - data. New data types can be declared incrementally at compile time. - There are separate declaration mechanisms for any of the three kinds - of contexts: theory, proof, generic. - - \paragraph{Theory data} may refer to destructive entities, which are - maintained in direct correspondence to the linear evolution of - theory values, including explicit copies.\footnote{Most existing - instances of destructive theory data are merely historical relics - (e.g.\ the destructive theorem storage, and destructive hints for - the Simplifier and Classical rules).} A theory data declaration - needs to implement the following SML signature: - - \medskip - \begin{tabular}{ll} - @{text "\ T"} & representing type \\ - @{text "\ empty: T"} & empty default value \\ - @{text "\ copy: T \ T"} & refresh impure data \\ - @{text "\ extend: T \ T"} & re-initialize on import \\ - @{text "\ merge: T \ T \ T"} & join on import \\ - \end{tabular} - \medskip - - \noindent The @{text "empty"} value acts as initial default for - \emph{any} theory that does not declare actual data content; @{text - "copy"} maintains persistent integrity for impure data, it is just - the identity for pure values; @{text "extend"} is acts like a - unitary version of @{text "merge"}, both operations should also - include the functionality of @{text "copy"} for impure data. - - \paragraph{Proof context data} is purely functional. A declaration - needs to implement the following SML signature: - - \medskip - \begin{tabular}{ll} - @{text "\ T"} & representing type \\ - @{text "\ init: theory \ T"} & produce initial value \\ - \end{tabular} - \medskip - - \noindent The @{text "init"} operation is supposed to produce a pure - value from the given background theory. - - \paragraph{Generic data} provides a hybrid interface for both theory - and proof data. The declaration is essentially the same as for - (pure) theory data, without @{text "copy"}. The @{text "init"} - operation for proof contexts merely selects the current data value - from the background theory. - - \bigskip A data declaration of type @{text "T"} results in the - following interface: - - \medskip - \begin{tabular}{ll} - @{text "init: theory \ theory"} \\ - @{text "get: context \ T"} \\ - @{text "put: T \ context \ context"} \\ - @{text "map: (T \ T) \ context \ context"} \\ - \end{tabular} - \medskip - - \noindent Here @{text "init"} is only applicable to impure theory - data to install a fresh copy persistently (destructive update on - uninitialized has no permanent effect). The other operations provide - access for the particular kind of context (theory, proof, or generic - context). Note that this is a safe interface: there is no other way - to access the corresponding data slot of a context. By keeping - these operations private, a component may maintain abstract values - authentically, without other components interfering. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML_functor TheoryDataFun} \\ - @{index_ML_functor ProofDataFun} \\ - @{index_ML_functor GenericDataFun} \\ - \end{mldecls} - - \begin{description} - - \item @{ML_functor TheoryDataFun}@{text "(spec)"} declares data for - type @{ML_type theory} according to the specification provided as - argument structure. The resulting structure provides data init and - access operations as described above. - - \item @{ML_functor ProofDataFun}@{text "(spec)"} is analogous to - @{ML_functor TheoryDataFun} for type @{ML_type Proof.context}. - - \item @{ML_functor GenericDataFun}@{text "(spec)"} is analogous to - @{ML_functor TheoryDataFun} for type @{ML_type Context.generic}. - - \end{description} -*} - - -section {* Names \label{sec:names} *} - -text {* - In principle, a name is just a string, but there are various - convention for encoding additional structure. For example, ``@{text - "Foo.bar.baz"}'' is considered as a qualified name consisting of - three basic name components. The individual constituents of a name - may have further substructure, e.g.\ the string - ``\verb,\,\verb,,'' encodes as a single symbol. -*} - - -subsection {* Strings of symbols *} - -text {* - \glossary{Symbol}{The smallest unit of text in Isabelle, subsumes - plain ASCII characters as well as an infinite collection of named - symbols (for greek, math etc.).} - - A \emph{symbol} constitutes the smallest textual unit in Isabelle - --- raw characters are normally not encountered at all. Isabelle - strings consist of a sequence of symbols, represented as a packed - string or a list of strings. Each symbol is in itself a small - string, which has either one of the following forms: - - \begin{enumerate} - - \item a single ASCII character ``@{text "c"}'', for example - ``\verb,a,'', - - \item a regular symbol ``\verb,\,\verb,<,@{text "ident"}\verb,>,'', - for example ``\verb,\,\verb,,'', - - \item a control symbol ``\verb,\,\verb,<^,@{text "ident"}\verb,>,'', - for example ``\verb,\,\verb,<^bold>,'', - - \item a raw symbol ``\verb,\,\verb,<^raw:,@{text text}\verb,>,'' - where @{text text} constists of printable characters excluding - ``\verb,.,'' and ``\verb,>,'', for example - ``\verb,\,\verb,<^raw:$\sum_{i = 1}^n$>,'', - - \item a numbered raw control symbol ``\verb,\,\verb,<^raw,@{text - n}\verb,>, where @{text n} consists of digits, for example - ``\verb,\,\verb,<^raw42>,''. - - \end{enumerate} - - \noindent The @{text "ident"} syntax for symbol names is @{text - "letter (letter | digit)\<^sup>*"}, where @{text "letter = - A..Za..z"} and @{text "digit = 0..9"}. There are infinitely many - regular symbols and control symbols, but a fixed collection of - standard symbols is treated specifically. For example, - ``\verb,\,\verb,,'' is classified as a letter, which means it - may occur within regular Isabelle identifiers. - - Since the character set underlying Isabelle symbols is 7-bit ASCII - and 8-bit characters are passed through transparently, Isabelle may - also process Unicode/UCS data in UTF-8 encoding. Unicode provides - its own collection of mathematical symbols, but there is no built-in - link to the standard collection of Isabelle. - - \medskip Output of Isabelle symbols depends on the print mode - (\secref{FIXME}). For example, the standard {\LaTeX} setup of the - Isabelle document preparation system would present - ``\verb,\,\verb,,'' as @{text "\"}, and - ``\verb,\,\verb,<^bold>,\verb,\,\verb,,'' as @{text - "\<^bold>\"}. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML_type "Symbol.symbol"} \\ - @{index_ML Symbol.explode: "string -> Symbol.symbol list"} \\ - @{index_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\ - @{index_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\ - @{index_ML Symbol.is_quasi: "Symbol.symbol -> bool"} \\ - @{index_ML Symbol.is_blank: "Symbol.symbol -> bool"} \\ - \end{mldecls} - \begin{mldecls} - @{index_ML_type "Symbol.sym"} \\ - @{index_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML_type "Symbol.symbol"} represents individual Isabelle - symbols; this is an alias for @{ML_type "string"}. - - \item @{ML "Symbol.explode"}~@{text "str"} produces a symbol list - from the packed form. This function supercedes @{ML - "String.explode"} for virtually all purposes of manipulating text in - Isabelle! - - \item @{ML "Symbol.is_letter"}, @{ML "Symbol.is_digit"}, @{ML - "Symbol.is_quasi"}, @{ML "Symbol.is_blank"} classify standard - symbols according to fixed syntactic conventions of Isabelle, cf.\ - \cite{isabelle-isar-ref}. - - \item @{ML_type "Symbol.sym"} is a concrete datatype that represents - the different kinds of symbols explicitly, with constructors @{ML - "Symbol.Char"}, @{ML "Symbol.Sym"}, @{ML "Symbol.Ctrl"}, @{ML - "Symbol.Raw"}. - - \item @{ML "Symbol.decode"} converts the string representation of a - symbol into the datatype version. - - \end{description} -*} - - -subsection {* Basic names \label{sec:basic-names} *} - -text {* - A \emph{basic name} essentially consists of a single Isabelle - identifier. There are conventions to mark separate classes of basic - names, by attaching a suffix of underscores (@{text "_"}): one - underscore means \emph{internal name}, two underscores means - \emph{Skolem name}, three underscores means \emph{internal Skolem - name}. - - For example, the basic name @{text "foo"} has the internal version - @{text "foo_"}, with Skolem versions @{text "foo__"} and @{text - "foo___"}, respectively. - - These special versions provide copies of the basic name space, apart - from anything that normally appears in the user text. For example, - system generated variables in Isar proof contexts are usually marked - as internal, which prevents mysterious name references like @{text - "xaa"} to appear in the text. - - \medskip Manipulating binding scopes often requires on-the-fly - renamings. A \emph{name context} contains a collection of already - used names. The @{text "declare"} operation adds names to the - context. - - The @{text "invents"} operation derives a number of fresh names from - a given starting point. For example, the first three names derived - from @{text "a"} are @{text "a"}, @{text "b"}, @{text "c"}. - - The @{text "variants"} operation produces fresh names by - incrementing tentative names as base-26 numbers (with digits @{text - "a..z"}) until all clashes are resolved. For example, name @{text - "foo"} results in variants @{text "fooa"}, @{text "foob"}, @{text - "fooc"}, \dots, @{text "fooaa"}, @{text "fooab"} etc.; each renaming - step picks the next unused variant from this sequence. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML Name.internal: "string -> string"} \\ - @{index_ML Name.skolem: "string -> string"} \\ - \end{mldecls} - \begin{mldecls} - @{index_ML_type Name.context} \\ - @{index_ML Name.context: Name.context} \\ - @{index_ML Name.declare: "string -> Name.context -> Name.context"} \\ - @{index_ML Name.invents: "Name.context -> string -> int -> string list"} \\ - @{index_ML Name.variants: "string list -> Name.context -> string list * Name.context"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML Name.internal}~@{text "name"} produces an internal name - by adding one underscore. - - \item @{ML Name.skolem}~@{text "name"} produces a Skolem name by - adding two underscores. - - \item @{ML_type Name.context} represents the context of already used - names; the initial value is @{ML "Name.context"}. - - \item @{ML Name.declare}~@{text "name"} enters a used name into the - context. - - \item @{ML Name.invents}~@{text "context name n"} produces @{text - "n"} fresh names derived from @{text "name"}. - - \item @{ML Name.variants}~@{text "names context"} produces fresh - varians of @{text "names"}; the result is entered into the context. - - \end{description} -*} - - -subsection {* Indexed names *} - -text {* - An \emph{indexed name} (or @{text "indexname"}) is a pair of a basic - name and a natural number. This representation allows efficient - renaming by incrementing the second component only. The canonical - way to rename two collections of indexnames apart from each other is - this: determine the maximum index @{text "maxidx"} of the first - collection, then increment all indexes of the second collection by - @{text "maxidx + 1"}; the maximum index of an empty collection is - @{text "-1"}. - - Occasionally, basic names and indexed names are injected into the - same pair type: the (improper) indexname @{text "(x, -1)"} is used - to encode basic names. - - \medskip Isabelle syntax observes the following rules for - representing an indexname @{text "(x, i)"} as a packed string: - - \begin{itemize} - - \item @{text "?x"} if @{text "x"} does not end with a digit and @{text "i = 0"}, - - \item @{text "?xi"} if @{text "x"} does not end with a digit, - - \item @{text "?x.i"} otherwise. - - \end{itemize} - - Indexnames may acquire large index numbers over time. Results are - normalized towards @{text "0"} at certain checkpoints, notably at - the end of a proof. This works by producing variants of the - corresponding basic name components. For example, the collection - @{text "?x1, ?x7, ?x42"} becomes @{text "?x, ?xa, ?xb"}. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML_type indexname} \\ - \end{mldecls} - - \begin{description} - - \item @{ML_type indexname} represents indexed names. This is an - abbreviation for @{ML_type "string * int"}. The second component is - usually non-negative, except for situations where @{text "(x, -1)"} - is used to embed basic names into this type. - - \end{description} -*} - - -subsection {* Qualified names and name spaces *} - -text {* - A \emph{qualified name} consists of a non-empty sequence of basic - name components. The packed representation uses a dot as separator, - as in ``@{text "A.b.c"}''. The last component is called \emph{base} - name, the remaining prefix \emph{qualifier} (which may be empty). - The idea of qualified names is to encode nested structures by - recording the access paths as qualifiers. For example, an item - named ``@{text "A.b.c"}'' may be understood as a local entity @{text - "c"}, within a local structure @{text "b"}, within a global - structure @{text "A"}. Typically, name space hierarchies consist of - 1--2 levels of qualification, but this need not be always so. - - The empty name is commonly used as an indication of unnamed - entities, whenever this makes any sense. The basic operations on - qualified names are smart enough to pass through such improper names - unchanged. - - \medskip A @{text "naming"} policy tells how to turn a name - specification into a fully qualified internal name (by the @{text - "full"} operation), and how fully qualified names may be accessed - externally. For example, the default naming policy is to prefix an - implicit path: @{text "full x"} produces @{text "path.x"}, and the - standard accesses for @{text "path.x"} include both @{text "x"} and - @{text "path.x"}. Normally, the naming is implicit in the theory or - proof context; there are separate versions of the corresponding. - - \medskip A @{text "name space"} manages a collection of fully - internalized names, together with a mapping between external names - and internal names (in both directions). The corresponding @{text - "intern"} and @{text "extern"} operations are mostly used for - parsing and printing only! The @{text "declare"} operation augments - a name space according to the accesses determined by the naming - policy. - - \medskip As a general principle, there is a separate name space for - each kind of formal entity, e.g.\ logical constant, type - constructor, type class, theorem. It is usually clear from the - occurrence in concrete syntax (or from the scope) which kind of - entity a name refers to. For example, the very same name @{text - "c"} may be used uniformly for a constant, type constructor, and - type class. - - There are common schemes to name theorems systematically, according - to the name of the main logical entity involved, e.g.\ @{text - "c.intro"} for a canonical theorem related to constant @{text "c"}. - This technique of mapping names from one space into another requires - some care in order to avoid conflicts. In particular, theorem names - derived from a type constructor or type class are better suffixed in - addition to the usual qualification, e.g.\ @{text "c_type.intro"} - and @{text "c_class.intro"} for theorems related to type @{text "c"} - and class @{text "c"}, respectively. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML NameSpace.base: "string -> string"} \\ - @{index_ML NameSpace.qualifier: "string -> string"} \\ - @{index_ML NameSpace.append: "string -> string -> string"} \\ - @{index_ML NameSpace.implode: "string list -> string"} \\ - @{index_ML NameSpace.explode: "string -> string list"} \\ - \end{mldecls} - \begin{mldecls} - @{index_ML_type NameSpace.naming} \\ - @{index_ML NameSpace.default_naming: NameSpace.naming} \\ - @{index_ML NameSpace.add_path: "string -> NameSpace.naming -> NameSpace.naming"} \\ - @{index_ML NameSpace.full_name: "NameSpace.naming -> binding -> string"} \\ - \end{mldecls} - \begin{mldecls} - @{index_ML_type NameSpace.T} \\ - @{index_ML NameSpace.empty: NameSpace.T} \\ - @{index_ML NameSpace.merge: "NameSpace.T * NameSpace.T -> NameSpace.T"} \\ - @{index_ML NameSpace.declare: "NameSpace.naming -> binding -> NameSpace.T -> string * NameSpace.T"} \\ - @{index_ML NameSpace.intern: "NameSpace.T -> string -> string"} \\ - @{index_ML NameSpace.extern: "NameSpace.T -> string -> string"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML NameSpace.base}~@{text "name"} returns the base name of a - qualified name. - - \item @{ML NameSpace.qualifier}~@{text "name"} returns the qualifier - of a qualified name. - - \item @{ML NameSpace.append}~@{text "name\<^isub>1 name\<^isub>2"} - appends two qualified names. - - \item @{ML NameSpace.implode}~@{text "name"} and @{ML - NameSpace.explode}~@{text "names"} convert between the packed string - representation and the explicit list form of qualified names. - - \item @{ML_type NameSpace.naming} represents the abstract concept of - a naming policy. - - \item @{ML NameSpace.default_naming} is the default naming policy. - In a theory context, this is usually augmented by a path prefix - consisting of the theory name. - - \item @{ML NameSpace.add_path}~@{text "path naming"} augments the - naming policy by extending its path component. - - \item @{ML NameSpace.full_name}@{text "naming binding"} turns a name - binding (usually a basic name) into the fully qualified - internal name, according to the given naming policy. - - \item @{ML_type NameSpace.T} represents name spaces. - - \item @{ML NameSpace.empty} and @{ML NameSpace.merge}~@{text - "(space\<^isub>1, space\<^isub>2)"} are the canonical operations for - maintaining name spaces according to theory data management - (\secref{sec:context-data}). - - \item @{ML NameSpace.declare}~@{text "naming bindings space"} enters a - name binding as fully qualified internal name into the name space, - with external accesses determined by the naming policy. - - \item @{ML NameSpace.intern}~@{text "space name"} internalizes a - (partially qualified) external name. - - This operation is mostly for parsing! Note that fully qualified - names stemming from declarations are produced via @{ML - "NameSpace.full_name"} and @{ML "NameSpace.declare"} - (or their derivatives for @{ML_type theory} and - @{ML_type Proof.context}). - - \item @{ML NameSpace.extern}~@{text "space name"} externalizes a - (fully qualified) internal name. - - This operation is mostly for printing! Note unqualified names are - produced via @{ML NameSpace.base}. - - \end{description} -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/proof.thy --- a/doc-src/IsarImplementation/Thy/proof.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,332 +0,0 @@ - -(* $Id$ *) - -theory "proof" imports base begin - -chapter {* Structured proofs *} - -section {* Variables \label{sec:variables} *} - -text {* - Any variable that is not explicitly bound by @{text "\"}-abstraction - is considered as ``free''. Logically, free variables act like - outermost universal quantification at the sequent level: @{text - "A\<^isub>1(x), \, A\<^isub>n(x) \ B(x)"} means that the result - holds \emph{for all} values of @{text "x"}. Free variables for - terms (not types) can be fully internalized into the logic: @{text - "\ B(x)"} and @{text "\ \x. B(x)"} are interchangeable, provided - that @{text "x"} does not occur elsewhere in the context. - Inspecting @{text "\ \x. B(x)"} more closely, we see that inside the - quantifier, @{text "x"} is essentially ``arbitrary, but fixed'', - while from outside it appears as a place-holder for instantiation - (thanks to @{text "\"} elimination). - - The Pure logic represents the idea of variables being either inside - or outside the current scope by providing separate syntactic - categories for \emph{fixed variables} (e.g.\ @{text "x"}) vs.\ - \emph{schematic variables} (e.g.\ @{text "?x"}). Incidently, a - universal result @{text "\ \x. B(x)"} has the HHF normal form @{text - "\ B(?x)"}, which represents its generality nicely without requiring - an explicit quantifier. The same principle works for type - variables: @{text "\ B(?\)"} represents the idea of ``@{text "\ - \\. B(\)"}'' without demanding a truly polymorphic framework. - - \medskip Additional care is required to treat type variables in a - way that facilitates type-inference. In principle, term variables - depend on type variables, which means that type variables would have - to be declared first. For example, a raw type-theoretic framework - would demand the context to be constructed in stages as follows: - @{text "\ = \: type, x: \, a: A(x\<^isub>\)"}. - - We allow a slightly less formalistic mode of operation: term - variables @{text "x"} are fixed without specifying a type yet - (essentially \emph{all} potential occurrences of some instance - @{text "x\<^isub>\"} are fixed); the first occurrence of @{text "x"} - within a specific term assigns its most general type, which is then - maintained consistently in the context. The above example becomes - @{text "\ = x: term, \: type, A(x\<^isub>\)"}, where type @{text - "\"} is fixed \emph{after} term @{text "x"}, and the constraint - @{text "x :: \"} is an implicit consequence of the occurrence of - @{text "x\<^isub>\"} in the subsequent proposition. - - This twist of dependencies is also accommodated by the reverse - operation of exporting results from a context: a type variable - @{text "\"} is considered fixed as long as it occurs in some fixed - term variable of the context. For example, exporting @{text "x: - term, \: type \ x\<^isub>\ = x\<^isub>\"} produces in the first step - @{text "x: term \ x\<^isub>\ = x\<^isub>\"} for fixed @{text "\"}, - and only in the second step @{text "\ ?x\<^isub>?\<^isub>\ = - ?x\<^isub>?\<^isub>\"} for schematic @{text "?x"} and @{text "?\"}. - - \medskip The Isabelle/Isar proof context manages the gory details of - term vs.\ type variables, with high-level principles for moving the - frontier between fixed and schematic variables. - - The @{text "add_fixes"} operation explictly declares fixed - variables; the @{text "declare_term"} operation absorbs a term into - a context by fixing new type variables and adding syntactic - constraints. - - The @{text "export"} operation is able to perform the main work of - generalizing term and type variables as sketched above, assuming - that fixing variables and terms have been declared properly. - - There @{text "import"} operation makes a generalized fact a genuine - part of the context, by inventing fixed variables for the schematic - ones. The effect can be reversed by using @{text "export"} later, - potentially with an extended context; the result is equivalent to - the original modulo renaming of schematic variables. - - The @{text "focus"} operation provides a variant of @{text "import"} - for nested propositions (with explicit quantification): @{text - "\x\<^isub>1 \ x\<^isub>n. B(x\<^isub>1, \, x\<^isub>n)"} is - decomposed by inventing fixed variables @{text "x\<^isub>1, \, - x\<^isub>n"} for the body. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML Variable.add_fixes: " - string list -> Proof.context -> string list * Proof.context"} \\ - @{index_ML Variable.variant_fixes: " - string list -> Proof.context -> string list * Proof.context"} \\ - @{index_ML Variable.declare_term: "term -> Proof.context -> Proof.context"} \\ - @{index_ML Variable.declare_constraints: "term -> Proof.context -> Proof.context"} \\ - @{index_ML Variable.export: "Proof.context -> Proof.context -> thm list -> thm list"} \\ - @{index_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\ - @{index_ML Variable.import_thms: "bool -> thm list -> Proof.context -> - ((ctyp list * cterm list) * thm list) * Proof.context"} \\ - @{index_ML Variable.focus: "cterm -> Proof.context -> (cterm list * cterm) * Proof.context"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML Variable.add_fixes}~@{text "xs ctxt"} fixes term - variables @{text "xs"}, returning the resulting internal names. By - default, the internal representation coincides with the external - one, which also means that the given variables must not be fixed - already. There is a different policy within a local proof body: the - given names are just hints for newly invented Skolem variables. - - \item @{ML Variable.variant_fixes} is similar to @{ML - Variable.add_fixes}, but always produces fresh variants of the given - names. - - \item @{ML Variable.declare_term}~@{text "t ctxt"} declares term - @{text "t"} to belong to the context. This automatically fixes new - type variables, but not term variables. Syntactic constraints for - type and term variables are declared uniformly, though. - - \item @{ML Variable.declare_constraints}~@{text "t ctxt"} declares - syntactic constraints from term @{text "t"}, without making it part - of the context yet. - - \item @{ML Variable.export}~@{text "inner outer thms"} generalizes - fixed type and term variables in @{text "thms"} according to the - difference of the @{text "inner"} and @{text "outer"} context, - following the principles sketched above. - - \item @{ML Variable.polymorphic}~@{text "ctxt ts"} generalizes type - variables in @{text "ts"} as far as possible, even those occurring - in fixed term variables. The default policy of type-inference is to - fix newly introduced type variables, which is essentially reversed - with @{ML Variable.polymorphic}: here the given terms are detached - from the context as far as possible. - - \item @{ML Variable.import_thms}~@{text "open thms ctxt"} invents fixed - type and term variables for the schematic ones occurring in @{text - "thms"}. The @{text "open"} flag indicates whether the fixed names - should be accessible to the user, otherwise newly introduced names - are marked as ``internal'' (\secref{sec:names}). - - \item @{ML Variable.focus}~@{text B} decomposes the outermost @{text - "\"} prefix of proposition @{text "B"}. - - \end{description} -*} - - -section {* Assumptions \label{sec:assumptions} *} - -text {* - An \emph{assumption} is a proposition that it is postulated in the - current context. Local conclusions may use assumptions as - additional facts, but this imposes implicit hypotheses that weaken - the overall statement. - - Assumptions are restricted to fixed non-schematic statements, i.e.\ - all generality needs to be expressed by explicit quantifiers. - Nevertheless, the result will be in HHF normal form with outermost - quantifiers stripped. For example, by assuming @{text "\x :: \. P - x"} we get @{text "\x :: \. P x \ P ?x"} for schematic @{text "?x"} - of fixed type @{text "\"}. Local derivations accumulate more and - more explicit references to hypotheses: @{text "A\<^isub>1, \, - A\<^isub>n \ B"} where @{text "A\<^isub>1, \, A\<^isub>n"} needs to - be covered by the assumptions of the current context. - - \medskip The @{text "add_assms"} operation augments the context by - local assumptions, which are parameterized by an arbitrary @{text - "export"} rule (see below). - - The @{text "export"} operation moves facts from a (larger) inner - context into a (smaller) outer context, by discharging the - difference of the assumptions as specified by the associated export - rules. Note that the discharged portion is determined by the - difference contexts, not the facts being exported! There is a - separate flag to indicate a goal context, where the result is meant - to refine an enclosing sub-goal of a structured proof state (cf.\ - \secref{sec:isar-proof-state}). - - \medskip The most basic export rule discharges assumptions directly - by means of the @{text "\"} introduction rule: - \[ - \infer[(@{text "\_intro"})]{@{text "\ \\ A \ A \ B"}}{@{text "\ \ B"}} - \] - - The variant for goal refinements marks the newly introduced - premises, which causes the canonical Isar goal refinement scheme to - enforce unification with local premises within the goal: - \[ - \infer[(@{text "#\_intro"})]{@{text "\ \\ A \ #A \ B"}}{@{text "\ \ B"}} - \] - - \medskip Alternative versions of assumptions may perform arbitrary - transformations on export, as long as the corresponding portion of - hypotheses is removed from the given facts. For example, a local - definition works by fixing @{text "x"} and assuming @{text "x \ t"}, - with the following export rule to reverse the effect: - \[ - \infer[(@{text "\-expand"})]{@{text "\ \\ x \ t \ B t"}}{@{text "\ \ B x"}} - \] - This works, because the assumption @{text "x \ t"} was introduced in - a context with @{text "x"} being fresh, so @{text "x"} does not - occur in @{text "\"} here. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML_type Assumption.export} \\ - @{index_ML Assumption.assume: "cterm -> thm"} \\ - @{index_ML Assumption.add_assms: - "Assumption.export -> - cterm list -> Proof.context -> thm list * Proof.context"} \\ - @{index_ML Assumption.add_assumes: " - cterm list -> Proof.context -> thm list * Proof.context"} \\ - @{index_ML Assumption.export: "bool -> Proof.context -> Proof.context -> thm -> thm"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML_type Assumption.export} represents arbitrary export - rules, which is any function of type @{ML_type "bool -> cterm list -> thm -> thm"}, - where the @{ML_type "bool"} indicates goal mode, and the @{ML_type - "cterm list"} the collection of assumptions to be discharged - simultaneously. - - \item @{ML Assumption.assume}~@{text "A"} turns proposition @{text - "A"} into a raw assumption @{text "A \ A'"}, where the conclusion - @{text "A'"} is in HHF normal form. - - \item @{ML Assumption.add_assms}~@{text "r As"} augments the context - by assumptions @{text "As"} with export rule @{text "r"}. The - resulting facts are hypothetical theorems as produced by the raw - @{ML Assumption.assume}. - - \item @{ML Assumption.add_assumes}~@{text "As"} is a special case of - @{ML Assumption.add_assms} where the export rule performs @{text - "\_intro"} or @{text "#\_intro"}, depending on goal mode. - - \item @{ML Assumption.export}~@{text "is_goal inner outer thm"} - exports result @{text "thm"} from the the @{text "inner"} context - back into the @{text "outer"} one; @{text "is_goal = true"} means - this is a goal context. The result is in HHF normal form. Note - that @{ML "ProofContext.export"} combines @{ML "Variable.export"} - and @{ML "Assumption.export"} in the canonical way. - - \end{description} -*} - - -section {* Results \label{sec:results} *} - -text {* - Local results are established by monotonic reasoning from facts - within a context. This allows common combinations of theorems, - e.g.\ via @{text "\/\"} elimination, resolution rules, or equational - reasoning, see \secref{sec:thms}. Unaccounted context manipulations - should be avoided, notably raw @{text "\/\"} introduction or ad-hoc - references to free variables or assumptions not present in the proof - context. - - \medskip The @{text "SUBPROOF"} combinator allows to structure a - tactical proof recursively by decomposing a selected sub-goal: - @{text "(\x. A(x) \ B(x)) \ \"} is turned into @{text "B(x) \ \"} - after fixing @{text "x"} and assuming @{text "A(x)"}. This means - the tactic needs to solve the conclusion, but may use the premise as - a local fact, for locally fixed variables. - - The @{text "prove"} operation provides an interface for structured - backwards reasoning under program control, with some explicit sanity - checks of the result. The goal context can be augmented by - additional fixed variables (cf.\ \secref{sec:variables}) and - assumptions (cf.\ \secref{sec:assumptions}), which will be available - as local facts during the proof and discharged into implications in - the result. Type and term variables are generalized as usual, - according to the context. - - The @{text "obtain"} operation produces results by eliminating - existing facts by means of a given tactic. This acts like a dual - conclusion: the proof demonstrates that the context may be augmented - by certain fixed variables and assumptions. See also - \cite{isabelle-isar-ref} for the user-level @{text "\"} and - @{text "\"} elements. Final results, which may not refer to - the parameters in the conclusion, need to exported explicitly into - the original context. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML SUBPROOF: - "({context: Proof.context, schematics: ctyp list * cterm list, - params: cterm list, asms: cterm list, concl: cterm, - prems: thm list} -> tactic) -> Proof.context -> int -> tactic"} \\ - \end{mldecls} - \begin{mldecls} - @{index_ML Goal.prove: "Proof.context -> string list -> term list -> term -> - ({prems: thm list, context: Proof.context} -> tactic) -> thm"} \\ - @{index_ML Goal.prove_multi: "Proof.context -> string list -> term list -> term list -> - ({prems: thm list, context: Proof.context} -> tactic) -> thm list"} \\ - \end{mldecls} - \begin{mldecls} - @{index_ML Obtain.result: "(Proof.context -> tactic) -> - thm list -> Proof.context -> (cterm list * thm list) * Proof.context"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML SUBPROOF}~@{text "tac"} decomposes the structure of a - particular sub-goal, producing an extended context and a reduced - goal, which needs to be solved by the given tactic. All schematic - parameters of the goal are imported into the context as fixed ones, - which may not be instantiated in the sub-proof. - - \item @{ML Goal.prove}~@{text "ctxt xs As C tac"} states goal @{text - "C"} in the context augmented by fixed variables @{text "xs"} and - assumptions @{text "As"}, and applies tactic @{text "tac"} to solve - it. The latter may depend on the local assumptions being presented - as facts. The result is in HHF normal form. - - \item @{ML Goal.prove_multi} is simular to @{ML Goal.prove}, but - states several conclusions simultaneously. The goal is encoded by - means of Pure conjunction; @{ML Goal.conjunction_tac} will turn this - into a collection of individual subgoals. - - \item @{ML Obtain.result}~@{text "tac thms ctxt"} eliminates the - given facts using a tactic, which results in additional fixed - variables and assumptions in the context. Final results need to be - exported explicitly. - - \end{description} -*} - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/tactic.thy --- a/doc-src/IsarImplementation/Thy/tactic.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,420 +0,0 @@ - -(* $Id$ *) - -theory tactic imports base begin - -chapter {* Tactical reasoning *} - -text {* - Tactical reasoning works by refining the initial claim in a - backwards fashion, until a solved form is reached. A @{text "goal"} - consists of several subgoals that need to be solved in order to - achieve the main statement; zero subgoals means that the proof may - be finished. A @{text "tactic"} is a refinement operation that maps - a goal to a lazy sequence of potential successors. A @{text - "tactical"} is a combinator for composing tactics. -*} - - -section {* Goals \label{sec:tactical-goals} *} - -text {* - Isabelle/Pure represents a goal\glossary{Tactical goal}{A theorem of - \seeglossary{Horn Clause} form stating that a number of subgoals - imply the main conclusion, which is marked as a protected - proposition.} as a theorem stating that the subgoals imply the main - goal: @{text "A\<^sub>1 \ \ \ A\<^sub>n \ C"}. The outermost goal - structure is that of a Horn Clause\glossary{Horn Clause}{An iterated - implication @{text "A\<^sub>1 \ \ \ A\<^sub>n \ C"}, without any - outermost quantifiers. Strictly speaking, propositions @{text - "A\<^sub>i"} need to be atomic in Horn Clauses, but Isabelle admits - arbitrary substructure here (nested @{text "\"} and @{text "\"} - connectives).}: i.e.\ an iterated implication without any - quantifiers\footnote{Recall that outermost @{text "\x. \[x]"} is - always represented via schematic variables in the body: @{text - "\[?x]"}. These variables may get instantiated during the course of - reasoning.}. For @{text "n = 0"} a goal is called ``solved''. - - The structure of each subgoal @{text "A\<^sub>i"} is that of a general - Hereditary Harrop Formula @{text "\x\<^sub>1 \ \x\<^sub>k. H\<^sub>1 \ \ \ H\<^sub>m \ B"} in - normal form. Here @{text "x\<^sub>1, \, x\<^sub>k"} are goal parameters, i.e.\ - arbitrary-but-fixed entities of certain types, and @{text "H\<^sub>1, \, - H\<^sub>m"} are goal hypotheses, i.e.\ facts that may be assumed locally. - Together, this forms the goal context of the conclusion @{text B} to - be established. The goal hypotheses may be again arbitrary - Hereditary Harrop Formulas, although the level of nesting rarely - exceeds 1--2 in practice. - - The main conclusion @{text C} is internally marked as a protected - proposition\glossary{Protected proposition}{An arbitrarily - structured proposition @{text "C"} which is forced to appear as - atomic by wrapping it into a propositional identity operator; - notation @{text "#C"}. Protecting a proposition prevents basic - inferences from entering into that structure for the time being.}, - which is represented explicitly by the notation @{text "#C"}. This - ensures that the decomposition into subgoals and main conclusion is - well-defined for arbitrarily structured claims. - - \medskip Basic goal management is performed via the following - Isabelle/Pure rules: - - \[ - \infer[@{text "(init)"}]{@{text "C \ #C"}}{} \qquad - \infer[@{text "(finish)"}]{@{text "C"}}{@{text "#C"}} - \] - - \medskip The following low-level variants admit general reasoning - with protected propositions: - - \[ - \infer[@{text "(protect)"}]{@{text "#C"}}{@{text "C"}} \qquad - \infer[@{text "(conclude)"}]{@{text "A\<^sub>1 \ \ \ A\<^sub>n \ C"}}{@{text "A\<^sub>1 \ \ \ A\<^sub>n \ #C"}} - \] -*} - -text %mlref {* - \begin{mldecls} - @{index_ML Goal.init: "cterm -> thm"} \\ - @{index_ML Goal.finish: "thm -> thm"} \\ - @{index_ML Goal.protect: "thm -> thm"} \\ - @{index_ML Goal.conclude: "thm -> thm"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML "Goal.init"}~@{text C} initializes a tactical goal from - the well-formed proposition @{text C}. - - \item @{ML "Goal.finish"}~@{text "thm"} checks whether theorem - @{text "thm"} is a solved goal (no subgoals), and concludes the - result by removing the goal protection. - - \item @{ML "Goal.protect"}~@{text "thm"} protects the full statement - of theorem @{text "thm"}. - - \item @{ML "Goal.conclude"}~@{text "thm"} removes the goal - protection, even if there are pending subgoals. - - \end{description} -*} - - -section {* Tactics *} - -text {* A @{text "tactic"} is a function @{text "goal \ goal\<^sup>*\<^sup>*"} that - maps a given goal state (represented as a theorem, cf.\ - \secref{sec:tactical-goals}) to a lazy sequence of potential - successor states. The underlying sequence implementation is lazy - both in head and tail, and is purely functional in \emph{not} - supporting memoing.\footnote{The lack of memoing and the strict - nature of SML requires some care when working with low-level - sequence operations, to avoid duplicate or premature evaluation of - results.} - - An \emph{empty result sequence} means that the tactic has failed: in - a compound tactic expressions other tactics might be tried instead, - or the whole refinement step might fail outright, producing a - toplevel error message. When implementing tactics from scratch, one - should take care to observe the basic protocol of mapping regular - error conditions to an empty result; only serious faults should - emerge as exceptions. - - By enumerating \emph{multiple results}, a tactic can easily express - the potential outcome of an internal search process. There are also - combinators for building proof tools that involve search - systematically, see also \secref{sec:tacticals}. - - \medskip As explained in \secref{sec:tactical-goals}, a goal state - essentially consists of a list of subgoals that imply the main goal - (conclusion). Tactics may operate on all subgoals or on a - particularly specified subgoal, but must not change the main - conclusion (apart from instantiating schematic goal variables). - - Tactics with explicit \emph{subgoal addressing} are of the form - @{text "int \ tactic"} and may be applied to a particular subgoal - (counting from 1). If the subgoal number is out of range, the - tactic should fail with an empty result sequence, but must not raise - an exception! - - Operating on a particular subgoal means to replace it by an interval - of zero or more subgoals in the same place; other subgoals must not - be affected, apart from instantiating schematic variables ranging - over the whole goal state. - - A common pattern of composing tactics with subgoal addressing is to - try the first one, and then the second one only if the subgoal has - not been solved yet. Special care is required here to avoid bumping - into unrelated subgoals that happen to come after the original - subgoal. Assuming that there is only a single initial subgoal is a - very common error when implementing tactics! - - Tactics with internal subgoal addressing should expose the subgoal - index as @{text "int"} argument in full generality; a hardwired - subgoal 1 inappropriate. - - \medskip The main well-formedness conditions for proper tactics are - summarized as follows. - - \begin{itemize} - - \item General tactic failure is indicated by an empty result, only - serious faults may produce an exception. - - \item The main conclusion must not be changed, apart from - instantiating schematic variables. - - \item A tactic operates either uniformly on all subgoals, or - specifically on a selected subgoal (without bumping into unrelated - subgoals). - - \item Range errors in subgoal addressing produce an empty result. - - \end{itemize} - - Some of these conditions are checked by higher-level goal - infrastructure (\secref{sec:results}); others are not checked - explicitly, and violating them merely results in ill-behaved tactics - experienced by the user (e.g.\ tactics that insist in being - applicable only to singleton goals, or disallow composition with - basic tacticals). -*} - -text %mlref {* - \begin{mldecls} - @{index_ML_type tactic: "thm -> thm Seq.seq"} \\ - @{index_ML no_tac: tactic} \\ - @{index_ML all_tac: tactic} \\ - @{index_ML print_tac: "string -> tactic"} \\[1ex] - @{index_ML PRIMITIVE: "(thm -> thm) -> tactic"} \\[1ex] - @{index_ML SUBGOAL: "(term * int -> tactic) -> int -> tactic"} \\ - @{index_ML CSUBGOAL: "(cterm * int -> tactic) -> int -> tactic"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML_type tactic} represents tactics. The well-formedness - conditions described above need to be observed. See also @{"file" - "~~/src/Pure/General/seq.ML"} for the underlying implementation of - lazy sequences. - - \item @{ML_type "int -> tactic"} represents tactics with explicit - subgoal addressing, with well-formedness conditions as described - above. - - \item @{ML no_tac} is a tactic that always fails, returning the - empty sequence. - - \item @{ML all_tac} is a tactic that always succeeds, returning a - singleton sequence with unchanged goal state. - - \item @{ML print_tac}~@{text "message"} is like @{ML all_tac}, but - prints a message together with the goal state on the tracing - channel. - - \item @{ML PRIMITIVE}~@{text rule} turns a primitive inference rule - into a tactic with unique result. Exception @{ML THM} is considered - a regular tactic failure and produces an empty result; other - exceptions are passed through. - - \item @{ML SUBGOAL}~@{text "(fn (subgoal, i) => tactic)"} is the - most basic form to produce a tactic with subgoal addressing. The - given abstraction over the subgoal term and subgoal number allows to - peek at the relevant information of the full goal state. The - subgoal range is checked as required above. - - \item @{ML CSUBGOAL} is similar to @{ML SUBGOAL}, but passes the - subgoal as @{ML_type cterm} instead of raw @{ML_type term}. This - avoids expensive re-certification in situations where the subgoal is - used directly for primitive inferences. - - \end{description} -*} - - -subsection {* Resolution and assumption tactics \label{sec:resolve-assume-tac} *} - -text {* \emph{Resolution} is the most basic mechanism for refining a - subgoal using a theorem as object-level rule. - \emph{Elim-resolution} is particularly suited for elimination rules: - it resolves with a rule, proves its first premise by assumption, and - finally deletes that assumption from any new subgoals. - \emph{Destruct-resolution} is like elim-resolution, but the given - destruction rules are first turned into canonical elimination - format. \emph{Forward-resolution} is like destruct-resolution, but - without deleting the selected assumption. The @{text "r/e/d/f"} - naming convention is maintained for several different kinds of - resolution rules and tactics. - - Assumption tactics close a subgoal by unifying some of its premises - against its conclusion. - - \medskip All the tactics in this section operate on a subgoal - designated by a positive integer. Other subgoals might be affected - indirectly, due to instantiation of schematic variables. - - There are various sources of non-determinism, the tactic result - sequence enumerates all possibilities of the following choices (if - applicable): - - \begin{enumerate} - - \item selecting one of the rules given as argument to the tactic; - - \item selecting a subgoal premise to eliminate, unifying it against - the first premise of the rule; - - \item unifying the conclusion of the subgoal to the conclusion of - the rule. - - \end{enumerate} - - Recall that higher-order unification may produce multiple results - that are enumerated here. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML resolve_tac: "thm list -> int -> tactic"} \\ - @{index_ML eresolve_tac: "thm list -> int -> tactic"} \\ - @{index_ML dresolve_tac: "thm list -> int -> tactic"} \\ - @{index_ML forward_tac: "thm list -> int -> tactic"} \\[1ex] - @{index_ML assume_tac: "int -> tactic"} \\ - @{index_ML eq_assume_tac: "int -> tactic"} \\[1ex] - @{index_ML match_tac: "thm list -> int -> tactic"} \\ - @{index_ML ematch_tac: "thm list -> int -> tactic"} \\ - @{index_ML dmatch_tac: "thm list -> int -> tactic"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML resolve_tac}~@{text "thms i"} refines the goal state - using the given theorems, which should normally be introduction - rules. The tactic resolves a rule's conclusion with subgoal @{text - i}, replacing it by the corresponding versions of the rule's - premises. - - \item @{ML eresolve_tac}~@{text "thms i"} performs elim-resolution - with the given theorems, which should normally be elimination rules. - - \item @{ML dresolve_tac}~@{text "thms i"} performs - destruct-resolution with the given theorems, which should normally - be destruction rules. This replaces an assumption by the result of - applying one of the rules. - - \item @{ML forward_tac} is like @{ML dresolve_tac} except that the - selected assumption is not deleted. It applies a rule to an - assumption, adding the result as a new assumption. - - \item @{ML assume_tac}~@{text i} attempts to solve subgoal @{text i} - by assumption (modulo higher-order unification). - - \item @{ML eq_assume_tac} is similar to @{ML assume_tac}, but checks - only for immediate @{text "\"}-convertibility instead of using - unification. It succeeds (with a unique next state) if one of the - assumptions is equal to the subgoal's conclusion. Since it does not - instantiate variables, it cannot make other subgoals unprovable. - - \item @{ML match_tac}, @{ML ematch_tac}, and @{ML dmatch_tac} are - similar to @{ML resolve_tac}, @{ML eresolve_tac}, and @{ML - dresolve_tac}, respectively, but do not instantiate schematic - variables in the goal state. - - Flexible subgoals are not updated at will, but are left alone. - Strictly speaking, matching means to treat the unknowns in the goal - state as constants; these tactics merely discard unifiers that would - update the goal state. - - \end{description} -*} - - -subsection {* Explicit instantiation within a subgoal context *} - -text {* The main resolution tactics (\secref{sec:resolve-assume-tac}) - use higher-order unification, which works well in many practical - situations despite its daunting theoretical properties. - Nonetheless, there are important problem classes where unguided - higher-order unification is not so useful. This typically involves - rules like universal elimination, existential introduction, or - equational substitution. Here the unification problem involves - fully flexible @{text "?P ?x"} schemes, which are hard to manage - without further hints. - - By providing a (small) rigid term for @{text "?x"} explicitly, the - remaining unification problem is to assign a (large) term to @{text - "?P"}, according to the shape of the given subgoal. This is - sufficiently well-behaved in most practical situations. - - \medskip Isabelle provides separate versions of the standard @{text - "r/e/d/f"} resolution tactics that allow to provide explicit - instantiations of unknowns of the given rule, wrt.\ terms that refer - to the implicit context of the selected subgoal. - - An instantiation consists of a list of pairs of the form @{text - "(?x, t)"}, where @{text ?x} is a schematic variable occurring in - the given rule, and @{text t} is a term from the current proof - context, augmented by the local goal parameters of the selected - subgoal; cf.\ the @{text "focus"} operation described in - \secref{sec:variables}. - - Entering the syntactic context of a subgoal is a brittle operation, - because its exact form is somewhat accidental, and the choice of - bound variable names depends on the presence of other local and - global names. Explicit renaming of subgoal parameters prior to - explicit instantiation might help to achieve a bit more robustness. - - Type instantiations may be given as well, via pairs like @{text - "(?'a, \)"}. Type instantiations are distinguished from term - instantiations by the syntactic form of the schematic variable. - Types are instantiated before terms are. Since term instantiation - already performs type-inference as expected, explicit type - instantiations are seldom necessary. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML res_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\ - @{index_ML eres_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\ - @{index_ML dres_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\ - @{index_ML forw_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\[1ex] - @{index_ML rename_tac: "string list -> int -> tactic"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML res_inst_tac}~@{text "ctxt insts thm i"} instantiates the - rule @{text thm} with the instantiations @{text insts}, as described - above, and then performs resolution on subgoal @{text i}. - - \item @{ML eres_inst_tac} is like @{ML res_inst_tac}, but performs - elim-resolution. - - \item @{ML dres_inst_tac} is like @{ML res_inst_tac}, but performs - destruct-resolution. - - \item @{ML forw_inst_tac} is like @{ML dres_inst_tac} except that - the selected assumption is not deleted. - - \item @{ML rename_tac}~@{text "names i"} renames the innermost - parameters of subgoal @{text i} according to the provided @{text - names} (which need to be distinct indentifiers). - - \end{description} -*} - - -section {* Tacticals \label{sec:tacticals} *} - -text {* - -FIXME - -\glossary{Tactical}{A functional combinator for building up complex -tactics from simpler ones. Typical tactical perform sequential -composition, disjunction (choice), iteration, or goal addressing. -Various search strategies may be expressed via tacticals.} - -*} - -end - diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/Thy/unused.thy --- a/doc-src/IsarImplementation/Thy/unused.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,79 +0,0 @@ - -section {* Sessions and document preparation *} - -section {* Structured output *} - -subsection {* Pretty printing *} - -text FIXME - -subsection {* Output channels *} - -text FIXME - -subsection {* Print modes \label{sec:print-mode} *} - -text FIXME - -text {* - - - \medskip The general concept supports block-structured reasoning - nicely, with arbitrary mechanisms for introducing local assumptions. - The common reasoning pattern is as follows: - - \medskip - \begin{tabular}{l} - @{text "add_assms e\<^isub>1 A\<^isub>1"} \\ - @{text "\"} \\ - @{text "add_assms e\<^isub>n A\<^isub>n"} \\ - @{text "export"} \\ - \end{tabular} - \medskip - - \noindent The final @{text "export"} will turn any fact @{text - "A\<^isub>1, \, A\<^isub>n \ B"} into some @{text "\ B'"}, by - applying the export rules @{text "e\<^isub>1, \, e\<^isub>n"} - inside-out. - - - A \emph{fixed variable} acts like a local constant in the current - context, representing some simple type @{text "\"}, or some value - @{text "x: \"} (for a fixed type expression @{text "\"}). A - \emph{schematic variable} acts like a placeholder for arbitrary - elements, similar to outermost quantification. The division between - fixed and schematic variables tells which abstract entities are - inside and outside the current context. - - - @{index_ML Variable.trade: "Proof.context -> (thm list -> thm list) -> thm list -> thm list"} \\ - - - - \item @{ML Variable.trade} composes @{ML Variable.import} and @{ML - Variable.export}, i.e.\ it provides a view on facts with all - variables being fixed in the current context. - - - In practice, super-contexts emerge either by merging existing ones, - or by adding explicit declarations. For example, new theories are - usually derived by importing existing theories from the library - @{text "\ = \\<^sub>1 + \ + \\<^isub>n"}, or - - - - The Isar toplevel works differently for interactive developments - vs.\ batch processing of theory sources. For example, diagnostic - commands produce a warning batch mode, because they are considered - alien to the final theory document being produced eventually. - Moreover, full @{text undo} with intermediate checkpoints to protect - against destroying theories accidentally are limited to interactive - mode. In batch mode there is only a single strictly linear stream - of potentially desctructive theory transformations. - - \item @{ML Toplevel.empty} is an empty transition; the Isar command - dispatcher internally applies @{ML Toplevel.name} (for the command) - name and @{ML Toplevel.position} for the source position. - -*} - diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/checkglossary --- a/doc-src/IsarImplementation/checkglossary Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,28 +0,0 @@ -#!/usr/bin/env perl -# $Id$ - -use strict; - -my %defs = (); -my %refs = (); - -while () { - if (m,\\glossaryentry\{\w*\\bf *((\w|\s)+)@,) { - $defs{lc $1} = 1; - } - while (m,\\seeglossary *\{((\w|\s)+)\},g) { - $refs{lc $1} = 1; - } -} - -print "Glossary definitions:\n"; -foreach (sort(keys(%defs))) { - print " \"$_\"\n"; -} - -foreach (keys(%refs)) { - s,s$,,; - if (!defined($defs{$_})) { - print "### Undefined glossary reference: \"$_\"\n"; - } -} diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/implementation.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/intro.tex --- a/doc-src/IsarImplementation/intro.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,13 +0,0 @@ - -%% $Id$ - -\chapter{Introduction} - -FIXME - -\nocite{Wenzel-PhD} - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "implementation" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/makeglossary --- a/doc-src/IsarImplementation/makeglossary Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,6 +0,0 @@ -#!/bin/sh -# $Id$ - -NAME="$1" -makeindex -s nomencl -o "${NAME}.gls" "${NAME}.glo" -./checkglossary "${NAME}.glo" diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarImplementation/style.sty diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarOverview/Isar/document/.cvsignore --- a/doc-src/IsarOverview/Isar/document/.cvsignore Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2 +0,0 @@ -*.sty -session.tex \ No newline at end of file diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/IsaMakefile diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/First_Order_Logic.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Thy/First_Order_Logic.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,520 @@ + +header {* Example: First-Order Logic *} + +theory %visible First_Order_Logic +imports Pure +begin + +text {* + \noindent In order to commence a new object-logic within + Isabelle/Pure we introduce abstract syntactic categories @{text "i"} + for individuals and @{text "o"} for object-propositions. The latter + is embedded into the language of Pure propositions by means of a + separate judgment. +*} + +typedecl i +typedecl o + +judgment + Trueprop :: "o \ prop" ("_" 5) + +text {* + \noindent Note that the object-logic judgement is implicit in the + syntax: writing @{prop A} produces @{term "Trueprop A"} internally. + From the Pure perspective this means ``@{prop A} is derivable in the + object-logic''. +*} + + +subsection {* Equational reasoning \label{sec:framework-ex-equal} *} + +text {* + Equality is axiomatized as a binary predicate on individuals, with + reflexivity as introduction, and substitution as elimination + principle. Note that the latter is particularly convenient in a + framework like Isabelle, because syntactic congruences are + implicitly produced by unification of @{term "B x"} against + expressions containing occurrences of @{term x}. +*} + +axiomatization + equal :: "i \ i \ o" (infix "=" 50) +where + refl [intro]: "x = x" and + subst [elim]: "x = y \ B x \ B y" + +text {* + \noindent Substitution is very powerful, but also hard to control in + full generality. We derive some common symmetry~/ transitivity + schemes of as particular consequences. +*} + +theorem sym [sym]: + assumes "x = y" + shows "y = x" +proof - + have "x = x" .. + with `x = y` show "y = x" .. +qed + +theorem forw_subst [trans]: + assumes "y = x" and "B x" + shows "B y" +proof - + from `y = x` have "x = y" .. + from this and `B x` show "B y" .. +qed + +theorem back_subst [trans]: + assumes "B x" and "x = y" + shows "B y" +proof - + from `x = y` and `B x` + show "B y" .. +qed + +theorem trans [trans]: + assumes "x = y" and "y = z" + shows "x = z" +proof - + from `y = z` and `x = y` + show "x = z" .. +qed + + +subsection {* Basic group theory *} + +text {* + As an example for equational reasoning we consider some bits of + group theory. The subsequent locale definition postulates group + operations and axioms; we also derive some consequences of this + specification. +*} + +locale group = + fixes prod :: "i \ i \ i" (infix "\" 70) + and inv :: "i \ i" ("(_\)" [1000] 999) + and unit :: i ("1") + assumes assoc: "(x \ y) \ z = x \ (y \ z)" + and left_unit: "1 \ x = x" + and left_inv: "x\ \ x = 1" +begin + +theorem right_inv: "x \ x\ = 1" +proof - + have "x \ x\ = 1 \ (x \ x\)" by (rule left_unit [symmetric]) + also have "\ = (1 \ x) \ x\" by (rule assoc [symmetric]) + also have "1 = (x\)\ \ x\" by (rule left_inv [symmetric]) + also have "\ \ x = (x\)\ \ (x\ \ x)" by (rule assoc) + also have "x\ \ x = 1" by (rule left_inv) + also have "((x\)\ \ \) \ x\ = (x\)\ \ (1 \ x\)" by (rule assoc) + also have "1 \ x\ = x\" by (rule left_unit) + also have "(x\)\ \ \ = 1" by (rule left_inv) + finally show "x \ x\ = 1" . +qed + +theorem right_unit: "x \ 1 = x" +proof - + have "1 = x\ \ x" by (rule left_inv [symmetric]) + also have "x \ \ = (x \ x\) \ x" by (rule assoc [symmetric]) + also have "x \ x\ = 1" by (rule right_inv) + also have "\ \ x = x" by (rule left_unit) + finally show "x \ 1 = x" . +qed + +text {* + \noindent Reasoning from basic axioms is often tedious. Our proofs + work by producing various instances of the given rules (potentially + the symmetric form) using the pattern ``@{command have}~@{text + eq}~@{command "by"}~@{text "(rule r)"}'' and composing the chain of + results via @{command also}/@{command finally}. These steps may + involve any of the transitivity rules declared in + \secref{sec:framework-ex-equal}, namely @{thm trans} in combining + the first two results in @{thm right_inv} and in the final steps of + both proofs, @{thm forw_subst} in the first combination of @{thm + right_unit}, and @{thm back_subst} in all other calculational steps. + + Occasional substitutions in calculations are adequate, but should + not be over-emphasized. The other extreme is to compose a chain by + plain transitivity only, with replacements occurring always in + topmost position. For example: +*} + +(*<*) +theorem "\A. PROP A \ PROP A" +proof - + assume [symmetric, defn]: "\x y. (x \ y) \ Trueprop (x = y)" +(*>*) + have "x \ 1 = x \ (x\ \ x)" unfolding left_inv .. + also have "\ = (x \ x\) \ x" unfolding assoc .. + also have "\ = 1 \ x" unfolding right_inv .. + also have "\ = x" unfolding left_unit .. + finally have "x \ 1 = x" . +(*<*) +qed +(*>*) + +text {* + \noindent Here we have re-used the built-in mechanism for unfolding + definitions in order to normalize each equational problem. A more + realistic object-logic would include proper setup for the Simplifier + (\secref{sec:simplifier}), the main automated tool for equational + reasoning in Isabelle. Then ``@{command unfolding}~@{thm + left_inv}~@{command ".."}'' would become ``@{command "by"}~@{text + "(simp only: left_inv)"}'' etc. +*} + +end + + +subsection {* Propositional logic \label{sec:framework-ex-prop} *} + +text {* + We axiomatize basic connectives of propositional logic: implication, + disjunction, and conjunction. The associated rules are modeled + after Gentzen's system of Natural Deduction \cite{Gentzen:1935}. +*} + +axiomatization + imp :: "o \ o \ o" (infixr "\" 25) where + impI [intro]: "(A \ B) \ A \ B" and + impD [dest]: "(A \ B) \ A \ B" + +axiomatization + disj :: "o \ o \ o" (infixr "\" 30) where + disjI\<^isub>1 [intro]: "A \ A \ B" and + disjI\<^isub>2 [intro]: "B \ A \ B" and + disjE [elim]: "A \ B \ (A \ C) \ (B \ C) \ C" + +axiomatization + conj :: "o \ o \ o" (infixr "\" 35) where + conjI [intro]: "A \ B \ A \ B" and + conjD\<^isub>1: "A \ B \ A" and + conjD\<^isub>2: "A \ B \ B" + +text {* + \noindent The conjunctive destructions have the disadvantage that + decomposing @{prop "A \ B"} involves an immediate decision which + component should be projected. The more convenient simultaneous + elimination @{prop "A \ B \ (A \ B \ C) \ C"} can be derived as + follows: +*} + +theorem conjE [elim]: + assumes "A \ B" + obtains A and B +proof + from `A \ B` show A by (rule conjD\<^isub>1) + from `A \ B` show B by (rule conjD\<^isub>2) +qed + +text {* + \noindent Here is an example of swapping conjuncts with a single + intermediate elimination step: +*} + +(*<*) +lemma "\A. PROP A \ PROP A" +proof - +(*>*) + assume "A \ B" + then obtain B and A .. + then have "B \ A" .. +(*<*) +qed +(*>*) + +text {* + \noindent Note that the analogous elimination rule for disjunction + ``@{text "\ A \ B \ A \ B"}'' coincides with + the original axiomatization of @{thm disjE}. + + \medskip We continue propositional logic by introducing absurdity + with its characteristic elimination. Plain truth may then be + defined as a proposition that is trivially true. +*} + +axiomatization + false :: o ("\") where + falseE [elim]: "\ \ A" + +definition + true :: o ("\") where + "\ \ \ \ \" + +theorem trueI [intro]: \ + unfolding true_def .. + +text {* + \medskip\noindent Now negation represents an implication towards + absurdity: +*} + +definition + not :: "o \ o" ("\ _" [40] 40) where + "\ A \ A \ \" + +theorem notI [intro]: + assumes "A \ \" + shows "\ A" +unfolding not_def +proof + assume A + then show \ by (rule `A \ \`) +qed + +theorem notE [elim]: + assumes "\ A" and A + shows B +proof - + from `\ A` have "A \ \" unfolding not_def . + from `A \ \` and `A` have \ .. + then show B .. +qed + + +subsection {* Classical logic *} + +text {* + Subsequently we state the principle of classical contradiction as a + local assumption. Thus we refrain from forcing the object-logic + into the classical perspective. Within that context, we may derive + well-known consequences of the classical principle. +*} + +locale classical = + assumes classical: "(\ C \ C) \ C" +begin + +theorem double_negation: + assumes "\ \ C" + shows C +proof (rule classical) + assume "\ C" + with `\ \ C` show C .. +qed + +theorem tertium_non_datur: "C \ \ C" +proof (rule double_negation) + show "\ \ (C \ \ C)" + proof + assume "\ (C \ \ C)" + have "\ C" + proof + assume C then have "C \ \ C" .. + with `\ (C \ \ C)` show \ .. + qed + then have "C \ \ C" .. + with `\ (C \ \ C)` show \ .. + qed +qed + +text {* + \noindent These examples illustrate both classical reasoning and + non-trivial propositional proofs in general. All three rules + characterize classical logic independently, but the original rule is + already the most convenient to use, because it leaves the conclusion + unchanged. Note that @{prop "(\ C \ C) \ C"} fits again into our + format for eliminations, despite the additional twist that the + context refers to the main conclusion. So we may write @{thm + classical} as the Isar statement ``@{text "\ \ thesis"}''. + This also explains nicely how classical reasoning really works: + whatever the main @{text thesis} might be, we may always assume its + negation! +*} + +end + + +subsection {* Quantifiers \label{sec:framework-ex-quant} *} + +text {* + Representing quantifiers is easy, thanks to the higher-order nature + of the underlying framework. According to the well-known technique + introduced by Church \cite{church40}, quantifiers are operators on + predicates, which are syntactically represented as @{text "\"}-terms + of type @{typ "i \ o"}. Binder notation turns @{text "All (\x. B + x)"} into @{text "\x. B x"} etc. +*} + +axiomatization + All :: "(i \ o) \ o" (binder "\" 10) where + allI [intro]: "(\x. B x) \ \x. B x" and + allD [dest]: "(\x. B x) \ B a" + +axiomatization + Ex :: "(i \ o) \ o" (binder "\" 10) where + exI [intro]: "B a \ (\x. B x)" and + exE [elim]: "(\x. B x) \ (\x. B x \ C) \ C" + +text {* + \noindent The statement of @{thm exE} corresponds to ``@{text + "\ \x. B x \ x \ B x"}'' in Isar. In the + subsequent example we illustrate quantifier reasoning involving all + four rules: +*} + +theorem + assumes "\x. \y. R x y" + shows "\y. \x. R x y" +proof -- {* @{text "\"} introduction *} + obtain x where "\y. R x y" using `\x. \y. R x y` .. -- {* @{text "\"} elimination *} + fix y have "R x y" using `\y. R x y` .. -- {* @{text "\"} destruction *} + then show "\x. R x y" .. -- {* @{text "\"} introduction *} +qed + + +subsection {* Canonical reasoning patterns *} + +text {* + The main rules of first-order predicate logic from + \secref{sec:framework-ex-prop} and \secref{sec:framework-ex-quant} + can now be summarized as follows, using the native Isar statement + format of \secref{sec:framework-stmt}. + + \medskip + \begin{tabular}{l} + @{text "impI: \ A \ B \ A \ B"} \\ + @{text "impD: \ A \ B \ A \ B"} \\[1ex] + + @{text "disjI\<^isub>1: \ A \ A \ B"} \\ + @{text "disjI\<^isub>2: \ B \ A \ B"} \\ + @{text "disjE: \ A \ B \ A \ B"} \\[1ex] + + @{text "conjI: \ A \ B \ A \ B"} \\ + @{text "conjE: \ A \ B \ A \ B"} \\[1ex] + + @{text "falseE: \ \ \ A"} \\ + @{text "trueI: \ \"} \\[1ex] + + @{text "notI: \ A \ \ \ \ A"} \\ + @{text "notE: \ \ A \ A \ B"} \\[1ex] + + @{text "allI: \ \x. B x \ \x. B x"} \\ + @{text "allE: \ \x. B x \ B a"} \\[1ex] + + @{text "exI: \ B a \ \x. B x"} \\ + @{text "exE: \ \x. B x \ a \ B a"} + \end{tabular} + \medskip + + \noindent This essentially provides a declarative reading of Pure + rules as Isar reasoning patterns: the rule statements tells how a + canonical proof outline shall look like. Since the above rules have + already been declared as @{attribute (Pure) intro}, @{attribute + (Pure) elim}, @{attribute (Pure) dest} --- each according to its + particular shape --- we can immediately write Isar proof texts as + follows: +*} + +(*<*) +theorem "\A. PROP A \ PROP A" +proof - +(*>*) + + txt_raw {*\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "A \ B" + proof + assume A + show B sorry %noproof + qed + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "A \ B" and A sorry %noproof + then have B .. + + txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have A sorry %noproof + then have "A \ B" .. + + have B sorry %noproof + then have "A \ B" .. + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "A \ B" sorry %noproof + then have C + proof + assume A + then show C sorry %noproof + next + assume B + then show C sorry %noproof + qed + + txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have A and B sorry %noproof + then have "A \ B" .. + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "A \ B" sorry %noproof + then obtain A and B .. + + txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\" sorry %noproof + then have A .. + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\" .. + + txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\ A" + proof + assume A + then show "\" sorry %noproof + qed + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\ A" and A sorry %noproof + then have B .. + + txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\x. B x" + proof + fix x + show "B x" sorry %noproof + qed + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\x. B x" sorry %noproof + then have "B a" .. + + txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\x. B x" + proof + show "B a" sorry %noproof + qed + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\x. B x" sorry %noproof + then obtain a where "B a" .. + + txt_raw {*\end{minipage}*} + +(*<*) +qed +(*>*) + +text {* + \bigskip\noindent Of course, these proofs are merely examples. As + sketched in \secref{sec:framework-subproof}, there is a fair amount + of flexibility in expressing Pure deductions in Isar. Here the user + is asked to express himself adequately, aiming at proof texts of + literary quality. +*} + +end %visible diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/Framework.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Thy/Framework.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,1017 @@ +theory Framework +imports Main +begin + +chapter {* The Isabelle/Isar Framework \label{ch:isar-framework} *} + +text {* + Isabelle/Isar + \cite{Wenzel:1999:TPHOL,Wenzel-PhD,Nipkow-TYPES02,Wenzel-Paulson:2006,Wenzel:2006:Festschrift} + is intended as a generic framework for developing formal + mathematical documents with full proof checking. Definitions and + proofs are organized as theories. An assembly of theory sources may + be presented as a printed document; see also + \chref{ch:document-prep}. + + The main objective of Isar is the design of a human-readable + structured proof language, which is called the ``primary proof + format'' in Isar terminology. Such a primary proof language is + somewhere in the middle between the extremes of primitive proof + objects and actual natural language. In this respect, Isar is a bit + more formalistic than Mizar + \cite{Trybulec:1993:MizarFeatures,Rudnicki:1992:MizarOverview,Wiedijk:1999:Mizar}, + using logical symbols for certain reasoning schemes where Mizar + would prefer English words; see \cite{Wenzel-Wiedijk:2002} for + further comparisons of these systems. + + So Isar challenges the traditional way of recording informal proofs + in mathematical prose, as well as the common tendency to see fully + formal proofs directly as objects of some logical calculus (e.g.\ + @{text "\"}-terms in a version of type theory). In fact, Isar is + better understood as an interpreter of a simple block-structured + language for describing the data flow of local facts and goals, + interspersed with occasional invocations of proof methods. + Everything is reduced to logical inferences internally, but these + steps are somewhat marginal compared to the overall bookkeeping of + the interpretation process. Thanks to careful design of the syntax + and semantics of Isar language elements, a formal record of Isar + instructions may later appear as an intelligible text to the + attentive reader. + + The Isar proof language has emerged from careful analysis of some + inherent virtues of the existing logical framework of Isabelle/Pure + \cite{paulson-found,paulson700}, notably composition of higher-order + natural deduction rules, which is a generalization of Gentzen's + original calculus \cite{Gentzen:1935}. The approach of generic + inference systems in Pure is continued by Isar towards actual proof + texts. + + Concrete applications require another intermediate layer: an + object-logic. Isabelle/HOL \cite{isa-tutorial} (simply-typed + set-theory) is being used most of the time; Isabelle/ZF + \cite{isabelle-ZF} is less extensively developed, although it would + probably fit better for classical mathematics. + + \medskip In order to illustrate natural deduction in Isar, we shall + refer to the background theory and library of Isabelle/HOL. This + includes common notions of predicate logic, naive set-theory etc.\ + using fairly standard mathematical notation. From the perspective + of generic natural deduction there is nothing special about the + logical connectives of HOL (@{text "\"}, @{text "\"}, @{text "\"}, + @{text "\"}, etc.), only the resulting reasoning principles are + relevant to the user. There are similar rules available for + set-theory operators (@{text "\"}, @{text "\"}, @{text "\"}, @{text + "\"}, etc.), or any other theory developed in the library (lattice + theory, topology etc.). + + Subsequently we briefly review fragments of Isar proof texts + corresponding directly to such general deduction schemes. The + examples shall refer to set-theory, to minimize the danger of + understanding connectives of predicate logic as something special. + + \medskip The following deduction performs @{text "\"}-introduction, + working forwards from assumptions towards the conclusion. We give + both the Isar text, and depict the primitive rule involved, as + determined by unification of the problem against rules that are + declared in the library context. +*} + +text_raw {*\medskip\begin{minipage}{0.6\textwidth}*} + +(*<*) +lemma True +proof +(*>*) + assume "x \ A" and "x \ B" + then have "x \ A \ B" .. +(*<*) +qed +(*>*) + +text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*} + +text {* + \infer{@{prop "x \ A \ B"}}{@{prop "x \ A"} & @{prop "x \ B"}} +*} + +text_raw {*\end{minipage}*} + +text {* + \medskip\noindent Note that @{command assume} augments the proof + context, @{command then} indicates that the current fact shall be + used in the next step, and @{command have} states an intermediate + goal. The two dots ``@{command ".."}'' refer to a complete proof of + this claim, using the indicated facts and a canonical rule from the + context. We could have been more explicit here by spelling out the + final proof step via the @{command "by"} command: +*} + +(*<*) +lemma True +proof +(*>*) + assume "x \ A" and "x \ B" + then have "x \ A \ B" by (rule IntI) +(*<*) +qed +(*>*) + +text {* + \noindent The format of the @{text "\"}-introduction rule represents + the most basic inference, which proceeds from given premises to a + conclusion, without any nested proof context involved. + + The next example performs backwards introduction on @{term "\\"}, + the intersection of all sets within a given set. This requires a + nested proof of set membership within a local context, where @{term + A} is an arbitrary-but-fixed member of the collection: +*} + +text_raw {*\medskip\begin{minipage}{0.6\textwidth}*} + +(*<*) +lemma True +proof +(*>*) + have "x \ \\" + proof + fix A + assume "A \ \" + show "x \ A" sorry %noproof + qed +(*<*) +qed +(*>*) + +text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*} + +text {* + \infer{@{prop "x \ \\"}}{\infer*{@{prop "x \ A"}}{@{text "[A][A \ \]"}}} +*} + +text_raw {*\end{minipage}*} + +text {* + \medskip\noindent This Isar reasoning pattern again refers to the + primitive rule depicted above. The system determines it in the + ``@{command proof}'' step, which could have been spelt out more + explicitly as ``@{command proof}~@{text "(rule InterI)"}''. Note + that the rule involves both a local parameter @{term "A"} and an + assumption @{prop "A \ \"} in the nested reasoning. This kind of + compound rule typically demands a genuine sub-proof in Isar, working + backwards rather than forwards as seen before. In the proof body we + encounter the @{command fix}-@{command assume}-@{command show} + outline of nested sub-proofs that is typical for Isar. The final + @{command show} is like @{command have} followed by an additional + refinement of the enclosing claim, using the rule derived from the + proof body. + + \medskip The next example involves @{term "\\"}, which can be + characterized as the set of all @{term "x"} such that @{prop "\A. x + \ A \ A \ \"}. The elimination rule for @{prop "x \ \\"} does + not mention @{text "\"} and @{text "\"} at all, but admits to obtain + directly a local @{term "A"} such that @{prop "x \ A"} and @{prop "A + \ \"} hold. This corresponds to the following Isar proof and + inference rule, respectively: +*} + +text_raw {*\medskip\begin{minipage}{0.6\textwidth}*} + +(*<*) +lemma True +proof +(*>*) + assume "x \ \\" + then have C + proof + fix A + assume "x \ A" and "A \ \" + show C sorry %noproof + qed +(*<*) +qed +(*>*) + +text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*} + +text {* + \infer{@{prop "C"}}{@{prop "x \ \\"} & \infer*{@{prop "C"}~}{@{text "[A][x \ A, A \ \]"}}} +*} + +text_raw {*\end{minipage}*} + +text {* + \medskip\noindent Although the Isar proof follows the natural + deduction rule closely, the text reads not as natural as + anticipated. There is a double occurrence of an arbitrary + conclusion @{prop "C"}, which represents the final result, but is + irrelevant for now. This issue arises for any elimination rule + involving local parameters. Isar provides the derived language + element @{command obtain}, which is able to perform the same + elimination proof more conveniently: +*} + +(*<*) +lemma True +proof +(*>*) + assume "x \ \\" + then obtain A where "x \ A" and "A \ \" .. +(*<*) +qed +(*>*) + +text {* + \noindent Here we avoid to mention the final conclusion @{prop "C"} + and return to plain forward reasoning. The rule involved in the + ``@{command ".."}'' proof is the same as before. +*} + + +section {* The Pure framework \label{sec:framework-pure} *} + +text {* + The Pure logic \cite{paulson-found,paulson700} is an intuitionistic + fragment of higher-order logic \cite{church40}. In type-theoretic + parlance, there are three levels of @{text "\"}-calculus with + corresponding arrows @{text "\"}/@{text "\"}/@{text "\"}: + + \medskip + \begin{tabular}{ll} + @{text "\ \ \"} & syntactic function space (terms depending on terms) \\ + @{text "\x. B(x)"} & universal quantification (proofs depending on terms) \\ + @{text "A \ B"} & implication (proofs depending on proofs) \\ + \end{tabular} + \medskip + + \noindent Here only the types of syntactic terms, and the + propositions of proof terms have been shown. The @{text + "\"}-structure of proofs can be recorded as an optional feature of + the Pure inference kernel \cite{Berghofer-Nipkow:2000:TPHOL}, but + the formal system can never depend on them due to \emph{proof + irrelevance}. + + On top of this most primitive layer of proofs, Pure implements a + generic calculus for nested natural deduction rules, similar to + \cite{Schroeder-Heister:1984}. Here object-logic inferences are + internalized as formulae over @{text "\"} and @{text "\"}. + Combining such rule statements may involve higher-order unification + \cite{paulson-natural}. +*} + + +subsection {* Primitive inferences *} + +text {* + Term syntax provides explicit notation for abstraction @{text "\x :: + \. b(x)"} and application @{text "b a"}, while types are usually + implicit thanks to type-inference; terms of type @{text "prop"} are + called propositions. Logical statements are composed via @{text "\x + :: \. B(x)"} and @{text "A \ B"}. Primitive reasoning operates on + judgments of the form @{text "\ \ \"}, with standard introduction + and elimination rules for @{text "\"} and @{text "\"} that refer to + fixed parameters @{text "x\<^isub>1, \, x\<^isub>m"} and hypotheses + @{text "A\<^isub>1, \, A\<^isub>n"} from the context @{text "\"}; + the corresponding proof terms are left implicit. The subsequent + inference rules define @{text "\ \ \"} inductively, relative to a + collection of axioms: + + \[ + \infer{@{text "\ A"}}{(@{text "A"} \text{~axiom})} + \qquad + \infer{@{text "A \ A"}}{} + \] + + \[ + \infer{@{text "\ \ \x. B(x)"}}{@{text "\ \ B(x)"} & @{text "x \ \"}} + \qquad + \infer{@{text "\ \ B(a)"}}{@{text "\ \ \x. B(x)"}} + \] + + \[ + \infer{@{text "\ - A \ A \ B"}}{@{text "\ \ B"}} + \qquad + \infer{@{text "\\<^sub>1 \ \\<^sub>2 \ B"}}{@{text "\\<^sub>1 \ A \ B"} & @{text "\\<^sub>2 \ A"}} + \] + + Furthermore, Pure provides a built-in equality @{text "\ :: \ \ \ \ + prop"} with axioms for reflexivity, substitution, extensionality, + and @{text "\\\"}-conversion on @{text "\"}-terms. + + \medskip An object-logic introduces another layer on top of Pure, + e.g.\ with types @{text "i"} for individuals and @{text "o"} for + propositions, term constants @{text "Trueprop :: o \ prop"} as + (implicit) derivability judgment and connectives like @{text "\ :: o + \ o \ o"} or @{text "\ :: (i \ o) \ o"}, and axioms for object-level + rules such as @{text "conjI: A \ B \ A \ B"} or @{text "allI: (\x. B + x) \ \x. B x"}. Derived object rules are represented as theorems of + Pure. After the initial object-logic setup, further axiomatizations + are usually avoided; plain definitions and derived principles are + used exclusively. +*} + + +subsection {* Reasoning with rules \label{sec:framework-resolution} *} + +text {* + Primitive inferences mostly serve foundational purposes. The main + reasoning mechanisms of Pure operate on nested natural deduction + rules expressed as formulae, using @{text "\"} to bind local + parameters and @{text "\"} to express entailment. Multiple + parameters and premises are represented by repeating these + connectives in a right-associative manner. + + Since @{text "\"} and @{text "\"} commute thanks to the theorem + @{prop "(A \ (\x. B x)) \ (\x. A \ B x)"}, we may assume w.l.o.g.\ + that rule statements always observe the normal form where + quantifiers are pulled in front of implications at each level of + nesting. This means that any Pure proposition may be presented as a + \emph{Hereditary Harrop Formula} \cite{Miller:1991} which is of the + form @{text "\x\<^isub>1 \ x\<^isub>m. H\<^isub>1 \ \ H\<^isub>n \ + A"} for @{text "m, n \ 0"}, and @{text "A"} atomic, and @{text + "H\<^isub>1, \, H\<^isub>n"} being recursively of the same format. + Following the convention that outermost quantifiers are implicit, + Horn clauses @{text "A\<^isub>1 \ \ A\<^isub>n \ A"} are a special + case of this. + + For example, @{text "\"}-introduction rule encountered before is + represented as a Pure theorem as follows: + \[ + @{text "IntI:"}~@{prop "x \ A \ x \ B \ x \ A \ B"} + \] + + \noindent This is a plain Horn clause, since no further nesting on + the left is involved. The general @{text "\"}-introduction + corresponds to a Hereditary Harrop Formula with one additional level + of nesting: + \[ + @{text "InterI:"}~@{prop "(\A. A \ \ \ x \ A) \ x \ \\"} + \] + + \medskip Goals are also represented as rules: @{text "A\<^isub>1 \ + \ A\<^isub>n \ C"} states that the sub-goals @{text "A\<^isub>1, \, + A\<^isub>n"} entail the result @{text "C"}; for @{text "n = 0"} the + goal is finished. To allow @{text "C"} being a rule statement + itself, we introduce the protective marker @{text "# :: prop \ + prop"}, which is defined as identity and hidden from the user. We + initialize and finish goal states as follows: + + \[ + \begin{array}{c@ {\qquad}c} + \infer[(@{inference_def init})]{@{text "C \ #C"}}{} & + \infer[(@{inference_def finish})]{@{text C}}{@{text "#C"}} + \end{array} + \] + + \noindent Goal states are refined in intermediate proof steps until + a finished form is achieved. Here the two main reasoning principles + are @{inference resolution}, for back-chaining a rule against a + sub-goal (replacing it by zero or more sub-goals), and @{inference + assumption}, for solving a sub-goal (finding a short-circuit with + local assumptions). Below @{text "\<^vec>x"} stands for @{text + "x\<^isub>1, \, x\<^isub>n"} (@{text "n \ 0"}). + + \[ + \infer[(@{inference_def resolution})] + {@{text "(\\<^vec>x. \<^vec>H \<^vec>x \ \<^vec>A (\<^vec>a \<^vec>x))\ \ C\"}} + {\begin{tabular}{rl} + @{text "rule:"} & + @{text "\<^vec>A \<^vec>a \ B \<^vec>a"} \\ + @{text "goal:"} & + @{text "(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C"} \\ + @{text "goal unifier:"} & + @{text "(\\<^vec>x. B (\<^vec>a \<^vec>x))\ = B'\"} \\ + \end{tabular}} + \] + + \medskip + + \[ + \infer[(@{inference_def assumption})]{@{text "C\"}} + {\begin{tabular}{rl} + @{text "goal:"} & + @{text "(\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x) \ C"} \\ + @{text "assm unifier:"} & @{text "A\ = H\<^sub>i\"}~~\text{(for some~@{text "H\<^sub>i"})} \\ + \end{tabular}} + \] + + The following trace illustrates goal-oriented reasoning in + Isabelle/Pure: + + {\footnotesize + \medskip + \begin{tabular}{r@ {\quad}l} + @{text "(A \ B \ B \ A) \ #(A \ B \ B \ A)"} & @{text "(init)"} \\ + @{text "(A \ B \ B) \ (A \ B \ A) \ #\"} & @{text "(resolution B \ A \ B \ A)"} \\ + @{text "(A \ B \ A \ B) \ (A \ B \ A) \ #\"} & @{text "(resolution A \ B \ B)"} \\ + @{text "(A \ B \ A) \ #\"} & @{text "(assumption)"} \\ + @{text "(A \ B \ B \ A) \ #\"} & @{text "(resolution A \ B \ A)"} \\ + @{text "#\"} & @{text "(assumption)"} \\ + @{text "A \ B \ B \ A"} & @{text "(finish)"} \\ + \end{tabular} + \medskip + } + + Compositions of @{inference assumption} after @{inference + resolution} occurs quite often, typically in elimination steps. + Traditional Isabelle tactics accommodate this by a combined + @{inference_def elim_resolution} principle. In contrast, Isar uses + a slightly more refined combination, where the assumptions to be + closed are marked explicitly, using again the protective marker + @{text "#"}: + + \[ + \infer[(@{inference refinement})] + {@{text "(\\<^vec>x. \<^vec>H \<^vec>x \ \<^vec>G' (\<^vec>a \<^vec>x))\ \ C\"}} + {\begin{tabular}{rl} + @{text "sub\proof:"} & + @{text "\<^vec>G \<^vec>a \ B \<^vec>a"} \\ + @{text "goal:"} & + @{text "(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C"} \\ + @{text "goal unifier:"} & + @{text "(\\<^vec>x. B (\<^vec>a \<^vec>x))\ = B'\"} \\ + @{text "assm unifiers:"} & + @{text "(\\<^vec>x. G\<^sub>j (\<^vec>a \<^vec>x))\ = #H\<^sub>i\"} \\ + & \quad (for each marked @{text "G\<^sub>j"} some @{text "#H\<^sub>i"}) \\ + \end{tabular}} + \] + + \noindent Here the @{text "sub\proof"} rule stems from the + main @{command fix}-@{command assume}-@{command show} outline of + Isar (cf.\ \secref{sec:framework-subproof}): each assumption + indicated in the text results in a marked premise @{text "G"} above. + The marking enforces resolution against one of the sub-goal's + premises. Consequently, @{command fix}-@{command assume}-@{command + show} enables to fit the result of a sub-proof quite robustly into a + pending sub-goal, while maintaining a good measure of flexibility. +*} + + +section {* The Isar proof language \label{sec:framework-isar} *} + +text {* + Structured proofs are presented as high-level expressions for + composing entities of Pure (propositions, facts, and goals). The + Isar proof language allows to organize reasoning within the + underlying rule calculus of Pure, but Isar is not another logical + calculus! + + Isar is an exercise in sound minimalism. Approximately half of the + language is introduced as primitive, the rest defined as derived + concepts. The following grammar describes the core language + (category @{text "proof"}), which is embedded into theory + specification elements such as @{command theorem}; see also + \secref{sec:framework-stmt} for the separate category @{text + "statement"}. + + \medskip + \begin{tabular}{rcl} + @{text "theory\stmt"} & = & @{command "theorem"}~@{text "statement proof |"}~~@{command "definition"}~@{text "\ | \"} \\[1ex] + + @{text "proof"} & = & @{text "prfx\<^sup>*"}~@{command "proof"}~@{text "method\<^sup>? stmt\<^sup>*"}~@{command "qed"}~@{text "method\<^sup>?"} \\[1ex] + + @{text prfx} & = & @{command "using"}~@{text "facts"} \\ + & @{text "|"} & @{command "unfolding"}~@{text "facts"} \\ + + @{text stmt} & = & @{command "{"}~@{text "stmt\<^sup>*"}~@{command "}"} \\ + & @{text "|"} & @{command "next"} \\ + & @{text "|"} & @{command "note"}~@{text "name = facts"} \\ + & @{text "|"} & @{command "let"}~@{text "term = term"} \\ + & @{text "|"} & @{command "fix"}~@{text "var\<^sup>+"} \\ + & @{text "|"} & @{command assume}~@{text "\inference\ name: props"} \\ + & @{text "|"} & @{command "then"}@{text "\<^sup>?"}~@{text goal} \\ + @{text goal} & = & @{command "have"}~@{text "name: props proof"} \\ + & @{text "|"} & @{command "show"}~@{text "name: props proof"} \\ + \end{tabular} + + \medskip Simultaneous propositions or facts may be separated by the + @{keyword "and"} keyword. + + \medskip The syntax for terms and propositions is inherited from + Pure (and the object-logic). A @{text "pattern"} is a @{text + "term"} with schematic variables, to be bound by higher-order + matching. + + \medskip Facts may be referenced by name or proposition. For + example, the result of ``@{command have}~@{text "a: A \proof\"}'' + becomes available both as @{text "a"} and + \isacharbackquoteopen@{text "A"}\isacharbackquoteclose. Moreover, + fact expressions may involve attributes that modify either the + theorem or the background context. For example, the expression + ``@{text "a [OF b]"}'' refers to the composition of two facts + according to the @{inference resolution} inference of + \secref{sec:framework-resolution}, while ``@{text "a [intro]"}'' + declares a fact as introduction rule in the context. + + The special fact called ``@{fact this}'' always refers to the last + result, as produced by @{command note}, @{command assume}, @{command + have}, or @{command show}. Since @{command note} occurs + frequently together with @{command then} we provide some + abbreviations: + + \medskip + \begin{tabular}{rcl} + @{command from}~@{text a} & @{text "\"} & @{command note}~@{text a}~@{command then} \\ + @{command with}~@{text a} & @{text "\"} & @{command from}~@{text "a \ this"} \\ + \end{tabular} + \medskip + + The @{text "method"} category is essentially a parameter and may be + populated later. Methods use the facts indicated by @{command + "then"} or @{command using}, and then operate on the goal state. + Some basic methods are predefined: ``@{method "-"}'' leaves the goal + unchanged, ``@{method this}'' applies the facts as rules to the + goal, ``@{method "rule"}'' applies the facts to another rule and the + result to the goal (both ``@{method this}'' and ``@{method rule}'' + refer to @{inference resolution} of + \secref{sec:framework-resolution}). The secondary arguments to + ``@{method rule}'' may be specified explicitly as in ``@{text "(rule + a)"}'', or picked from the context. In the latter case, the system + first tries rules declared as @{attribute (Pure) elim} or + @{attribute (Pure) dest}, followed by those declared as @{attribute + (Pure) intro}. + + The default method for @{command proof} is ``@{method rule}'' + (arguments picked from the context), for @{command qed} it is + ``@{method "-"}''. Further abbreviations for terminal proof steps + are ``@{command "by"}~@{text "method\<^sub>1 method\<^sub>2"}'' for + ``@{command proof}~@{text "method\<^sub>1"}~@{command qed}~@{text + "method\<^sub>2"}'', and ``@{command ".."}'' for ``@{command + "by"}~@{method rule}, and ``@{command "."}'' for ``@{command + "by"}~@{method this}''. The @{command unfolding} element operates + directly on the current facts and goal by applying equalities. + + \medskip Block structure can be indicated explicitly by ``@{command + "{"}~@{text "\"}~@{command "}"}'', although the body of a sub-proof + already involves implicit nesting. In any case, @{command next} + jumps into the next section of a block, i.e.\ it acts like closing + an implicit block scope and opening another one; there is no direct + correspondence to subgoals here. + + The remaining elements @{command fix} and @{command assume} build up + a local context (see \secref{sec:framework-context}), while + @{command show} refines a pending sub-goal by the rule resulting + from a nested sub-proof (see \secref{sec:framework-subproof}). + Further derived concepts will support calculational reasoning (see + \secref{sec:framework-calc}). +*} + + +subsection {* Context elements \label{sec:framework-context} *} + +text {* + In judgments @{text "\ \ \"} of the primitive framework, @{text "\"} + essentially acts like a proof context. Isar elaborates this idea + towards a higher-level notion, with additional information for + type-inference, term abbreviations, local facts, hypotheses etc. + + The element @{command fix}~@{text "x :: \"} declares a local + parameter, i.e.\ an arbitrary-but-fixed entity of a given type; in + results exported from the context, @{text "x"} may become anything. + The @{command assume}~@{text "\inference\"} element provides a + general interface to hypotheses: ``@{command assume}~@{text + "\inference\ A"}'' produces @{text "A \ A"} locally, while the + included inference tells how to discharge @{text A} from results + @{text "A \ B"} later on. There is no user-syntax for @{text + "\inference\"}, i.e.\ it may only occur internally when derived + commands are defined in ML. + + At the user-level, the default inference for @{command assume} is + @{inference discharge} as given below. The additional variants + @{command presume} and @{command def} are defined as follows: + + \medskip + \begin{tabular}{rcl} + @{command presume}~@{text A} & @{text "\"} & @{command assume}~@{text "\weak\discharge\ A"} \\ + @{command def}~@{text "x \ a"} & @{text "\"} & @{command fix}~@{text x}~@{command assume}~@{text "\expansion\ x \ a"} \\ + \end{tabular} + \medskip + + \[ + \infer[(@{inference_def discharge})]{@{text "\\ - A \ #A \ B"}}{@{text "\\ \ B"}} + \] + \[ + \infer[(@{inference_def "weak\discharge"})]{@{text "\\ - A \ A \ B"}}{@{text "\\ \ B"}} + \] + \[ + \infer[(@{inference_def expansion})]{@{text "\\ - (x \ a) \ B a"}}{@{text "\\ \ B x"}} + \] + + \medskip Note that @{inference discharge} and @{inference + "weak\discharge"} differ in the marker for @{prop A}, which is + relevant when the result of a @{command fix}-@{command + assume}-@{command show} outline is composed with a pending goal, + cf.\ \secref{sec:framework-subproof}. + + The most interesting derived context element in Isar is @{command + obtain} \cite[\S5.3]{Wenzel-PhD}, which supports generalized + elimination steps in a purely forward manner. The @{command obtain} + command takes a specification of parameters @{text "\<^vec>x"} and + assumptions @{text "\<^vec>A"} to be added to the context, together + with a proof of a case rule stating that this extension is + conservative (i.e.\ may be removed from closed results later on): + + \medskip + \begin{tabular}{l} + @{text "\facts\"}~~@{command obtain}~@{text "\<^vec>x \ \<^vec>A \<^vec>x \proof\ \"} \\[0.5ex] + \quad @{command have}~@{text "case: \thesis. (\\<^vec>x. \<^vec>A \<^vec>x \ thesis) \ thesis\"} \\ + \quad @{command proof}~@{method "-"} \\ + \qquad @{command fix}~@{text thesis} \\ + \qquad @{command assume}~@{text "[intro]: \\<^vec>x. \<^vec>A \<^vec>x \ thesis"} \\ + \qquad @{command show}~@{text thesis}~@{command using}~@{text "\facts\ \proof\"} \\ + \quad @{command qed} \\ + \quad @{command fix}~@{text "\<^vec>x"}~@{command assume}~@{text "\elimination case\ \<^vec>A \<^vec>x"} \\ + \end{tabular} + \medskip + + \[ + \infer[(@{inference elimination})]{@{text "\ \ B"}}{ + \begin{tabular}{rl} + @{text "case:"} & + @{text "\ \ \thesis. (\\<^vec>x. \<^vec>A \<^vec>x \ thesis) \ thesis"} \\[0.2ex] + @{text "result:"} & + @{text "\ \ \<^vec>A \<^vec>y \ B"} \\[0.2ex] + \end{tabular}} + \] + + \noindent Here the name ``@{text thesis}'' is a specific convention + for an arbitrary-but-fixed proposition; in the primitive natural + deduction rules shown before we have occasionally used @{text C}. + The whole statement of ``@{command obtain}~@{text x}~@{keyword + "where"}~@{text "A x"}'' may be read as a claim that @{text "A x"} + may be assumed for some arbitrary-but-fixed @{text "x"}. Also note + that ``@{command obtain}~@{text "A \ B"}'' without parameters + is similar to ``@{command have}~@{text "A \ B"}'', but the + latter involves multiple sub-goals. + + \medskip The subsequent Isar proof texts explain all context + elements introduced above using the formal proof language itself. + After finishing a local proof within a block, we indicate the + exported result via @{command note}. +*} + +(*<*) +theorem True +proof +(*>*) + txt_raw {* \begin{minipage}[t]{0.4\textwidth} *} + { + fix x + have "B x" sorry %noproof + } + note `\x. B x` + txt_raw {* \end{minipage}\quad\begin{minipage}[t]{0.4\textwidth} *}(*<*)next(*>*) + { + assume A + have B sorry %noproof + } + note `A \ B` + txt_raw {* \end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} *}(*<*)next(*>*) + { + def x \ a + have "B x" sorry %noproof + } + note `B a` + txt_raw {* \end{minipage}\quad\begin{minipage}[t]{0.4\textwidth} *}(*<*)next(*>*) + { + obtain x where "A x" sorry %noproof + have B sorry %noproof + } + note `B` + txt_raw {* \end{minipage} *} +(*<*) +qed +(*>*) + +text {* + \bigskip\noindent This illustrates the meaning of Isar context + elements without goals getting in between. +*} + +subsection {* Structured statements \label{sec:framework-stmt} *} + +text {* + The category @{text "statement"} of top-level theorem specifications + is defined as follows: + + \medskip + \begin{tabular}{rcl} + @{text "statement"} & @{text "\"} & @{text "name: props \ \"} \\ + & @{text "|"} & @{text "context\<^sup>* conclusion"} \\[0.5ex] + + @{text "context"} & @{text "\"} & @{text "\ vars \ \"} \\ + & @{text "|"} & @{text "\ name: props \ \"} \\ + + @{text "conclusion"} & @{text "\"} & @{text "\ name: props \ \"} \\ + & @{text "|"} & @{text "\ vars \ \ \ name: props \ \"} \\ + & & \quad @{text "\ \"} \\ + \end{tabular} + + \medskip\noindent A simple @{text "statement"} consists of named + propositions. The full form admits local context elements followed + by the actual conclusions, such as ``@{keyword "fixes"}~@{text + x}~@{keyword "assumes"}~@{text "A x"}~@{keyword "shows"}~@{text "B + x"}''. The final result emerges as a Pure rule after discharging + the context: @{prop "\x. A x \ B x"}. + + The @{keyword "obtains"} variant is another abbreviation defined + below; unlike @{command obtain} (cf.\ + \secref{sec:framework-context}) there may be several ``cases'' + separated by ``@{text "\"}'', each consisting of several + parameters (@{text "vars"}) and several premises (@{text "props"}). + This specifies multi-branch elimination rules. + + \medskip + \begin{tabular}{l} + @{text "\ \<^vec>x \ \<^vec>A \<^vec>x \ \ \"} \\[0.5ex] + \quad @{text "\ thesis"} \\ + \quad @{text "\ [intro]: \\<^vec>x. \<^vec>A \<^vec>x \ thesis \ \"} \\ + \quad @{text "\ thesis"} \\ + \end{tabular} + \medskip + + Presenting structured statements in such an ``open'' format usually + simplifies the subsequent proof, because the outer structure of the + problem is already laid out directly. E.g.\ consider the following + canonical patterns for @{text "\"} and @{text "\"}, + respectively: +*} + +text_raw {*\begin{minipage}{0.5\textwidth}*} + +theorem + fixes x and y + assumes "A x" and "B y" + shows "C x y" +proof - + from `A x` and `B y` + show "C x y" sorry %noproof +qed + +text_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*} + +theorem + obtains x and y + where "A x" and "B y" +proof - + have "A a" and "B b" sorry %noproof + then show thesis .. +qed + +text_raw {*\end{minipage}*} + +text {* + \medskip\noindent Here local facts \isacharbackquoteopen@{text "A + x"}\isacharbackquoteclose\ and \isacharbackquoteopen@{text "B + y"}\isacharbackquoteclose\ are referenced immediately; there is no + need to decompose the logical rule structure again. In the second + proof the final ``@{command then}~@{command show}~@{text + thesis}~@{command ".."}'' involves the local rule case @{text "\x + y. A x \ B y \ thesis"} for the particular instance of terms @{text + "a"} and @{text "b"} produced in the body. +*} + + +subsection {* Structured proof refinement \label{sec:framework-subproof} *} + +text {* + By breaking up the grammar for the Isar proof language, we may + understand a proof text as a linear sequence of individual proof + commands. These are interpreted as transitions of the Isar virtual + machine (Isar/VM), which operates on a block-structured + configuration in single steps. This allows users to write proof + texts in an incremental manner, and inspect intermediate + configurations for debugging. + + The basic idea is analogous to evaluating algebraic expressions on a + stack machine: @{text "(a + b) \ c"} then corresponds to a sequence + of single transitions for each symbol @{text "(, a, +, b, ), \, c"}. + In Isar the algebraic values are facts or goals, and the operations + are inferences. + + \medskip The Isar/VM state maintains a stack of nodes, each node + contains the local proof context, the linguistic mode, and a pending + goal (optional). The mode determines the type of transition that + may be performed next, it essentially alternates between forward and + backward reasoning, with an intermediate stage for chained facts + (see \figref{fig:isar-vm}). + + \begin{figure}[htb] + \begin{center} + \includegraphics[width=0.8\textwidth]{Thy/document/isar-vm} + \end{center} + \caption{Isar/VM modes}\label{fig:isar-vm} + \end{figure} + + For example, in @{text "state"} mode Isar acts like a mathematical + scratch-pad, accepting declarations like @{command fix}, @{command + assume}, and claims like @{command have}, @{command show}. A goal + statement changes the mode to @{text "prove"}, which means that we + may now refine the problem via @{command unfolding} or @{command + proof}. Then we are again in @{text "state"} mode of a proof body, + which may issue @{command show} statements to solve pending + sub-goals. A concluding @{command qed} will return to the original + @{text "state"} mode one level upwards. The subsequent Isar/VM + trace indicates block structure, linguistic mode, goal state, and + inferences: +*} + +text_raw {* \begingroup\footnotesize *} +(*<*)lemma True +proof +(*>*) + txt_raw {* \begin{minipage}[t]{0.18\textwidth} *} + have "A \ B" + proof + assume A + show B + sorry %noproof + qed + txt_raw {* \end{minipage}\quad +\begin{minipage}[t]{0.06\textwidth} +@{text "begin"} \\ +\\ +\\ +@{text "begin"} \\ +@{text "end"} \\ +@{text "end"} \\ +\end{minipage} +\begin{minipage}[t]{0.08\textwidth} +@{text "prove"} \\ +@{text "state"} \\ +@{text "state"} \\ +@{text "prove"} \\ +@{text "state"} \\ +@{text "state"} \\ +\end{minipage}\begin{minipage}[t]{0.35\textwidth} +@{text "(A \ B) \ #(A \ B)"} \\ +@{text "(A \ B) \ #(A \ B)"} \\ +\\ +\\ +@{text "#(A \ B)"} \\ +@{text "A \ B"} \\ +\end{minipage}\begin{minipage}[t]{0.4\textwidth} +@{text "(init)"} \\ +@{text "(resolution impI)"} \\ +\\ +\\ +@{text "(refinement #A \ B)"} \\ +@{text "(finish)"} \\ +\end{minipage} *} +(*<*) +qed +(*>*) +text_raw {* \endgroup *} + +text {* + \noindent Here the @{inference refinement} inference from + \secref{sec:framework-resolution} mediates composition of Isar + sub-proofs nicely. Observe that this principle incorporates some + degree of freedom in proof composition. In particular, the proof + body allows parameters and assumptions to be re-ordered, or commuted + according to Hereditary Harrop Form. Moreover, context elements + that are not used in a sub-proof may be omitted altogether. For + example: +*} + +text_raw {*\begin{minipage}{0.5\textwidth}*} + +(*<*) +lemma True +proof +(*>*) + have "\x y. A x \ B y \ C x y" + proof - + fix x and y + assume "A x" and "B y" + show "C x y" sorry %noproof + qed + +txt_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*} + +(*<*) +next +(*>*) + have "\x y. A x \ B y \ C x y" + proof - + fix x assume "A x" + fix y assume "B y" + show "C x y" sorry %noproof + qed + +txt_raw {*\end{minipage}\\[3ex]\begin{minipage}{0.5\textwidth}*} + +(*<*) +next +(*>*) + have "\x y. A x \ B y \ C x y" + proof - + fix y assume "B y" + fix x assume "A x" + show "C x y" sorry + qed + +txt_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*} +(*<*) +next +(*>*) + have "\x y. A x \ B y \ C x y" + proof - + fix y assume "B y" + fix x + show "C x y" sorry + qed +(*<*) +qed +(*>*) + +text_raw {*\end{minipage}*} + +text {* + \medskip\noindent Such ``peephole optimizations'' of Isar texts are + practically important to improve readability, by rearranging + contexts elements according to the natural flow of reasoning in the + body, while still observing the overall scoping rules. + + \medskip This illustrates the basic idea of structured proof + processing in Isar. The main mechanisms are based on natural + deduction rule composition within the Pure framework. In + particular, there are no direct operations on goal states within the + proof body. Moreover, there is no hidden automated reasoning + involved, just plain unification. +*} + + +subsection {* Calculational reasoning \label{sec:framework-calc} *} + +text {* + The existing Isar infrastructure is sufficiently flexible to support + calculational reasoning (chains of transitivity steps) as derived + concept. The generic proof elements introduced below depend on + rules declared as @{attribute trans} in the context. It is left to + the object-logic to provide a suitable rule collection for mixed + relations of @{text "="}, @{text "<"}, @{text "\"}, @{text "\"}, + @{text "\"} etc. Due to the flexibility of rule composition + (\secref{sec:framework-resolution}), substitution of equals by + equals is covered as well, even substitution of inequalities + involving monotonicity conditions; see also \cite[\S6]{Wenzel-PhD} + and \cite{Bauer-Wenzel:2001}. + + The generic calculational mechanism is based on the observation that + rules such as @{text "trans:"}~@{prop "x = y \ y = z \ x = z"} + proceed from the premises towards the conclusion in a deterministic + fashion. Thus we may reason in forward mode, feeding intermediate + results into rules selected from the context. The course of + reasoning is organized by maintaining a secondary fact called + ``@{fact calculation}'', apart from the primary ``@{fact this}'' + already provided by the Isar primitives. In the definitions below, + @{attribute OF} refers to @{inference resolution} + (\secref{sec:framework-resolution}) with multiple rule arguments, + and @{text "trans"} represents to a suitable rule from the context: + + \begin{matharray}{rcl} + @{command "also"}@{text "\<^sub>0"} & \equiv & @{command "note"}~@{text "calculation = this"} \\ + @{command "also"}@{text "\<^sub>n\<^sub>+\<^sub>1"} & \equiv & @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\[0.5ex] + @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~@{text calculation} \\ + \end{matharray} + + \noindent The start of a calculation is determined implicitly in the + text: here @{command also} sets @{fact calculation} to the current + result; any subsequent occurrence will update @{fact calculation} by + combination with the next result and a transitivity rule. The + calculational sequence is concluded via @{command finally}, where + the final result is exposed for use in a concluding claim. + + Here is a canonical proof pattern, using @{command have} to + establish the intermediate results: +*} + +(*<*) +lemma True +proof +(*>*) + have "a = b" sorry + also have "\ = c" sorry + also have "\ = d" sorry + finally have "a = d" . +(*<*) +qed +(*>*) + +text {* + \noindent The term ``@{text "\"}'' above is a special abbreviation + provided by the Isabelle/Isar syntax layer: it statically refers to + the right-hand side argument of the previous statement given in the + text. Thus it happens to coincide with relevant sub-expressions in + the calculational chain, but the exact correspondence is dependent + on the transitivity rules being involved. + + \medskip Symmetry rules such as @{prop "x = y \ y = x"} are like + transitivities with only one premise. Isar maintains a separate + rule collection declared via the @{attribute sym} attribute, to be + used in fact expressions ``@{text "a [symmetric]"}'', or single-step + proofs ``@{command assume}~@{text "x = y"}~@{command then}~@{command + have}~@{text "y = x"}~@{command ".."}''. +*} + +end \ No newline at end of file diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/HOL_Specific.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/Introduction.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/Outer_Syntax.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/Proof.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/Quick_Reference.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/ROOT.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/Spec.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/Symbols.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/document/First_Order_Logic.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Thy/document/First_Order_Logic.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,1417 @@ +% +\begin{isabellebody}% +\def\isabellecontext{First{\isacharunderscore}Order{\isacharunderscore}Logic}% +% +\isamarkupheader{Example: First-Order Logic% +} +\isamarkuptrue% +% +\isadelimvisible +% +\endisadelimvisible +% +\isatagvisible +\isacommand{theory}\isamarkupfalse% +\ First{\isacharunderscore}Order{\isacharunderscore}Logic\isanewline +\isakeyword{imports}\ Pure\isanewline +\isakeyword{begin}% +\endisatagvisible +{\isafoldvisible}% +% +\isadelimvisible +% +\endisadelimvisible +% +\begin{isamarkuptext}% +\noindent In order to commence a new object-logic within + Isabelle/Pure we introduce abstract syntactic categories \isa{{\isachardoublequote}i{\isachardoublequote}} + for individuals and \isa{{\isachardoublequote}o{\isachardoublequote}} for object-propositions. The latter + is embedded into the language of Pure propositions by means of a + separate judgment.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{typedecl}\isamarkupfalse% +\ i\isanewline +\isacommand{typedecl}\isamarkupfalse% +\ o\isanewline +\isanewline +\isacommand{judgment}\isamarkupfalse% +\isanewline +\ \ Trueprop\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}o\ {\isasymRightarrow}\ prop{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}{\isachardoublequoteopen}{\isacharunderscore}{\isachardoublequoteclose}\ {\isadigit{5}}{\isacharparenright}% +\begin{isamarkuptext}% +\noindent Note that the object-logic judgement is implicit in the + syntax: writing \isa{A} produces \isa{{\isachardoublequote}Trueprop\ A{\isachardoublequote}} internally. + From the Pure perspective this means ``\isa{A} is derivable in the + object-logic''.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Equational reasoning \label{sec:framework-ex-equal}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Equality is axiomatized as a binary predicate on individuals, with + reflexivity as introduction, and substitution as elimination + principle. Note that the latter is particularly convenient in a + framework like Isabelle, because syntactic congruences are + implicitly produced by unification of \isa{{\isachardoublequote}B\ x{\isachardoublequote}} against + expressions containing occurrences of \isa{x}.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{axiomatization}\isamarkupfalse% +\isanewline +\ \ equal\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}i\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ o{\isachardoublequoteclose}\ \ {\isacharparenleft}\isakeyword{infix}\ {\isachardoublequoteopen}{\isacharequal}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\isanewline +\isakeyword{where}\isanewline +\ \ refl\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isakeyword{and}\isanewline +\ \ subst\ {\isacharbrackleft}elim{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isacharequal}\ y\ {\isasymLongrightarrow}\ B\ x\ {\isasymLongrightarrow}\ B\ y{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent Substitution is very powerful, but also hard to control in + full generality. We derive some common symmetry~/ transitivity + schemes of as particular consequences.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{theorem}\isamarkupfalse% +\ sym\ {\isacharbrackleft}sym{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}y\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{with}\isamarkupfalse% +\ {\isacharbackquoteopen}x\ {\isacharequal}\ y{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ forw{\isacharunderscore}subst\ {\isacharbrackleft}trans{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}y\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}B\ x{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}B\ y{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ {\isacharbackquoteopen}y\ {\isacharequal}\ x{\isacharbackquoteclose}\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isacharequal}\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ this\ \isakeyword{and}\ {\isacharbackquoteopen}B\ x{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}B\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ back{\isacharunderscore}subst\ {\isacharbrackleft}trans{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}B\ x{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}B\ y{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ {\isacharbackquoteopen}x\ {\isacharequal}\ y{\isacharbackquoteclose}\ \isakeyword{and}\ {\isacharbackquoteopen}B\ x{\isacharbackquoteclose}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}B\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ trans\ {\isacharbrackleft}trans{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}x\ {\isacharequal}\ y{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}x\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ {\isacharbackquoteopen}y\ {\isacharequal}\ z{\isacharbackquoteclose}\ \isakeyword{and}\ {\isacharbackquoteopen}x\ {\isacharequal}\ y{\isacharbackquoteclose}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isacharequal}\ z{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Basic group theory% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +As an example for equational reasoning we consider some bits of + group theory. The subsequent locale definition postulates group + operations and axioms; we also derive some consequences of this + specification.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{locale}\isamarkupfalse% +\ group\ {\isacharequal}\isanewline +\ \ \isakeyword{fixes}\ prod\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}i\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ i{\isachardoublequoteclose}\ \ {\isacharparenleft}\isakeyword{infix}\ {\isachardoublequoteopen}{\isasymcirc}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline +\ \ \ \ \isakeyword{and}\ inv\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}i\ {\isasymRightarrow}\ i{\isachardoublequoteclose}\ \ {\isacharparenleft}{\isachardoublequoteopen}{\isacharparenleft}{\isacharunderscore}{\isasyminverse}{\isacharparenright}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline +\ \ \ \ \isakeyword{and}\ unit\ {\isacharcolon}{\isacharcolon}\ i\ \ {\isacharparenleft}{\isachardoublequoteopen}{\isadigit{1}}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymcirc}\ y{\isacharparenright}\ {\isasymcirc}\ z\ {\isacharequal}\ x\ {\isasymcirc}\ {\isacharparenleft}y\ {\isasymcirc}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ left{\isacharunderscore}unit{\isacharcolon}\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isasymcirc}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ left{\isacharunderscore}inv{\isacharcolon}\ {\isachardoublequoteopen}x{\isasyminverse}\ {\isasymcirc}\ x\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ right{\isacharunderscore}inv{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymcirc}\ x{\isasyminverse}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymcirc}\ x{\isasyminverse}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymcirc}\ {\isacharparenleft}x\ {\isasymcirc}\ x{\isasyminverse}{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ left{\isacharunderscore}unit\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{1}}\ {\isasymcirc}\ x{\isacharparenright}\ {\isasymcirc}\ x{\isasyminverse}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ assoc\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminverse}{\isacharparenright}{\isasyminverse}\ {\isasymcirc}\ x{\isasyminverse}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ left{\isacharunderscore}inv\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymcirc}\ x\ {\isacharequal}\ {\isacharparenleft}x{\isasyminverse}{\isacharparenright}{\isasyminverse}\ {\isasymcirc}\ {\isacharparenleft}x{\isasyminverse}\ {\isasymcirc}\ x{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ assoc{\isacharparenright}\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x{\isasyminverse}\ {\isasymcirc}\ x\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ left{\isacharunderscore}inv{\isacharparenright}\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharparenleft}x{\isasyminverse}{\isacharparenright}{\isasyminverse}\ {\isasymcirc}\ {\isasymdots}{\isacharparenright}\ {\isasymcirc}\ x{\isasyminverse}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminverse}{\isacharparenright}{\isasyminverse}\ {\isasymcirc}\ {\isacharparenleft}{\isadigit{1}}\ {\isasymcirc}\ x{\isasyminverse}{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ assoc{\isacharparenright}\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isadigit{1}}\ {\isasymcirc}\ x{\isasyminverse}\ {\isacharequal}\ x{\isasyminverse}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ left{\isacharunderscore}unit{\isacharparenright}\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}x{\isasyminverse}{\isacharparenright}{\isasyminverse}\ {\isasymcirc}\ {\isasymdots}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ left{\isacharunderscore}inv{\isacharparenright}\isanewline +\ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymcirc}\ x{\isasyminverse}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ right{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymcirc}\ {\isadigit{1}}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharequal}\ x{\isasyminverse}\ {\isasymcirc}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ left{\isacharunderscore}inv\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymcirc}\ {\isasymdots}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymcirc}\ x{\isasyminverse}{\isacharparenright}\ {\isasymcirc}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ assoc\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymcirc}\ x{\isasyminverse}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ right{\isacharunderscore}inv{\isacharparenright}\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isasymcirc}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ left{\isacharunderscore}unit{\isacharparenright}\isanewline +\ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymcirc}\ {\isadigit{1}}\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Reasoning from basic axioms is often tedious. Our proofs + work by producing various instances of the given rules (potentially + the symmetric form) using the pattern ``\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{eq}~\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isachardoublequote}{\isacharparenleft}rule\ r{\isacharparenright}{\isachardoublequote}}'' and composing the chain of + results via \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}/\hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}}. These steps may + involve any of the transitivity rules declared in + \secref{sec:framework-ex-equal}, namely \isa{trans} in combining + the first two results in \isa{right{\isacharunderscore}inv} and in the final steps of + both proofs, \isa{forw{\isacharunderscore}subst} in the first combination of \isa{right{\isacharunderscore}unit}, and \isa{back{\isacharunderscore}subst} in all other calculational steps. + + Occasional substitutions in calculations are adequate, but should + not be over-emphasized. The other extreme is to compose a chain by + plain transitivity only, with replacements occurring always in + topmost position. For example:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymcirc}\ {\isadigit{1}}\ {\isacharequal}\ x\ {\isasymcirc}\ {\isacharparenleft}x{\isasyminverse}\ {\isasymcirc}\ x{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse% +\ left{\isacharunderscore}inv\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymcirc}\ x{\isasyminverse}{\isacharparenright}\ {\isasymcirc}\ x{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse% +\ assoc\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymcirc}\ x{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse% +\ right{\isacharunderscore}inv\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse% +\ left{\isacharunderscore}unit\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymcirc}\ {\isadigit{1}}\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Here we have re-used the built-in mechanism for unfolding + definitions in order to normalize each equational problem. A more + realistic object-logic would include proper setup for the Simplifier + (\secref{sec:simplifier}), the main automated tool for equational + reasoning in Isabelle. Then ``\hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}}~\isa{left{\isacharunderscore}inv}~\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isachardot}{\isachardot}}}}}'' would become ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isachardoublequote}{\isacharparenleft}simp\ only{\isacharcolon}\ left{\isacharunderscore}inv{\isacharparenright}{\isachardoublequote}}'' etc.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsection{Propositional logic \label{sec:framework-ex-prop}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We axiomatize basic connectives of propositional logic: implication, + disjunction, and conjunction. The associated rules are modeled + after Gentzen's system of Natural Deduction \cite{Gentzen:1935}.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{axiomatization}\isamarkupfalse% +\isanewline +\ \ imp\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}o\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ o{\isachardoublequoteclose}\ \ {\isacharparenleft}\isakeyword{infixr}\ {\isachardoublequoteopen}{\isasymlongrightarrow}{\isachardoublequoteclose}\ {\isadigit{2}}{\isadigit{5}}{\isacharparenright}\ \isakeyword{where}\isanewline +\ \ impI\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}A\ {\isasymLongrightarrow}\ B{\isacharparenright}\ {\isasymLongrightarrow}\ A\ {\isasymlongrightarrow}\ B{\isachardoublequoteclose}\ \isakeyword{and}\isanewline +\ \ impD\ {\isacharbrackleft}dest{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}A\ {\isasymlongrightarrow}\ B{\isacharparenright}\ {\isasymLongrightarrow}\ A\ {\isasymLongrightarrow}\ B{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{axiomatization}\isamarkupfalse% +\isanewline +\ \ disj\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}o\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ o{\isachardoublequoteclose}\ \ {\isacharparenleft}\isakeyword{infixr}\ {\isachardoublequoteopen}{\isasymor}{\isachardoublequoteclose}\ {\isadigit{3}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\isanewline +\ \ disjI\isactrlisub {\isadigit{1}}\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymLongrightarrow}\ A\ {\isasymor}\ B{\isachardoublequoteclose}\ \isakeyword{and}\isanewline +\ \ disjI\isactrlisub {\isadigit{2}}\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}B\ {\isasymLongrightarrow}\ A\ {\isasymor}\ B{\isachardoublequoteclose}\ \isakeyword{and}\isanewline +\ \ disjE\ {\isacharbrackleft}elim{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymor}\ B\ {\isasymLongrightarrow}\ {\isacharparenleft}A\ {\isasymLongrightarrow}\ C{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}B\ {\isasymLongrightarrow}\ C{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{axiomatization}\isamarkupfalse% +\isanewline +\ \ conj\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}o\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ o{\isachardoublequoteclose}\ \ {\isacharparenleft}\isakeyword{infixr}\ {\isachardoublequoteopen}{\isasymand}{\isachardoublequoteclose}\ {\isadigit{3}}{\isadigit{5}}{\isacharparenright}\ \isakeyword{where}\isanewline +\ \ conjI\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymLongrightarrow}\ B\ {\isasymLongrightarrow}\ A\ {\isasymand}\ B{\isachardoublequoteclose}\ \isakeyword{and}\isanewline +\ \ conjD\isactrlisub {\isadigit{1}}{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ A{\isachardoublequoteclose}\ \isakeyword{and}\isanewline +\ \ conjD\isactrlisub {\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ B{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent The conjunctive destructions have the disadvantage that + decomposing \isa{{\isachardoublequote}A\ {\isasymand}\ B{\isachardoublequote}} involves an immediate decision which + component should be projected. The more convenient simultaneous + elimination \isa{{\isachardoublequote}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ {\isacharparenleft}A\ {\isasymLongrightarrow}\ B\ {\isasymLongrightarrow}\ C{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isachardoublequote}} can be derived as + follows:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{theorem}\isamarkupfalse% +\ conjE\ {\isacharbrackleft}elim{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{obtains}\ A\ \isakeyword{and}\ B\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ {\isacharbackquoteopen}A\ {\isasymand}\ B{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ A\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ conjD\isactrlisub {\isadigit{1}}{\isacharparenright}\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ {\isacharbackquoteopen}A\ {\isasymand}\ B{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ B\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ conjD\isactrlisub {\isadigit{2}}{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Here is an example of swapping conjuncts with a single + intermediate elimination step:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ B\ \isakeyword{and}\ A\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}B\ {\isasymand}\ A{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Note that the analogous elimination rule for disjunction + ``\isa{{\isachardoublequote}{\isasymASSUMES}\ A\ {\isasymor}\ B\ {\isasymOBTAINS}\ A\ {\isasymBBAR}\ B{\isachardoublequote}}'' coincides with + the original axiomatization of \isa{disjE}. + + \medskip We continue propositional logic by introducing absurdity + with its characteristic elimination. Plain truth may then be + defined as a proposition that is trivially true.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{axiomatization}\isamarkupfalse% +\isanewline +\ \ false\ {\isacharcolon}{\isacharcolon}\ o\ \ {\isacharparenleft}{\isachardoublequoteopen}{\isasymbottom}{\isachardoublequoteclose}{\isacharparenright}\ \isakeyword{where}\isanewline +\ \ falseE\ {\isacharbrackleft}elim{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymbottom}\ {\isasymLongrightarrow}\ A{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ true\ {\isacharcolon}{\isacharcolon}\ o\ \ {\isacharparenleft}{\isachardoublequoteopen}{\isasymtop}{\isachardoublequoteclose}{\isacharparenright}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}{\isasymtop}\ {\isasymequiv}\ {\isasymbottom}\ {\isasymlongrightarrow}\ {\isasymbottom}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ trueI\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isasymtop}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ true{\isacharunderscore}def\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\medskip\noindent Now negation represents an implication towards + absurdity:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ not\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}o\ {\isasymRightarrow}\ o{\isachardoublequoteclose}\ \ {\isacharparenleft}{\isachardoublequoteopen}{\isasymnot}\ {\isacharunderscore}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{4}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{4}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}{\isasymnot}\ A\ {\isasymequiv}\ A\ {\isasymlongrightarrow}\ {\isasymbottom}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ notI\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}A\ {\isasymLongrightarrow}\ {\isasymbottom}{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isasymnot}\ A{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ not{\isacharunderscore}def\isanewline +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ A\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isasymbottom}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ {\isacharbackquoteopen}A\ {\isasymLongrightarrow}\ {\isasymbottom}{\isacharbackquoteclose}{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ notE\ {\isacharbrackleft}elim{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymnot}\ A{\isachardoublequoteclose}\ \isakeyword{and}\ A\isanewline +\ \ \isakeyword{shows}\ B\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ {\isacharbackquoteopen}{\isasymnot}\ A{\isacharbackquoteclose}\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymlongrightarrow}\ {\isasymbottom}{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse% +\ not{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ {\isacharbackquoteopen}A\ {\isasymlongrightarrow}\ {\isasymbottom}{\isacharbackquoteclose}\ \isakeyword{and}\ {\isacharbackquoteopen}A{\isacharbackquoteclose}\ \isacommand{have}\isamarkupfalse% +\ {\isasymbottom}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ B\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Classical logic% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Subsequently we state the principle of classical contradiction as a + local assumption. Thus we refrain from forcing the object-logic + into the classical perspective. Within that context, we may derive + well-known consequences of the classical principle.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{locale}\isamarkupfalse% +\ classical\ {\isacharequal}\isanewline +\ \ \isakeyword{assumes}\ classical{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isasymnot}\ C\ {\isasymLongrightarrow}\ C{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isachardoublequoteclose}\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ double{\isacharunderscore}negation{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymnot}\ {\isasymnot}\ C{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ C\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ classical{\isacharparenright}\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymnot}\ C{\isachardoublequoteclose}\isanewline +\ \ \isacommand{with}\isamarkupfalse% +\ {\isacharbackquoteopen}{\isasymnot}\ {\isasymnot}\ C{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ C\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ tertium{\isacharunderscore}non{\isacharunderscore}datur{\isacharcolon}\ {\isachardoublequoteopen}C\ {\isasymor}\ {\isasymnot}\ C{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ double{\isacharunderscore}negation{\isacharparenright}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymnot}\ {\isasymnot}\ {\isacharparenleft}C\ {\isasymor}\ {\isasymnot}\ C{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymnot}\ {\isacharparenleft}C\ {\isasymor}\ {\isasymnot}\ C{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymnot}\ C{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse% +\ C\ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}C\ {\isasymor}\ {\isasymnot}\ C{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{with}\isamarkupfalse% +\ {\isacharbackquoteopen}{\isasymnot}\ {\isacharparenleft}C\ {\isasymor}\ {\isasymnot}\ C{\isacharparenright}{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ {\isasymbottom}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}C\ {\isasymor}\ {\isasymnot}\ C{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ {\isacharbackquoteopen}{\isasymnot}\ {\isacharparenleft}C\ {\isasymor}\ {\isasymnot}\ C{\isacharparenright}{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse% +\ {\isasymbottom}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent These examples illustrate both classical reasoning and + non-trivial propositional proofs in general. All three rules + characterize classical logic independently, but the original rule is + already the most convenient to use, because it leaves the conclusion + unchanged. Note that \isa{{\isachardoublequote}{\isacharparenleft}{\isasymnot}\ C\ {\isasymLongrightarrow}\ C{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isachardoublequote}} fits again into our + format for eliminations, despite the additional twist that the + context refers to the main conclusion. So we may write \isa{classical} as the Isar statement ``\isa{{\isachardoublequote}{\isasymOBTAINS}\ {\isasymnot}\ thesis{\isachardoublequote}}''. + This also explains nicely how classical reasoning really works: + whatever the main \isa{thesis} might be, we may always assume its + negation!% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsection{Quantifiers \label{sec:framework-ex-quant}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Representing quantifiers is easy, thanks to the higher-order nature + of the underlying framework. According to the well-known technique + introduced by Church \cite{church40}, quantifiers are operators on + predicates, which are syntactically represented as \isa{{\isachardoublequote}{\isasymlambda}{\isachardoublequote}}-terms + of type \isa{{\isachardoublequote}i\ {\isasymRightarrow}\ o{\isachardoublequote}}. Binder notation turns \isa{{\isachardoublequote}All\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ B\ x{\isacharparenright}{\isachardoublequote}} into \isa{{\isachardoublequote}{\isasymforall}x{\isachardot}\ B\ x{\isachardoublequote}} etc.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{axiomatization}\isamarkupfalse% +\isanewline +\ \ All\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isasymRightarrow}\ o{\isacharparenright}\ {\isasymRightarrow}\ o{\isachardoublequoteclose}\ \ {\isacharparenleft}\isakeyword{binder}\ {\isachardoublequoteopen}{\isasymforall}{\isachardoublequoteclose}\ {\isadigit{1}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\isanewline +\ \ allI\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isasymAnd}x{\isachardot}\ B\ x{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymforall}x{\isachardot}\ B\ x{\isachardoublequoteclose}\ \isakeyword{and}\isanewline +\ \ allD\ {\isacharbrackleft}dest{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isasymforall}x{\isachardot}\ B\ x{\isacharparenright}\ {\isasymLongrightarrow}\ B\ a{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{axiomatization}\isamarkupfalse% +\isanewline +\ \ Ex\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isasymRightarrow}\ o{\isacharparenright}\ {\isasymRightarrow}\ o{\isachardoublequoteclose}\ \ {\isacharparenleft}\isakeyword{binder}\ {\isachardoublequoteopen}{\isasymexists}{\isachardoublequoteclose}\ {\isadigit{1}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\isanewline +\ \ exI\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}B\ a\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymexists}x{\isachardot}\ B\ x{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{and}\isanewline +\ \ exE\ {\isacharbrackleft}elim{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isasymexists}x{\isachardot}\ B\ x{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ B\ x\ {\isasymLongrightarrow}\ C{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent The statement of \isa{exE} corresponds to ``\isa{{\isachardoublequote}{\isasymASSUMES}\ {\isasymexists}x{\isachardot}\ B\ x\ {\isasymOBTAINS}\ x\ {\isasymWHERE}\ B\ x{\isachardoublequote}}'' in Isar. In the + subsequent example we illustrate quantifier reasoning involving all + four rules:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{theorem}\isamarkupfalse% +\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymexists}x{\isachardot}\ {\isasymforall}y{\isachardot}\ R\ x\ y{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isasymforall}y{\isachardot}\ {\isasymexists}x{\isachardot}\ R\ x\ y{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ \ \ \ % +\isamarkupcmt{\isa{{\isachardoublequote}{\isasymforall}{\isachardoublequote}} introduction% +} +\isanewline +\ \ \isacommand{obtain}\isamarkupfalse% +\ x\ \isakeyword{where}\ {\isachardoublequoteopen}{\isasymforall}y{\isachardot}\ R\ x\ y{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% +\ {\isacharbackquoteopen}{\isasymexists}x{\isachardot}\ {\isasymforall}y{\isachardot}\ R\ x\ y{\isacharbackquoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\ \ \ \ % +\isamarkupcmt{\isa{{\isachardoublequote}{\isasymexists}{\isachardoublequote}} elimination% +} +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ y\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}R\ x\ y{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% +\ {\isacharbackquoteopen}{\isasymforall}y{\isachardot}\ R\ x\ y{\isacharbackquoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\ \ \ \ % +\isamarkupcmt{\isa{{\isachardoublequote}{\isasymforall}{\isachardoublequote}} destruction% +} +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymexists}x{\isachardot}\ R\ x\ y{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\ \ \ \ % +\isamarkupcmt{\isa{{\isachardoublequote}{\isasymexists}{\isachardoublequote}} introduction% +} +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Canonical reasoning patterns% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The main rules of first-order predicate logic from + \secref{sec:framework-ex-prop} and \secref{sec:framework-ex-quant} + can now be summarized as follows, using the native Isar statement + format of \secref{sec:framework-stmt}. + + \medskip + \begin{tabular}{l} + \isa{{\isachardoublequote}impI{\isacharcolon}\ {\isasymASSUMES}\ A\ {\isasymLongrightarrow}\ B\ {\isasymSHOWS}\ A\ {\isasymlongrightarrow}\ B{\isachardoublequote}} \\ + \isa{{\isachardoublequote}impD{\isacharcolon}\ {\isasymASSUMES}\ A\ {\isasymlongrightarrow}\ B\ {\isasymAND}\ A\ {\isasymSHOWS}\ B{\isachardoublequote}} \\[1ex] + + \isa{{\isachardoublequote}disjI\isactrlisub {\isadigit{1}}{\isacharcolon}\ {\isasymASSUMES}\ A\ {\isasymSHOWS}\ A\ {\isasymor}\ B{\isachardoublequote}} \\ + \isa{{\isachardoublequote}disjI\isactrlisub {\isadigit{2}}{\isacharcolon}\ {\isasymASSUMES}\ B\ {\isasymSHOWS}\ A\ {\isasymor}\ B{\isachardoublequote}} \\ + \isa{{\isachardoublequote}disjE{\isacharcolon}\ {\isasymASSUMES}\ A\ {\isasymor}\ B\ {\isasymOBTAINS}\ A\ {\isasymBBAR}\ B{\isachardoublequote}} \\[1ex] + + \isa{{\isachardoublequote}conjI{\isacharcolon}\ {\isasymASSUMES}\ A\ {\isasymAND}\ B\ {\isasymSHOWS}\ A\ {\isasymand}\ B{\isachardoublequote}} \\ + \isa{{\isachardoublequote}conjE{\isacharcolon}\ {\isasymASSUMES}\ A\ {\isasymand}\ B\ {\isasymOBTAINS}\ A\ {\isasymAND}\ B{\isachardoublequote}} \\[1ex] + + \isa{{\isachardoublequote}falseE{\isacharcolon}\ {\isasymASSUMES}\ {\isasymbottom}\ {\isasymSHOWS}\ A{\isachardoublequote}} \\ + \isa{{\isachardoublequote}trueI{\isacharcolon}\ {\isasymSHOWS}\ {\isasymtop}{\isachardoublequote}} \\[1ex] + + \isa{{\isachardoublequote}notI{\isacharcolon}\ {\isasymASSUMES}\ A\ {\isasymLongrightarrow}\ {\isasymbottom}\ {\isasymSHOWS}\ {\isasymnot}\ A{\isachardoublequote}} \\ + \isa{{\isachardoublequote}notE{\isacharcolon}\ {\isasymASSUMES}\ {\isasymnot}\ A\ {\isasymAND}\ A\ {\isasymSHOWS}\ B{\isachardoublequote}} \\[1ex] + + \isa{{\isachardoublequote}allI{\isacharcolon}\ {\isasymASSUMES}\ {\isasymAnd}x{\isachardot}\ B\ x\ {\isasymSHOWS}\ {\isasymforall}x{\isachardot}\ B\ x{\isachardoublequote}} \\ + \isa{{\isachardoublequote}allE{\isacharcolon}\ {\isasymASSUMES}\ {\isasymforall}x{\isachardot}\ B\ x\ {\isasymSHOWS}\ B\ a{\isachardoublequote}} \\[1ex] + + \isa{{\isachardoublequote}exI{\isacharcolon}\ {\isasymASSUMES}\ B\ a\ {\isasymSHOWS}\ {\isasymexists}x{\isachardot}\ B\ x{\isachardoublequote}} \\ + \isa{{\isachardoublequote}exE{\isacharcolon}\ {\isasymASSUMES}\ {\isasymexists}x{\isachardot}\ B\ x\ {\isasymOBTAINS}\ a\ {\isasymWHERE}\ B\ a{\isachardoublequote}} + \end{tabular} + \medskip + + \noindent This essentially provides a declarative reading of Pure + rules as Isar reasoning patterns: the rule statements tells how a + canonical proof outline shall look like. Since the above rules have + already been declared as \hyperlink{attribute.Pure.intro}{\mbox{\isa{intro}}}, \hyperlink{attribute.Pure.elim}{\mbox{\isa{elim}}}, \hyperlink{attribute.Pure.dest}{\mbox{\isa{dest}}} --- each according to its + particular shape --- we can immediately write Isar proof texts as + follows:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymlongrightarrow}\ B{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ A\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ B% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymlongrightarrow}\ B{\isachardoublequoteclose}\ \isakeyword{and}\ A% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ B\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ A% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymor}\ B{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ B% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymor}\ B{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymor}\ B{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ C\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ A\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ C% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ B\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ C% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ A\ \isakeyword{and}\ B% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ A\ \isakeyword{and}\ B\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymbottom}{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ A\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymtop}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymnot}\ A{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ A\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymbottom}{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymnot}\ A{\isachardoublequoteclose}\ \isakeyword{and}\ A% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ B\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymforall}x{\isachardot}\ B\ x{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}B\ x{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymforall}x{\isachardot}\ B\ x{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}B\ a{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymexists}x{\isachardot}\ B\ x{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}B\ a{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymexists}x{\isachardot}\ B\ x{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ a\ \isakeyword{where}\ {\isachardoublequoteopen}B\ a{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\end{minipage} +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\bigskip\noindent Of course, these proofs are merely examples. As + sketched in \secref{sec:framework-subproof}, there is a fair amount + of flexibility in expressing Pure deductions in Isar. Here the user + is asked to express himself adequately, aiming at proof texts of + literary quality.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimvisible +% +\endisadelimvisible +% +\isatagvisible +\isacommand{end}\isamarkupfalse% +% +\endisatagvisible +{\isafoldvisible}% +% +\isadelimvisible +% +\endisadelimvisible +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/document/Framework.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Thy/document/Framework.tex Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,1518 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Framework}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Framework\isanewline +\isakeyword{imports}\ Main\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupchapter{The Isabelle/Isar Framework \label{ch:isar-framework}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle/Isar + \cite{Wenzel:1999:TPHOL,Wenzel-PhD,Nipkow-TYPES02,Wenzel-Paulson:2006,Wenzel:2006:Festschrift} + is intended as a generic framework for developing formal + mathematical documents with full proof checking. Definitions and + proofs are organized as theories. An assembly of theory sources may + be presented as a printed document; see also + \chref{ch:document-prep}. + + The main objective of Isar is the design of a human-readable + structured proof language, which is called the ``primary proof + format'' in Isar terminology. Such a primary proof language is + somewhere in the middle between the extremes of primitive proof + objects and actual natural language. In this respect, Isar is a bit + more formalistic than Mizar + \cite{Trybulec:1993:MizarFeatures,Rudnicki:1992:MizarOverview,Wiedijk:1999:Mizar}, + using logical symbols for certain reasoning schemes where Mizar + would prefer English words; see \cite{Wenzel-Wiedijk:2002} for + further comparisons of these systems. + + So Isar challenges the traditional way of recording informal proofs + in mathematical prose, as well as the common tendency to see fully + formal proofs directly as objects of some logical calculus (e.g.\ + \isa{{\isachardoublequote}{\isasymlambda}{\isachardoublequote}}-terms in a version of type theory). In fact, Isar is + better understood as an interpreter of a simple block-structured + language for describing the data flow of local facts and goals, + interspersed with occasional invocations of proof methods. + Everything is reduced to logical inferences internally, but these + steps are somewhat marginal compared to the overall bookkeeping of + the interpretation process. Thanks to careful design of the syntax + and semantics of Isar language elements, a formal record of Isar + instructions may later appear as an intelligible text to the + attentive reader. + + The Isar proof language has emerged from careful analysis of some + inherent virtues of the existing logical framework of Isabelle/Pure + \cite{paulson-found,paulson700}, notably composition of higher-order + natural deduction rules, which is a generalization of Gentzen's + original calculus \cite{Gentzen:1935}. The approach of generic + inference systems in Pure is continued by Isar towards actual proof + texts. + + Concrete applications require another intermediate layer: an + object-logic. Isabelle/HOL \cite{isa-tutorial} (simply-typed + set-theory) is being used most of the time; Isabelle/ZF + \cite{isabelle-ZF} is less extensively developed, although it would + probably fit better for classical mathematics. + + \medskip In order to illustrate natural deduction in Isar, we shall + refer to the background theory and library of Isabelle/HOL. This + includes common notions of predicate logic, naive set-theory etc.\ + using fairly standard mathematical notation. From the perspective + of generic natural deduction there is nothing special about the + logical connectives of HOL (\isa{{\isachardoublequote}{\isasymand}{\isachardoublequote}}, \isa{{\isachardoublequote}{\isasymor}{\isachardoublequote}}, \isa{{\isachardoublequote}{\isasymforall}{\isachardoublequote}}, + \isa{{\isachardoublequote}{\isasymexists}{\isachardoublequote}}, etc.), only the resulting reasoning principles are + relevant to the user. There are similar rules available for + set-theory operators (\isa{{\isachardoublequote}{\isasyminter}{\isachardoublequote}}, \isa{{\isachardoublequote}{\isasymunion}{\isachardoublequote}}, \isa{{\isachardoublequote}{\isasymInter}{\isachardoublequote}}, \isa{{\isachardoublequote}{\isasymUnion}{\isachardoublequote}}, etc.), or any other theory developed in the library (lattice + theory, topology etc.). + + Subsequently we briefly review fragments of Isar proof texts + corresponding directly to such general deduction schemes. The + examples shall refer to set-theory, to minimize the danger of + understanding connectives of predicate logic as something special. + + \medskip The following deduction performs \isa{{\isachardoublequote}{\isasyminter}{\isachardoublequote}}-introduction, + working forwards from assumptions towards the conclusion. We give + both the Isar text, and depict the primitive rule involved, as + determined by unification of the problem against rules that are + declared in the library context.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\medskip\begin{minipage}{0.6\textwidth} +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ A{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}x\ {\isasymin}\ B{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ A\ {\isasyminter}\ B{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\end{minipage}\begin{minipage}{0.4\textwidth} +% +\begin{isamarkuptext}% +\infer{\isa{{\isachardoublequote}x\ {\isasymin}\ A\ {\isasyminter}\ B{\isachardoublequote}}}{\isa{{\isachardoublequote}x\ {\isasymin}\ A{\isachardoublequote}} & \isa{{\isachardoublequote}x\ {\isasymin}\ B{\isachardoublequote}}}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\end{minipage} +% +\begin{isamarkuptext}% +\medskip\noindent Note that \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} augments the proof + context, \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} indicates that the current fact shall be + used in the next step, and \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}} states an intermediate + goal. The two dots ``\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isachardot}{\isachardot}}}}}'' refer to a complete proof of + this claim, using the indicated facts and a canonical rule from the + context. We could have been more explicit here by spelling out the + final proof step via the \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}} command:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ A{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}x\ {\isasymin}\ B{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ A\ {\isasyminter}\ B{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ IntI{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent The format of the \isa{{\isachardoublequote}{\isasyminter}{\isachardoublequote}}-introduction rule represents + the most basic inference, which proceeds from given premises to a + conclusion, without any nested proof context involved. + + The next example performs backwards introduction on \isa{{\isachardoublequote}{\isasymInter}{\isasymA}{\isachardoublequote}}, + the intersection of all sets within a given set. This requires a + nested proof of set membership within a local context, where \isa{A} is an arbitrary-but-fixed member of the collection:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\medskip\begin{minipage}{0.6\textwidth} +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isasymInter}{\isasymA}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ A\isanewline +\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymin}\ {\isasymA}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ A{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\end{minipage}\begin{minipage}{0.4\textwidth} +% +\begin{isamarkuptext}% +\infer{\isa{{\isachardoublequote}x\ {\isasymin}\ {\isasymInter}{\isasymA}{\isachardoublequote}}}{\infer*{\isa{{\isachardoublequote}x\ {\isasymin}\ A{\isachardoublequote}}}{\isa{{\isachardoublequote}{\isacharbrackleft}A{\isacharbrackright}{\isacharbrackleft}A\ {\isasymin}\ {\isasymA}{\isacharbrackright}{\isachardoublequote}}}}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\end{minipage} +% +\begin{isamarkuptext}% +\medskip\noindent This Isar reasoning pattern again refers to the + primitive rule depicted above. The system determines it in the + ``\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}'' step, which could have been spelt out more + explicitly as ``\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isachardoublequote}{\isacharparenleft}rule\ InterI{\isacharparenright}{\isachardoublequote}}''. Note + that the rule involves both a local parameter \isa{{\isachardoublequote}A{\isachardoublequote}} and an + assumption \isa{{\isachardoublequote}A\ {\isasymin}\ {\isasymA}{\isachardoublequote}} in the nested reasoning. This kind of + compound rule typically demands a genuine sub-proof in Isar, working + backwards rather than forwards as seen before. In the proof body we + encounter the \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}-\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}-\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} + outline of nested sub-proofs that is typical for Isar. The final + \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} is like \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}} followed by an additional + refinement of the enclosing claim, using the rule derived from the + proof body. + + \medskip The next example involves \isa{{\isachardoublequote}{\isasymUnion}{\isasymA}{\isachardoublequote}}, which can be + characterized as the set of all \isa{{\isachardoublequote}x{\isachardoublequote}} such that \isa{{\isachardoublequote}{\isasymexists}A{\isachardot}\ x\ {\isasymin}\ A\ {\isasymand}\ A\ {\isasymin}\ {\isasymA}{\isachardoublequote}}. The elimination rule for \isa{{\isachardoublequote}x\ {\isasymin}\ {\isasymUnion}{\isasymA}{\isachardoublequote}} does + not mention \isa{{\isachardoublequote}{\isasymexists}{\isachardoublequote}} and \isa{{\isachardoublequote}{\isasymand}{\isachardoublequote}} at all, but admits to obtain + directly a local \isa{{\isachardoublequote}A{\isachardoublequote}} such that \isa{{\isachardoublequote}x\ {\isasymin}\ A{\isachardoublequote}} and \isa{{\isachardoublequote}A\ {\isasymin}\ {\isasymA}{\isachardoublequote}} hold. This corresponds to the following Isar proof and + inference rule, respectively:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\medskip\begin{minipage}{0.6\textwidth} +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isasymUnion}{\isasymA}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ C\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ A\isanewline +\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ A{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}A\ {\isasymin}\ {\isasymA}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ C% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\end{minipage}\begin{minipage}{0.4\textwidth} +% +\begin{isamarkuptext}% +\infer{\isa{{\isachardoublequote}C{\isachardoublequote}}}{\isa{{\isachardoublequote}x\ {\isasymin}\ {\isasymUnion}{\isasymA}{\isachardoublequote}} & \infer*{\isa{{\isachardoublequote}C{\isachardoublequote}}~}{\isa{{\isachardoublequote}{\isacharbrackleft}A{\isacharbrackright}{\isacharbrackleft}x\ {\isasymin}\ A{\isacharcomma}\ A\ {\isasymin}\ {\isasymA}{\isacharbrackright}{\isachardoublequote}}}}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\end{minipage} +% +\begin{isamarkuptext}% +\medskip\noindent Although the Isar proof follows the natural + deduction rule closely, the text reads not as natural as + anticipated. There is a double occurrence of an arbitrary + conclusion \isa{{\isachardoublequote}C{\isachardoublequote}}, which represents the final result, but is + irrelevant for now. This issue arises for any elimination rule + involving local parameters. Isar provides the derived language + element \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}, which is able to perform the same + elimination proof more conveniently:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isasymUnion}{\isasymA}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ A\ \isakeyword{where}\ {\isachardoublequoteopen}x\ {\isasymin}\ A{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}A\ {\isasymin}\ {\isasymA}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Here we avoid to mention the final conclusion \isa{{\isachardoublequote}C{\isachardoublequote}} + and return to plain forward reasoning. The rule involved in the + ``\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isachardot}{\isachardot}}}}}'' proof is the same as before.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{The Pure framework \label{sec:framework-pure}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The Pure logic \cite{paulson-found,paulson700} is an intuitionistic + fragment of higher-order logic \cite{church40}. In type-theoretic + parlance, there are three levels of \isa{{\isachardoublequote}{\isasymlambda}{\isachardoublequote}}-calculus with + corresponding arrows \isa{{\isachardoublequote}{\isasymRightarrow}{\isachardoublequote}}/\isa{{\isachardoublequote}{\isasymAnd}{\isachardoublequote}}/\isa{{\isachardoublequote}{\isasymLongrightarrow}{\isachardoublequote}}: + + \medskip + \begin{tabular}{ll} + \isa{{\isachardoublequote}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymbeta}{\isachardoublequote}} & syntactic function space (terms depending on terms) \\ + \isa{{\isachardoublequote}{\isasymAnd}x{\isachardot}\ B{\isacharparenleft}x{\isacharparenright}{\isachardoublequote}} & universal quantification (proofs depending on terms) \\ + \isa{{\isachardoublequote}A\ {\isasymLongrightarrow}\ B{\isachardoublequote}} & implication (proofs depending on proofs) \\ + \end{tabular} + \medskip + + \noindent Here only the types of syntactic terms, and the + propositions of proof terms have been shown. The \isa{{\isachardoublequote}{\isasymlambda}{\isachardoublequote}}-structure of proofs can be recorded as an optional feature of + the Pure inference kernel \cite{Berghofer-Nipkow:2000:TPHOL}, but + the formal system can never depend on them due to \emph{proof + irrelevance}. + + On top of this most primitive layer of proofs, Pure implements a + generic calculus for nested natural deduction rules, similar to + \cite{Schroeder-Heister:1984}. Here object-logic inferences are + internalized as formulae over \isa{{\isachardoublequote}{\isasymAnd}{\isachardoublequote}} and \isa{{\isachardoublequote}{\isasymLongrightarrow}{\isachardoublequote}}. + Combining such rule statements may involve higher-order unification + \cite{paulson-natural}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Primitive inferences% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Term syntax provides explicit notation for abstraction \isa{{\isachardoublequote}{\isasymlambda}x\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}{\isachardot}\ b{\isacharparenleft}x{\isacharparenright}{\isachardoublequote}} and application \isa{{\isachardoublequote}b\ a{\isachardoublequote}}, while types are usually + implicit thanks to type-inference; terms of type \isa{{\isachardoublequote}prop{\isachardoublequote}} are + called propositions. Logical statements are composed via \isa{{\isachardoublequote}{\isasymAnd}x\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}{\isachardot}\ B{\isacharparenleft}x{\isacharparenright}{\isachardoublequote}} and \isa{{\isachardoublequote}A\ {\isasymLongrightarrow}\ B{\isachardoublequote}}. Primitive reasoning operates on + judgments of the form \isa{{\isachardoublequote}{\isasymGamma}\ {\isasymturnstile}\ {\isasymphi}{\isachardoublequote}}, with standard introduction + and elimination rules for \isa{{\isachardoublequote}{\isasymAnd}{\isachardoublequote}} and \isa{{\isachardoublequote}{\isasymLongrightarrow}{\isachardoublequote}} that refer to + fixed parameters \isa{{\isachardoublequote}x\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ x\isactrlisub m{\isachardoublequote}} and hypotheses + \isa{{\isachardoublequote}A\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ A\isactrlisub n{\isachardoublequote}} from the context \isa{{\isachardoublequote}{\isasymGamma}{\isachardoublequote}}; + the corresponding proof terms are left implicit. The subsequent + inference rules define \isa{{\isachardoublequote}{\isasymGamma}\ {\isasymturnstile}\ {\isasymphi}{\isachardoublequote}} inductively, relative to a + collection of axioms: + + \[ + \infer{\isa{{\isachardoublequote}{\isasymturnstile}\ A{\isachardoublequote}}}{(\isa{{\isachardoublequote}A{\isachardoublequote}} \text{~axiom})} + \qquad + \infer{\isa{{\isachardoublequote}A\ {\isasymturnstile}\ A{\isachardoublequote}}}{} + \] + + \[ + \infer{\isa{{\isachardoublequote}{\isasymGamma}\ {\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ B{\isacharparenleft}x{\isacharparenright}{\isachardoublequote}}}{\isa{{\isachardoublequote}{\isasymGamma}\ {\isasymturnstile}\ B{\isacharparenleft}x{\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}x\ {\isasymnotin}\ {\isasymGamma}{\isachardoublequote}}} + \qquad + \infer{\isa{{\isachardoublequote}{\isasymGamma}\ {\isasymturnstile}\ B{\isacharparenleft}a{\isacharparenright}{\isachardoublequote}}}{\isa{{\isachardoublequote}{\isasymGamma}\ {\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ B{\isacharparenleft}x{\isacharparenright}{\isachardoublequote}}} + \] + + \[ + \infer{\isa{{\isachardoublequote}{\isasymGamma}\ {\isacharminus}\ A\ {\isasymturnstile}\ A\ {\isasymLongrightarrow}\ B{\isachardoublequote}}}{\isa{{\isachardoublequote}{\isasymGamma}\ {\isasymturnstile}\ B{\isachardoublequote}}} + \qquad + \infer{\isa{{\isachardoublequote}{\isasymGamma}\isactrlsub {\isadigit{1}}\ {\isasymunion}\ {\isasymGamma}\isactrlsub {\isadigit{2}}\ {\isasymturnstile}\ B{\isachardoublequote}}}{\isa{{\isachardoublequote}{\isasymGamma}\isactrlsub {\isadigit{1}}\ {\isasymturnstile}\ A\ {\isasymLongrightarrow}\ B{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymGamma}\isactrlsub {\isadigit{2}}\ {\isasymturnstile}\ A{\isachardoublequote}}} + \] + + Furthermore, Pure provides a built-in equality \isa{{\isachardoublequote}{\isasymequiv}\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ prop{\isachardoublequote}} with axioms for reflexivity, substitution, extensionality, + and \isa{{\isachardoublequote}{\isasymalpha}{\isasymbeta}{\isasymeta}{\isachardoublequote}}-conversion on \isa{{\isachardoublequote}{\isasymlambda}{\isachardoublequote}}-terms. + + \medskip An object-logic introduces another layer on top of Pure, + e.g.\ with types \isa{{\isachardoublequote}i{\isachardoublequote}} for individuals and \isa{{\isachardoublequote}o{\isachardoublequote}} for + propositions, term constants \isa{{\isachardoublequote}Trueprop\ {\isacharcolon}{\isacharcolon}\ o\ {\isasymRightarrow}\ prop{\isachardoublequote}} as + (implicit) derivability judgment and connectives like \isa{{\isachardoublequote}{\isasymand}\ {\isacharcolon}{\isacharcolon}\ o\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ o{\isachardoublequote}} or \isa{{\isachardoublequote}{\isasymforall}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}i\ {\isasymRightarrow}\ o{\isacharparenright}\ {\isasymRightarrow}\ o{\isachardoublequote}}, and axioms for object-level + rules such as \isa{{\isachardoublequote}conjI{\isacharcolon}\ A\ {\isasymLongrightarrow}\ B\ {\isasymLongrightarrow}\ A\ {\isasymand}\ B{\isachardoublequote}} or \isa{{\isachardoublequote}allI{\isacharcolon}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ B\ x{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymforall}x{\isachardot}\ B\ x{\isachardoublequote}}. Derived object rules are represented as theorems of + Pure. After the initial object-logic setup, further axiomatizations + are usually avoided; plain definitions and derived principles are + used exclusively.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Reasoning with rules \label{sec:framework-resolution}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Primitive inferences mostly serve foundational purposes. The main + reasoning mechanisms of Pure operate on nested natural deduction + rules expressed as formulae, using \isa{{\isachardoublequote}{\isasymAnd}{\isachardoublequote}} to bind local + parameters and \isa{{\isachardoublequote}{\isasymLongrightarrow}{\isachardoublequote}} to express entailment. Multiple + parameters and premises are represented by repeating these + connectives in a right-associative manner. + + Since \isa{{\isachardoublequote}{\isasymAnd}{\isachardoublequote}} and \isa{{\isachardoublequote}{\isasymLongrightarrow}{\isachardoublequote}} commute thanks to the theorem + \isa{{\isachardoublequote}{\isacharparenleft}A\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ B\ x{\isacharparenright}{\isacharparenright}\ {\isasymequiv}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ A\ {\isasymLongrightarrow}\ B\ x{\isacharparenright}{\isachardoublequote}}, we may assume w.l.o.g.\ + that rule statements always observe the normal form where + quantifiers are pulled in front of implications at each level of + nesting. This means that any Pure proposition may be presented as a + \emph{Hereditary Harrop Formula} \cite{Miller:1991} which is of the + form \isa{{\isachardoublequote}{\isasymAnd}x\isactrlisub {\isadigit{1}}\ {\isasymdots}\ x\isactrlisub m{\isachardot}\ H\isactrlisub {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isasymdots}\ H\isactrlisub n\ {\isasymLongrightarrow}\ A{\isachardoublequote}} for \isa{{\isachardoublequote}m{\isacharcomma}\ n\ {\isasymge}\ {\isadigit{0}}{\isachardoublequote}}, and \isa{{\isachardoublequote}A{\isachardoublequote}} atomic, and \isa{{\isachardoublequote}H\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ H\isactrlisub n{\isachardoublequote}} being recursively of the same format. + Following the convention that outermost quantifiers are implicit, + Horn clauses \isa{{\isachardoublequote}A\isactrlisub {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isasymdots}\ A\isactrlisub n\ {\isasymLongrightarrow}\ A{\isachardoublequote}} are a special + case of this. + + For example, \isa{{\isachardoublequote}{\isasyminter}{\isachardoublequote}}-introduction rule encountered before is + represented as a Pure theorem as follows: + \[ + \isa{{\isachardoublequote}IntI{\isacharcolon}{\isachardoublequote}}~\isa{{\isachardoublequote}x\ {\isasymin}\ A\ {\isasymLongrightarrow}\ x\ {\isasymin}\ B\ {\isasymLongrightarrow}\ x\ {\isasymin}\ A\ {\isasyminter}\ B{\isachardoublequote}} + \] + + \noindent This is a plain Horn clause, since no further nesting on + the left is involved. The general \isa{{\isachardoublequote}{\isasymInter}{\isachardoublequote}}-introduction + corresponds to a Hereditary Harrop Formula with one additional level + of nesting: + \[ + \isa{{\isachardoublequote}InterI{\isacharcolon}{\isachardoublequote}}~\isa{{\isachardoublequote}{\isacharparenleft}{\isasymAnd}A{\isachardot}\ A\ {\isasymin}\ {\isasymA}\ {\isasymLongrightarrow}\ x\ {\isasymin}\ A{\isacharparenright}\ {\isasymLongrightarrow}\ x\ {\isasymin}\ {\isasymInter}{\isasymA}{\isachardoublequote}} + \] + + \medskip Goals are also represented as rules: \isa{{\isachardoublequote}A\isactrlisub {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isasymdots}\ A\isactrlisub n\ {\isasymLongrightarrow}\ C{\isachardoublequote}} states that the sub-goals \isa{{\isachardoublequote}A\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ A\isactrlisub n{\isachardoublequote}} entail the result \isa{{\isachardoublequote}C{\isachardoublequote}}; for \isa{{\isachardoublequote}n\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequote}} the + goal is finished. To allow \isa{{\isachardoublequote}C{\isachardoublequote}} being a rule statement + itself, we introduce the protective marker \isa{{\isachardoublequote}{\isacharhash}\ {\isacharcolon}{\isacharcolon}\ prop\ {\isasymRightarrow}\ prop{\isachardoublequote}}, which is defined as identity and hidden from the user. We + initialize and finish goal states as follows: + + \[ + \begin{array}{c@ {\qquad}c} + \infer[(\indexdef{}{inference}{init}\hypertarget{inference.init}{\hyperlink{inference.init}{\mbox{\isa{init}}}})]{\isa{{\isachardoublequote}C\ {\isasymLongrightarrow}\ {\isacharhash}C{\isachardoublequote}}}{} & + \infer[(\indexdef{}{inference}{finish}\hypertarget{inference.finish}{\hyperlink{inference.finish}{\mbox{\isa{finish}}}})]{\isa{C}}{\isa{{\isachardoublequote}{\isacharhash}C{\isachardoublequote}}} + \end{array} + \] + + \noindent Goal states are refined in intermediate proof steps until + a finished form is achieved. Here the two main reasoning principles + are \hyperlink{inference.resolution}{\mbox{\isa{resolution}}}, for back-chaining a rule against a + sub-goal (replacing it by zero or more sub-goals), and \hyperlink{inference.assumption}{\mbox{\isa{assumption}}}, for solving a sub-goal (finding a short-circuit with + local assumptions). Below \isa{{\isachardoublequote}\isactrlvec x{\isachardoublequote}} stands for \isa{{\isachardoublequote}x\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ x\isactrlisub n{\isachardoublequote}} (\isa{{\isachardoublequote}n\ {\isasymge}\ {\isadigit{0}}{\isachardoublequote}}). + + \[ + \infer[(\indexdef{}{inference}{resolution}\hypertarget{inference.resolution}{\hyperlink{inference.resolution}{\mbox{\isa{resolution}}}})] + {\isa{{\isachardoublequote}{\isacharparenleft}{\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec H\ \isactrlvec x\ {\isasymLongrightarrow}\ \isactrlvec A\ {\isacharparenleft}\isactrlvec a\ \isactrlvec x{\isacharparenright}{\isacharparenright}{\isasymvartheta}\ {\isasymLongrightarrow}\ C{\isasymvartheta}{\isachardoublequote}}} + {\begin{tabular}{rl} + \isa{{\isachardoublequote}rule{\isacharcolon}{\isachardoublequote}} & + \isa{{\isachardoublequote}\isactrlvec A\ \isactrlvec a\ {\isasymLongrightarrow}\ B\ \isactrlvec a{\isachardoublequote}} \\ + \isa{{\isachardoublequote}goal{\isacharcolon}{\isachardoublequote}} & + \isa{{\isachardoublequote}{\isacharparenleft}{\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec H\ \isactrlvec x\ {\isasymLongrightarrow}\ B{\isacharprime}\ \isactrlvec x{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isachardoublequote}} \\ + \isa{{\isachardoublequote}goal\ unifier{\isacharcolon}{\isachardoublequote}} & + \isa{{\isachardoublequote}{\isacharparenleft}{\isasymlambda}\isactrlvec x{\isachardot}\ B\ {\isacharparenleft}\isactrlvec a\ \isactrlvec x{\isacharparenright}{\isacharparenright}{\isasymvartheta}\ {\isacharequal}\ B{\isacharprime}{\isasymvartheta}{\isachardoublequote}} \\ + \end{tabular}} + \] + + \medskip + + \[ + \infer[(\indexdef{}{inference}{assumption}\hypertarget{inference.assumption}{\hyperlink{inference.assumption}{\mbox{\isa{assumption}}}})]{\isa{{\isachardoublequote}C{\isasymvartheta}{\isachardoublequote}}} + {\begin{tabular}{rl} + \isa{{\isachardoublequote}goal{\isacharcolon}{\isachardoublequote}} & + \isa{{\isachardoublequote}{\isacharparenleft}{\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec H\ \isactrlvec x\ {\isasymLongrightarrow}\ A\ \isactrlvec x{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isachardoublequote}} \\ + \isa{{\isachardoublequote}assm\ unifier{\isacharcolon}{\isachardoublequote}} & \isa{{\isachardoublequote}A{\isasymvartheta}\ {\isacharequal}\ H\isactrlsub i{\isasymvartheta}{\isachardoublequote}}~~\text{(for some~\isa{{\isachardoublequote}H\isactrlsub i{\isachardoublequote}})} \\ + \end{tabular}} + \] + + The following trace illustrates goal-oriented reasoning in + Isabelle/Pure: + + {\footnotesize + \medskip + \begin{tabular}{r@ {\quad}l} + \isa{{\isachardoublequote}{\isacharparenleft}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ B\ {\isasymand}\ A{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharhash}{\isacharparenleft}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ B\ {\isasymand}\ A{\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}init{\isacharparenright}{\isachardoublequote}} \\ + \isa{{\isachardoublequote}{\isacharparenleft}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ B{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ A{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharhash}{\isasymdots}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}resolution\ B\ {\isasymLongrightarrow}\ A\ {\isasymLongrightarrow}\ B\ {\isasymand}\ A{\isacharparenright}{\isachardoublequote}} \\ + \isa{{\isachardoublequote}{\isacharparenleft}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ A\ {\isasymand}\ B{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ A{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharhash}{\isasymdots}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}resolution\ A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ B{\isacharparenright}{\isachardoublequote}} \\ + \isa{{\isachardoublequote}{\isacharparenleft}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ A{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharhash}{\isasymdots}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}assumption{\isacharparenright}{\isachardoublequote}} \\ + \isa{{\isachardoublequote}{\isacharparenleft}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ B\ {\isasymand}\ A{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharhash}{\isasymdots}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}resolution\ A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ A{\isacharparenright}{\isachardoublequote}} \\ + \isa{{\isachardoublequote}{\isacharhash}{\isasymdots}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}assumption{\isacharparenright}{\isachardoublequote}} \\ + \isa{{\isachardoublequote}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}finish{\isacharparenright}{\isachardoublequote}} \\ + \end{tabular} + \medskip + } + + Compositions of \hyperlink{inference.assumption}{\mbox{\isa{assumption}}} after \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} occurs quite often, typically in elimination steps. + Traditional Isabelle tactics accommodate this by a combined + \indexdef{}{inference}{elim\_resolution}\hypertarget{inference.elim-resolution}{\hyperlink{inference.elim-resolution}{\mbox{\isa{elim{\isacharunderscore}resolution}}}} principle. In contrast, Isar uses + a slightly more refined combination, where the assumptions to be + closed are marked explicitly, using again the protective marker + \isa{{\isachardoublequote}{\isacharhash}{\isachardoublequote}}: + + \[ + \infer[(\hyperlink{inference.refinement}{\mbox{\isa{refinement}}})] + {\isa{{\isachardoublequote}{\isacharparenleft}{\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec H\ \isactrlvec x\ {\isasymLongrightarrow}\ \isactrlvec G{\isacharprime}\ {\isacharparenleft}\isactrlvec a\ \isactrlvec x{\isacharparenright}{\isacharparenright}{\isasymvartheta}\ {\isasymLongrightarrow}\ C{\isasymvartheta}{\isachardoublequote}}} + {\begin{tabular}{rl} + \isa{{\isachardoublequote}sub{\isasymdash}proof{\isacharcolon}{\isachardoublequote}} & + \isa{{\isachardoublequote}\isactrlvec G\ \isactrlvec a\ {\isasymLongrightarrow}\ B\ \isactrlvec a{\isachardoublequote}} \\ + \isa{{\isachardoublequote}goal{\isacharcolon}{\isachardoublequote}} & + \isa{{\isachardoublequote}{\isacharparenleft}{\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec H\ \isactrlvec x\ {\isasymLongrightarrow}\ B{\isacharprime}\ \isactrlvec x{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isachardoublequote}} \\ + \isa{{\isachardoublequote}goal\ unifier{\isacharcolon}{\isachardoublequote}} & + \isa{{\isachardoublequote}{\isacharparenleft}{\isasymlambda}\isactrlvec x{\isachardot}\ B\ {\isacharparenleft}\isactrlvec a\ \isactrlvec x{\isacharparenright}{\isacharparenright}{\isasymvartheta}\ {\isacharequal}\ B{\isacharprime}{\isasymvartheta}{\isachardoublequote}} \\ + \isa{{\isachardoublequote}assm\ unifiers{\isacharcolon}{\isachardoublequote}} & + \isa{{\isachardoublequote}{\isacharparenleft}{\isasymlambda}\isactrlvec x{\isachardot}\ G\isactrlsub j\ {\isacharparenleft}\isactrlvec a\ \isactrlvec x{\isacharparenright}{\isacharparenright}{\isasymvartheta}\ {\isacharequal}\ {\isacharhash}H\isactrlsub i{\isasymvartheta}{\isachardoublequote}} \\ + & \quad (for each marked \isa{{\isachardoublequote}G\isactrlsub j{\isachardoublequote}} some \isa{{\isachardoublequote}{\isacharhash}H\isactrlsub i{\isachardoublequote}}) \\ + \end{tabular}} + \] + + \noindent Here the \isa{{\isachardoublequote}sub{\isasymdash}proof{\isachardoublequote}} rule stems from the + main \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}-\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}-\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} outline of + Isar (cf.\ \secref{sec:framework-subproof}): each assumption + indicated in the text results in a marked premise \isa{{\isachardoublequote}G{\isachardoublequote}} above. + The marking enforces resolution against one of the sub-goal's + premises. Consequently, \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}-\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}-\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} enables to fit the result of a sub-proof quite robustly into a + pending sub-goal, while maintaining a good measure of flexibility.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{The Isar proof language \label{sec:framework-isar}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Structured proofs are presented as high-level expressions for + composing entities of Pure (propositions, facts, and goals). The + Isar proof language allows to organize reasoning within the + underlying rule calculus of Pure, but Isar is not another logical + calculus! + + Isar is an exercise in sound minimalism. Approximately half of the + language is introduced as primitive, the rest defined as derived + concepts. The following grammar describes the core language + (category \isa{{\isachardoublequote}proof{\isachardoublequote}}), which is embedded into theory + specification elements such as \hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}; see also + \secref{sec:framework-stmt} for the separate category \isa{{\isachardoublequote}statement{\isachardoublequote}}. + + \medskip + \begin{tabular}{rcl} + \isa{{\isachardoublequote}theory{\isasymdash}stmt{\isachardoublequote}} & = & \hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}~\isa{{\isachardoublequote}statement\ proof\ \ {\isacharbar}{\isachardoublequote}}~~\hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}~\isa{{\isachardoublequote}{\isasymdots}\ \ {\isacharbar}\ \ {\isasymdots}{\isachardoublequote}} \\[1ex] + + \isa{{\isachardoublequote}proof{\isachardoublequote}} & = & \isa{{\isachardoublequote}prfx\isactrlsup {\isacharasterisk}{\isachardoublequote}}~\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isachardoublequote}method\isactrlsup {\isacharquery}\ stmt\isactrlsup {\isacharasterisk}{\isachardoublequote}}~\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}~\isa{{\isachardoublequote}method\isactrlsup {\isacharquery}{\isachardoublequote}} \\[1ex] + + \isa{prfx} & = & \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{{\isachardoublequote}facts{\isachardoublequote}} \\ + & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}}~\isa{{\isachardoublequote}facts{\isachardoublequote}} \\ + + \isa{stmt} & = & \hyperlink{command.braceleft}{\mbox{\isa{\isacommand{{\isacharbraceleft}}}}}~\isa{{\isachardoublequote}stmt\isactrlsup {\isacharasterisk}{\isachardoublequote}}~\hyperlink{command.braceright}{\mbox{\isa{\isacommand{{\isacharbraceright}}}}} \\ + & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \hyperlink{command.next}{\mbox{\isa{\isacommand{next}}}} \\ + & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isachardoublequote}name\ {\isacharequal}\ facts{\isachardoublequote}} \\ + & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}}~\isa{{\isachardoublequote}term\ {\isacharequal}\ term{\isachardoublequote}} \\ + & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{{\isachardoublequote}var\isactrlsup {\isacharplus}{\isachardoublequote}} \\ + & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isachardoublequote}{\isasymguillemotleft}inference{\isasymguillemotright}\ name{\isacharcolon}\ props{\isachardoublequote}} \\ + & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharquery}{\isachardoublequote}}~\isa{goal} \\ + \isa{goal} & = & \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isachardoublequote}name{\isacharcolon}\ props\ proof{\isachardoublequote}} \\ + & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}~\isa{{\isachardoublequote}name{\isacharcolon}\ props\ proof{\isachardoublequote}} \\ + \end{tabular} + + \medskip Simultaneous propositions or facts may be separated by the + \hyperlink{keyword.and}{\mbox{\isa{\isakeyword{and}}}} keyword. + + \medskip The syntax for terms and propositions is inherited from + Pure (and the object-logic). A \isa{{\isachardoublequote}pattern{\isachardoublequote}} is a \isa{{\isachardoublequote}term{\isachardoublequote}} with schematic variables, to be bound by higher-order + matching. + + \medskip Facts may be referenced by name or proposition. For + example, the result of ``\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isachardoublequote}a{\isacharcolon}\ A\ {\isasymlangle}proof{\isasymrangle}{\isachardoublequote}}'' + becomes available both as \isa{{\isachardoublequote}a{\isachardoublequote}} and + \isacharbackquoteopen\isa{{\isachardoublequote}A{\isachardoublequote}}\isacharbackquoteclose. Moreover, + fact expressions may involve attributes that modify either the + theorem or the background context. For example, the expression + ``\isa{{\isachardoublequote}a\ {\isacharbrackleft}OF\ b{\isacharbrackright}{\isachardoublequote}}'' refers to the composition of two facts + according to the \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} inference of + \secref{sec:framework-resolution}, while ``\isa{{\isachardoublequote}a\ {\isacharbrackleft}intro{\isacharbrackright}{\isachardoublequote}}'' + declares a fact as introduction rule in the context. + + The special fact called ``\hyperlink{fact.this}{\mbox{\isa{this}}}'' always refers to the last + result, as produced by \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}, \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}, \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}, or \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}. Since \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}} occurs + frequently together with \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} we provide some + abbreviations: + + \medskip + \begin{tabular}{rcl} + \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{a} & \isa{{\isachardoublequote}{\isasymequiv}{\isachardoublequote}} & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{a}~\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} \\ + \hyperlink{command.with}{\mbox{\isa{\isacommand{with}}}}~\isa{a} & \isa{{\isachardoublequote}{\isasymequiv}{\isachardoublequote}} & \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{{\isachardoublequote}a\ {\isasymAND}\ this{\isachardoublequote}} \\ + \end{tabular} + \medskip + + The \isa{{\isachardoublequote}method{\isachardoublequote}} category is essentially a parameter and may be + populated later. Methods use the facts indicated by \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} or \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}, and then operate on the goal state. + Some basic methods are predefined: ``\hyperlink{method.-}{\mbox{\isa{{\isacharminus}}}}'' leaves the goal + unchanged, ``\hyperlink{method.this}{\mbox{\isa{this}}}'' applies the facts as rules to the + goal, ``\hyperlink{method.rule}{\mbox{\isa{rule}}}'' applies the facts to another rule and the + result to the goal (both ``\hyperlink{method.this}{\mbox{\isa{this}}}'' and ``\hyperlink{method.rule}{\mbox{\isa{rule}}}'' + refer to \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} of + \secref{sec:framework-resolution}). The secondary arguments to + ``\hyperlink{method.rule}{\mbox{\isa{rule}}}'' may be specified explicitly as in ``\isa{{\isachardoublequote}{\isacharparenleft}rule\ a{\isacharparenright}{\isachardoublequote}}'', or picked from the context. In the latter case, the system + first tries rules declared as \hyperlink{attribute.Pure.elim}{\mbox{\isa{elim}}} or + \hyperlink{attribute.Pure.dest}{\mbox{\isa{dest}}}, followed by those declared as \hyperlink{attribute.Pure.intro}{\mbox{\isa{intro}}}. + + The default method for \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}} is ``\hyperlink{method.rule}{\mbox{\isa{rule}}}'' + (arguments picked from the context), for \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}} it is + ``\hyperlink{method.-}{\mbox{\isa{{\isacharminus}}}}''. Further abbreviations for terminal proof steps + are ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isachardoublequote}method\isactrlsub {\isadigit{1}}\ method\isactrlsub {\isadigit{2}}{\isachardoublequote}}'' for + ``\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isachardoublequote}method\isactrlsub {\isadigit{1}}{\isachardoublequote}}~\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}~\isa{{\isachardoublequote}method\isactrlsub {\isadigit{2}}{\isachardoublequote}}'', and ``\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isachardot}{\isachardot}}}}}'' for ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\hyperlink{method.rule}{\mbox{\isa{rule}}}, and ``\hyperlink{command.dot}{\mbox{\isa{\isacommand{{\isachardot}}}}}'' for ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\hyperlink{method.this}{\mbox{\isa{this}}}''. The \hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}} element operates + directly on the current facts and goal by applying equalities. + + \medskip Block structure can be indicated explicitly by ``\hyperlink{command.braceleft}{\mbox{\isa{\isacommand{{\isacharbraceleft}}}}}~\isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}}~\hyperlink{command.braceright}{\mbox{\isa{\isacommand{{\isacharbraceright}}}}}'', although the body of a sub-proof + already involves implicit nesting. In any case, \hyperlink{command.next}{\mbox{\isa{\isacommand{next}}}} + jumps into the next section of a block, i.e.\ it acts like closing + an implicit block scope and opening another one; there is no direct + correspondence to subgoals here. + + The remaining elements \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}} and \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} build up + a local context (see \secref{sec:framework-context}), while + \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} refines a pending sub-goal by the rule resulting + from a nested sub-proof (see \secref{sec:framework-subproof}). + Further derived concepts will support calculational reasoning (see + \secref{sec:framework-calc}).% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Context elements \label{sec:framework-context}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In judgments \isa{{\isachardoublequote}{\isasymGamma}\ {\isasymturnstile}\ {\isasymphi}{\isachardoublequote}} of the primitive framework, \isa{{\isachardoublequote}{\isasymGamma}{\isachardoublequote}} + essentially acts like a proof context. Isar elaborates this idea + towards a higher-level notion, with additional information for + type-inference, term abbreviations, local facts, hypotheses etc. + + The element \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{{\isachardoublequote}x\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}{\isachardoublequote}} declares a local + parameter, i.e.\ an arbitrary-but-fixed entity of a given type; in + results exported from the context, \isa{{\isachardoublequote}x{\isachardoublequote}} may become anything. + The \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isachardoublequote}{\isasymguillemotleft}inference{\isasymguillemotright}{\isachardoublequote}} element provides a + general interface to hypotheses: ``\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isachardoublequote}{\isasymguillemotleft}inference{\isasymguillemotright}\ A{\isachardoublequote}}'' produces \isa{{\isachardoublequote}A\ {\isasymturnstile}\ A{\isachardoublequote}} locally, while the + included inference tells how to discharge \isa{A} from results + \isa{{\isachardoublequote}A\ {\isasymturnstile}\ B{\isachardoublequote}} later on. There is no user-syntax for \isa{{\isachardoublequote}{\isasymguillemotleft}inference{\isasymguillemotright}{\isachardoublequote}}, i.e.\ it may only occur internally when derived + commands are defined in ML. + + At the user-level, the default inference for \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} is + \hyperlink{inference.discharge}{\mbox{\isa{discharge}}} as given below. The additional variants + \hyperlink{command.presume}{\mbox{\isa{\isacommand{presume}}}} and \hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}} are defined as follows: + + \medskip + \begin{tabular}{rcl} + \hyperlink{command.presume}{\mbox{\isa{\isacommand{presume}}}}~\isa{A} & \isa{{\isachardoublequote}{\isasymequiv}{\isachardoublequote}} & \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isachardoublequote}{\isasymguillemotleft}weak{\isasymdash}discharge{\isasymguillemotright}\ A{\isachardoublequote}} \\ + \hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}}~\isa{{\isachardoublequote}x\ {\isasymequiv}\ a{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymequiv}{\isachardoublequote}} & \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{x}~\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isachardoublequote}{\isasymguillemotleft}expansion{\isasymguillemotright}\ x\ {\isasymequiv}\ a{\isachardoublequote}} \\ + \end{tabular} + \medskip + + \[ + \infer[(\indexdef{}{inference}{discharge}\hypertarget{inference.discharge}{\hyperlink{inference.discharge}{\mbox{\isa{discharge}}}})]{\isa{{\isachardoublequote}{\isasymstrut}{\isasymGamma}\ {\isacharminus}\ A\ {\isasymturnstile}\ {\isacharhash}A\ {\isasymLongrightarrow}\ B{\isachardoublequote}}}{\isa{{\isachardoublequote}{\isasymstrut}{\isasymGamma}\ {\isasymturnstile}\ B{\isachardoublequote}}} + \] + \[ + \infer[(\indexdef{}{inference}{weak-discharge}\hypertarget{inference.weak-discharge}{\hyperlink{inference.weak-discharge}{\mbox{\isa{weak{\isasymdash}discharge}}}})]{\isa{{\isachardoublequote}{\isasymstrut}{\isasymGamma}\ {\isacharminus}\ A\ {\isasymturnstile}\ A\ {\isasymLongrightarrow}\ B{\isachardoublequote}}}{\isa{{\isachardoublequote}{\isasymstrut}{\isasymGamma}\ {\isasymturnstile}\ B{\isachardoublequote}}} + \] + \[ + \infer[(\indexdef{}{inference}{expansion}\hypertarget{inference.expansion}{\hyperlink{inference.expansion}{\mbox{\isa{expansion}}}})]{\isa{{\isachardoublequote}{\isasymstrut}{\isasymGamma}\ {\isacharminus}\ {\isacharparenleft}x\ {\isasymequiv}\ a{\isacharparenright}\ {\isasymturnstile}\ B\ a{\isachardoublequote}}}{\isa{{\isachardoublequote}{\isasymstrut}{\isasymGamma}\ {\isasymturnstile}\ B\ x{\isachardoublequote}}} + \] + + \medskip Note that \hyperlink{inference.discharge}{\mbox{\isa{discharge}}} and \hyperlink{inference.weak-discharge}{\mbox{\isa{weak{\isasymdash}discharge}}} differ in the marker for \isa{A}, which is + relevant when the result of a \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}-\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}-\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} outline is composed with a pending goal, + cf.\ \secref{sec:framework-subproof}. + + The most interesting derived context element in Isar is \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} \cite[\S5.3]{Wenzel-PhD}, which supports generalized + elimination steps in a purely forward manner. The \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} + command takes a specification of parameters \isa{{\isachardoublequote}\isactrlvec x{\isachardoublequote}} and + assumptions \isa{{\isachardoublequote}\isactrlvec A{\isachardoublequote}} to be added to the context, together + with a proof of a case rule stating that this extension is + conservative (i.e.\ may be removed from closed results later on): + + \medskip + \begin{tabular}{l} + \isa{{\isachardoublequote}{\isasymlangle}facts{\isasymrangle}{\isachardoublequote}}~~\hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}~\isa{{\isachardoublequote}\isactrlvec x\ {\isasymWHERE}\ \isactrlvec A\ \isactrlvec x\ \ {\isasymlangle}proof{\isasymrangle}\ {\isasymequiv}{\isachardoublequote}} \\[0.5ex] + \quad \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isachardoublequote}case{\isacharcolon}\ {\isasymAnd}thesis{\isachardot}\ {\isacharparenleft}{\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec A\ \isactrlvec x\ {\isasymLongrightarrow}\ thesis{\isacharparenright}\ {\isasymLongrightarrow}\ thesis{\isasymrangle}{\isachardoublequote}} \\ + \quad \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\hyperlink{method.-}{\mbox{\isa{{\isacharminus}}}} \\ + \qquad \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{thesis} \\ + \qquad \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isachardoublequote}{\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec A\ \isactrlvec x\ {\isasymLongrightarrow}\ thesis{\isachardoublequote}} \\ + \qquad \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}~\isa{thesis}~\hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{{\isachardoublequote}{\isasymlangle}facts{\isasymrangle}\ {\isasymlangle}proof{\isasymrangle}{\isachardoublequote}} \\ + \quad \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}} \\ + \quad \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{{\isachardoublequote}\isactrlvec x{\isachardoublequote}}~\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isachardoublequote}{\isasymguillemotleft}elimination\ case{\isasymguillemotright}\ \isactrlvec A\ \isactrlvec x{\isachardoublequote}} \\ + \end{tabular} + \medskip + + \[ + \infer[(\hyperlink{inference.elimination}{\mbox{\isa{elimination}}})]{\isa{{\isachardoublequote}{\isasymGamma}\ {\isasymturnstile}\ B{\isachardoublequote}}}{ + \begin{tabular}{rl} + \isa{{\isachardoublequote}case{\isacharcolon}{\isachardoublequote}} & + \isa{{\isachardoublequote}{\isasymGamma}\ {\isasymturnstile}\ {\isasymAnd}thesis{\isachardot}\ {\isacharparenleft}{\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec A\ \isactrlvec x\ {\isasymLongrightarrow}\ thesis{\isacharparenright}\ {\isasymLongrightarrow}\ thesis{\isachardoublequote}} \\[0.2ex] + \isa{{\isachardoublequote}result{\isacharcolon}{\isachardoublequote}} & + \isa{{\isachardoublequote}{\isasymGamma}\ {\isasymunion}\ \isactrlvec A\ \isactrlvec y\ {\isasymturnstile}\ B{\isachardoublequote}} \\[0.2ex] + \end{tabular}} + \] + + \noindent Here the name ``\isa{thesis}'' is a specific convention + for an arbitrary-but-fixed proposition; in the primitive natural + deduction rules shown before we have occasionally used \isa{C}. + The whole statement of ``\hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}~\isa{x}~\hyperlink{keyword.where}{\mbox{\isa{\isakeyword{where}}}}~\isa{{\isachardoublequote}A\ x{\isachardoublequote}}'' may be read as a claim that \isa{{\isachardoublequote}A\ x{\isachardoublequote}} + may be assumed for some arbitrary-but-fixed \isa{{\isachardoublequote}x{\isachardoublequote}}. Also note + that ``\hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}~\isa{{\isachardoublequote}A\ {\isasymAND}\ B{\isachardoublequote}}'' without parameters + is similar to ``\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isachardoublequote}A\ {\isasymAND}\ B{\isachardoublequote}}'', but the + latter involves multiple sub-goals. + + \medskip The subsequent Isar proof texts explain all context + elements introduced above using the formal proof language itself. + After finishing a local proof within a block, we indicate the + exported result via \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{minipage}[t]{0.4\textwidth} +\ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}B\ x{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isacharbraceright}}\isamarkupfalse% +\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isacharbackquoteopen}{\isasymAnd}x{\isachardot}\ B\ x{\isacharbackquoteclose}% +\end{minipage}\quad\begin{minipage}[t]{0.4\textwidth} +\ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ A\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ B% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isacharbraceright}}\isamarkupfalse% +\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isacharbackquoteopen}A\ {\isasymLongrightarrow}\ B{\isacharbackquoteclose}% +\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} +\ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{def}\isamarkupfalse% +\ x\ {\isasymequiv}\ a\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}B\ x{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isacharbraceright}}\isamarkupfalse% +\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isacharbackquoteopen}B\ a{\isacharbackquoteclose}% +\end{minipage}\quad\begin{minipage}[t]{0.4\textwidth} +\ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{obtain}\isamarkupfalse% +\ x\ \isakeyword{where}\ {\isachardoublequoteopen}A\ x{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{have}\isamarkupfalse% +\ B% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isacharbraceright}}\isamarkupfalse% +\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isacharbackquoteopen}B{\isacharbackquoteclose}% +\end{minipage} +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\bigskip\noindent This illustrates the meaning of Isar context + elements without goals getting in between.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Structured statements \label{sec:framework-stmt}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The category \isa{{\isachardoublequote}statement{\isachardoublequote}} of top-level theorem specifications + is defined as follows: + + \medskip + \begin{tabular}{rcl} + \isa{{\isachardoublequote}statement{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymequiv}{\isachardoublequote}} & \isa{{\isachardoublequote}name{\isacharcolon}\ props\ {\isasymAND}\ {\isasymdots}{\isachardoublequote}} \\ + & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}context\isactrlsup {\isacharasterisk}\ conclusion{\isachardoublequote}} \\[0.5ex] + + \isa{{\isachardoublequote}context{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymequiv}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymFIXES}\ vars\ {\isasymAND}\ {\isasymdots}{\isachardoublequote}} \\ + & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymASSUMES}\ name{\isacharcolon}\ props\ {\isasymAND}\ {\isasymdots}{\isachardoublequote}} \\ + + \isa{{\isachardoublequote}conclusion{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymequiv}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymSHOWS}\ name{\isacharcolon}\ props\ {\isasymAND}\ {\isasymdots}{\isachardoublequote}} \\ + & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymOBTAINS}\ vars\ {\isasymAND}\ {\isasymdots}\ {\isasymWHERE}\ name{\isacharcolon}\ props\ {\isasymAND}\ {\isasymdots}{\isachardoublequote}} \\ + & & \quad \isa{{\isachardoublequote}{\isasymBBAR}\ {\isasymdots}{\isachardoublequote}} \\ + \end{tabular} + + \medskip\noindent A simple \isa{{\isachardoublequote}statement{\isachardoublequote}} consists of named + propositions. The full form admits local context elements followed + by the actual conclusions, such as ``\hyperlink{keyword.fixes}{\mbox{\isa{\isakeyword{fixes}}}}~\isa{x}~\hyperlink{keyword.assumes}{\mbox{\isa{\isakeyword{assumes}}}}~\isa{{\isachardoublequote}A\ x{\isachardoublequote}}~\hyperlink{keyword.shows}{\mbox{\isa{\isakeyword{shows}}}}~\isa{{\isachardoublequote}B\ x{\isachardoublequote}}''. The final result emerges as a Pure rule after discharging + the context: \isa{{\isachardoublequote}{\isasymAnd}x{\isachardot}\ A\ x\ {\isasymLongrightarrow}\ B\ x{\isachardoublequote}}. + + The \hyperlink{keyword.obtains}{\mbox{\isa{\isakeyword{obtains}}}} variant is another abbreviation defined + below; unlike \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} (cf.\ + \secref{sec:framework-context}) there may be several ``cases'' + separated by ``\isa{{\isachardoublequote}{\isasymBBAR}{\isachardoublequote}}'', each consisting of several + parameters (\isa{{\isachardoublequote}vars{\isachardoublequote}}) and several premises (\isa{{\isachardoublequote}props{\isachardoublequote}}). + This specifies multi-branch elimination rules. + + \medskip + \begin{tabular}{l} + \isa{{\isachardoublequote}{\isasymOBTAINS}\ \isactrlvec x\ {\isasymWHERE}\ \isactrlvec A\ \isactrlvec x\ \ \ {\isasymBBAR}\ \ \ {\isasymdots}\ \ \ {\isasymequiv}{\isachardoublequote}} \\[0.5ex] + \quad \isa{{\isachardoublequote}{\isasymFIXES}\ thesis{\isachardoublequote}} \\ + \quad \isa{{\isachardoublequote}{\isasymASSUMES}\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isasymAnd}\isactrlvec x{\isachardot}\ \isactrlvec A\ \isactrlvec x\ {\isasymLongrightarrow}\ thesis\ \ {\isasymAND}\ \ {\isasymdots}{\isachardoublequote}} \\ + \quad \isa{{\isachardoublequote}{\isasymSHOWS}\ thesis{\isachardoublequote}} \\ + \end{tabular} + \medskip + + Presenting structured statements in such an ``open'' format usually + simplifies the subsequent proof, because the outer structure of the + problem is already laid out directly. E.g.\ consider the following + canonical patterns for \isa{{\isachardoublequote}{\isasymSHOWS}{\isachardoublequote}} and \isa{{\isachardoublequote}{\isasymOBTAINS}{\isachardoublequote}}, + respectively:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{minipage}{0.5\textwidth} +\isacommand{theorem}\isamarkupfalse% +\isanewline +\ \ \isakeyword{fixes}\ x\ \isakeyword{and}\ y\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}A\ x{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}B\ y{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}C\ x\ y{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ {\isacharbackquoteopen}A\ x{\isacharbackquoteclose}\ \isakeyword{and}\ {\isacharbackquoteopen}B\ y{\isacharbackquoteclose}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}C\ x\ y{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\end{minipage}\begin{minipage}{0.5\textwidth} +\isacommand{theorem}\isamarkupfalse% +\isanewline +\ \ \isakeyword{obtains}\ x\ \isakeyword{and}\ y\isanewline +\ \ \isakeyword{where}\ {\isachardoublequoteopen}A\ x{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}B\ y{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ a{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}B\ b{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\end{minipage} +% +\begin{isamarkuptext}% +\medskip\noindent Here local facts \isacharbackquoteopen\isa{{\isachardoublequote}A\ x{\isachardoublequote}}\isacharbackquoteclose\ and \isacharbackquoteopen\isa{{\isachardoublequote}B\ y{\isachardoublequote}}\isacharbackquoteclose\ are referenced immediately; there is no + need to decompose the logical rule structure again. In the second + proof the final ``\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}~\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}~\isa{thesis}~\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isachardot}{\isachardot}}}}}'' involves the local rule case \isa{{\isachardoublequote}{\isasymAnd}x\ y{\isachardot}\ A\ x\ {\isasymLongrightarrow}\ B\ y\ {\isasymLongrightarrow}\ thesis{\isachardoublequote}} for the particular instance of terms \isa{{\isachardoublequote}a{\isachardoublequote}} and \isa{{\isachardoublequote}b{\isachardoublequote}} produced in the body.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Structured proof refinement \label{sec:framework-subproof}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +By breaking up the grammar for the Isar proof language, we may + understand a proof text as a linear sequence of individual proof + commands. These are interpreted as transitions of the Isar virtual + machine (Isar/VM), which operates on a block-structured + configuration in single steps. This allows users to write proof + texts in an incremental manner, and inspect intermediate + configurations for debugging. + + The basic idea is analogous to evaluating algebraic expressions on a + stack machine: \isa{{\isachardoublequote}{\isacharparenleft}a\ {\isacharplus}\ b{\isacharparenright}\ {\isasymcdot}\ c{\isachardoublequote}} then corresponds to a sequence + of single transitions for each symbol \isa{{\isachardoublequote}{\isacharparenleft}{\isacharcomma}\ a{\isacharcomma}\ {\isacharplus}{\isacharcomma}\ b{\isacharcomma}\ {\isacharparenright}{\isacharcomma}\ {\isasymcdot}{\isacharcomma}\ c{\isachardoublequote}}. + In Isar the algebraic values are facts or goals, and the operations + are inferences. + + \medskip The Isar/VM state maintains a stack of nodes, each node + contains the local proof context, the linguistic mode, and a pending + goal (optional). The mode determines the type of transition that + may be performed next, it essentially alternates between forward and + backward reasoning, with an intermediate stage for chained facts + (see \figref{fig:isar-vm}). + + \begin{figure}[htb] + \begin{center} + \includegraphics[width=0.8\textwidth]{Thy/document/isar-vm} + \end{center} + \caption{Isar/VM modes}\label{fig:isar-vm} + \end{figure} + + For example, in \isa{{\isachardoublequote}state{\isachardoublequote}} mode Isar acts like a mathematical + scratch-pad, accepting declarations like \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}, \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}, and claims like \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}, \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}. A goal + statement changes the mode to \isa{{\isachardoublequote}prove{\isachardoublequote}}, which means that we + may now refine the problem via \hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}} or \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}. Then we are again in \isa{{\isachardoublequote}state{\isachardoublequote}} mode of a proof body, + which may issue \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} statements to solve pending + sub-goals. A concluding \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}} will return to the original + \isa{{\isachardoublequote}state{\isachardoublequote}} mode one level upwards. The subsequent Isar/VM + trace indicates block structure, linguistic mode, goal state, and + inferences:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begingroup\footnotesize +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{minipage}[t]{0.18\textwidth} +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymlongrightarrow}\ B{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ A\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ B% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +% +\isadelimnoproof +\ \ \ \ \ \ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\end{minipage}\quad +\begin{minipage}[t]{0.06\textwidth} +\isa{{\isachardoublequote}begin{\isachardoublequote}} \\ +\\ +\\ +\isa{{\isachardoublequote}begin{\isachardoublequote}} \\ +\isa{{\isachardoublequote}end{\isachardoublequote}} \\ +\isa{{\isachardoublequote}end{\isachardoublequote}} \\ +\end{minipage} +\begin{minipage}[t]{0.08\textwidth} +\isa{{\isachardoublequote}prove{\isachardoublequote}} \\ +\isa{{\isachardoublequote}state{\isachardoublequote}} \\ +\isa{{\isachardoublequote}state{\isachardoublequote}} \\ +\isa{{\isachardoublequote}prove{\isachardoublequote}} \\ +\isa{{\isachardoublequote}state{\isachardoublequote}} \\ +\isa{{\isachardoublequote}state{\isachardoublequote}} \\ +\end{minipage}\begin{minipage}[t]{0.35\textwidth} +\isa{{\isachardoublequote}{\isacharparenleft}A\ {\isasymlongrightarrow}\ B{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharhash}{\isacharparenleft}A\ {\isasymlongrightarrow}\ B{\isacharparenright}{\isachardoublequote}} \\ +\isa{{\isachardoublequote}{\isacharparenleft}A\ {\isasymLongrightarrow}\ B{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharhash}{\isacharparenleft}A\ {\isasymlongrightarrow}\ B{\isacharparenright}{\isachardoublequote}} \\ +\\ +\\ +\isa{{\isachardoublequote}{\isacharhash}{\isacharparenleft}A\ {\isasymlongrightarrow}\ B{\isacharparenright}{\isachardoublequote}} \\ +\isa{{\isachardoublequote}A\ {\isasymlongrightarrow}\ B{\isachardoublequote}} \\ +\end{minipage}\begin{minipage}[t]{0.4\textwidth} +\isa{{\isachardoublequote}{\isacharparenleft}init{\isacharparenright}{\isachardoublequote}} \\ +\isa{{\isachardoublequote}{\isacharparenleft}resolution\ impI{\isacharparenright}{\isachardoublequote}} \\ +\\ +\\ +\isa{{\isachardoublequote}{\isacharparenleft}refinement\ {\isacharhash}A\ {\isasymLongrightarrow}\ B{\isacharparenright}{\isachardoublequote}} \\ +\isa{{\isachardoublequote}{\isacharparenleft}finish{\isacharparenright}{\isachardoublequote}} \\ +\end{minipage} +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\endgroup +% +\begin{isamarkuptext}% +\noindent Here the \hyperlink{inference.refinement}{\mbox{\isa{refinement}}} inference from + \secref{sec:framework-resolution} mediates composition of Isar + sub-proofs nicely. Observe that this principle incorporates some + degree of freedom in proof composition. In particular, the proof + body allows parameters and assumptions to be re-ordered, or commuted + according to Hereditary Harrop Form. Moreover, context elements + that are not used in a sub-proof may be omitted altogether. For + example:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{minipage}{0.5\textwidth} +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymAnd}x\ y{\isachardot}\ A\ x\ {\isasymLongrightarrow}\ B\ y\ {\isasymLongrightarrow}\ C\ x\ y{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\ \isakeyword{and}\ y\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ x{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}B\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}C\ x\ y{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\end{minipage}\begin{minipage}{0.5\textwidth} +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymAnd}x\ y{\isachardot}\ A\ x\ {\isasymLongrightarrow}\ B\ y\ {\isasymLongrightarrow}\ C\ x\ y{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ y\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}B\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}C\ x\ y{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimnoproof +\ % +\endisadelimnoproof +% +\isatagnoproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagnoproof +{\isafoldnoproof}% +% +\isadelimnoproof +\isanewline +% +\endisadelimnoproof +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\end{minipage}\\[3ex]\begin{minipage}{0.5\textwidth} +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymAnd}x\ y{\isachardot}\ A\ x\ {\isasymLongrightarrow}\ B\ y\ {\isasymLongrightarrow}\ C\ x\ y{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ y\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}B\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}C\ x\ y{\isachardoublequoteclose}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\end{minipage}\begin{minipage}{0.5\textwidth} +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymAnd}x\ y{\isachardot}\ A\ x\ {\isasymLongrightarrow}\ B\ y\ {\isasymLongrightarrow}\ C\ x\ y{\isachardoublequoteclose}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ y\ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}B\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}C\ x\ y{\isachardoublequoteclose}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\end{minipage} +% +\begin{isamarkuptext}% +\medskip\noindent Such ``peephole optimizations'' of Isar texts are + practically important to improve readability, by rearranging + contexts elements according to the natural flow of reasoning in the + body, while still observing the overall scoping rules. + + \medskip This illustrates the basic idea of structured proof + processing in Isar. The main mechanisms are based on natural + deduction rule composition within the Pure framework. In + particular, there are no direct operations on goal states within the + proof body. Moreover, there is no hidden automated reasoning + involved, just plain unification.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Calculational reasoning \label{sec:framework-calc}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The existing Isar infrastructure is sufficiently flexible to support + calculational reasoning (chains of transitivity steps) as derived + concept. The generic proof elements introduced below depend on + rules declared as \hyperlink{attribute.trans}{\mbox{\isa{trans}}} in the context. It is left to + the object-logic to provide a suitable rule collection for mixed + relations of \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}}, \isa{{\isachardoublequote}{\isacharless}{\isachardoublequote}}, \isa{{\isachardoublequote}{\isasymle}{\isachardoublequote}}, \isa{{\isachardoublequote}{\isasymsubset}{\isachardoublequote}}, + \isa{{\isachardoublequote}{\isasymsubseteq}{\isachardoublequote}} etc. Due to the flexibility of rule composition + (\secref{sec:framework-resolution}), substitution of equals by + equals is covered as well, even substitution of inequalities + involving monotonicity conditions; see also \cite[\S6]{Wenzel-PhD} + and \cite{Bauer-Wenzel:2001}. + + The generic calculational mechanism is based on the observation that + rules such as \isa{{\isachardoublequote}trans{\isacharcolon}{\isachardoublequote}}~\isa{{\isachardoublequote}x\ {\isacharequal}\ y\ {\isasymLongrightarrow}\ y\ {\isacharequal}\ z\ {\isasymLongrightarrow}\ x\ {\isacharequal}\ z{\isachardoublequote}} + proceed from the premises towards the conclusion in a deterministic + fashion. Thus we may reason in forward mode, feeding intermediate + results into rules selected from the context. The course of + reasoning is organized by maintaining a secondary fact called + ``\hyperlink{fact.calculation}{\mbox{\isa{calculation}}}'', apart from the primary ``\hyperlink{fact.this}{\mbox{\isa{this}}}'' + already provided by the Isar primitives. In the definitions below, + \hyperlink{attribute.OF}{\mbox{\isa{OF}}} refers to \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} + (\secref{sec:framework-resolution}) with multiple rule arguments, + and \isa{{\isachardoublequote}trans{\isachardoublequote}} represents to a suitable rule from the context: + + \begin{matharray}{rcl} + \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}\isa{{\isachardoublequote}\isactrlsub {\isadigit{0}}{\isachardoublequote}} & \equiv & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isachardoublequote}calculation\ {\isacharequal}\ this{\isachardoublequote}} \\ + \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}\isa{{\isachardoublequote}\isactrlsub n\isactrlsub {\isacharplus}\isactrlsub {\isadigit{1}}{\isachardoublequote}} & \equiv & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isachardoublequote}calculation\ {\isacharequal}\ trans\ {\isacharbrackleft}OF\ calculation\ this{\isacharbrackright}{\isachardoublequote}} \\[0.5ex] + \hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}} & \equiv & \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}~\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{calculation} \\ + \end{matharray} + + \noindent The start of a calculation is determined implicitly in the + text: here \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}} sets \hyperlink{fact.calculation}{\mbox{\isa{calculation}}} to the current + result; any subsequent occurrence will update \hyperlink{fact.calculation}{\mbox{\isa{calculation}}} by + combination with the next result and a transitivity rule. The + calculational sequence is concluded via \hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}}, where + the final result is exposed for use in a concluding claim. + + Here is a canonical proof pattern, using \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}} to + establish the intermediate results:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}a\ {\isacharequal}\ b{\isachardoublequoteclose}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ c{\isachardoublequoteclose}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ d{\isachardoublequoteclose}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}a\ {\isacharequal}\ d{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent The term ``\isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}}'' above is a special abbreviation + provided by the Isabelle/Isar syntax layer: it statically refers to + the right-hand side argument of the previous statement given in the + text. Thus it happens to coincide with relevant sub-expressions in + the calculational chain, but the exact correspondence is dependent + on the transitivity rules being involved. + + \medskip Symmetry rules such as \isa{{\isachardoublequote}x\ {\isacharequal}\ y\ {\isasymLongrightarrow}\ y\ {\isacharequal}\ x{\isachardoublequote}} are like + transitivities with only one premise. Isar maintains a separate + rule collection declared via the \hyperlink{attribute.sym}{\mbox{\isa{sym}}} attribute, to be + used in fact expressions ``\isa{{\isachardoublequote}a\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isachardoublequote}}'', or single-step + proofs ``\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isachardoublequote}x\ {\isacharequal}\ y{\isachardoublequote}}~\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}~\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isachardoublequote}y\ {\isacharequal}\ x{\isachardoublequote}}~\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isachardot}{\isachardot}}}}}''.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/document/Inner_Syntax.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/document/Introduction.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/document/Outer_Syntax.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/document/Proof.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/document/Quick_Reference.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/document/Spec.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/document/Symbols.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/document/isar-vm.eps --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Thy/document/isar-vm.eps Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,2694 @@ +%!PS-Adobe-3.0 EPSF-3.0 +%%Creator: inkscape 0.46 +%%Pages: 1 +%%Orientation: Portrait +%%BoundingBox: 0 0 435 173 +%%HiResBoundingBox: 0 0 435 173 +%%EndComments +%%BeginSetup +%%EndSetup +%%Page: 1 1 +0 173 translate +0.8 -0.8 scale +0 0 0 setrgbcolor +[] 0 setdash +1 setlinewidth +0 setlinejoin +0 setlinecap +gsave [1 0 0 1 0 0] concat +gsave [1 0 0 1 -44.641342 -76.87234] concat +gsave [1 0 0 1 70.838012 79.725562] concat +0 0 0 setrgbcolor +[] 0 setdash +0.99921262 setlinewidth +1 setlinejoin +1 setlinecap +newpath +229.77649 131.52507 moveto +265.28729 131.52507 lineto +275.08072 131.52507 282.96496 139.40931 282.96496 149.20274 curveto +282.96496 166.99701 lineto +282.96496 176.79043 275.08072 184.67467 265.28729 184.67467 curveto +229.77649 184.67467 lineto +219.98306 184.67467 212.09882 176.79043 212.09882 166.99701 curveto +212.09882 149.20274 lineto +212.09882 139.40931 219.98306 131.52507 229.77649 131.52507 curveto +closepath +stroke +gsave +0 0 0 setrgbcolor +newpath +231.92252 155.58815 moveto +231.92252 157.8694 lineto +231.5423 157.60899 231.15949 157.41628 230.77408 157.29128 curveto +230.39386 157.16628 229.99803 157.10378 229.58658 157.10378 curveto +228.80532 157.10378 228.19595 157.33295 227.75845 157.79128 curveto +227.32616 158.24441 227.11001 158.87982 227.11002 159.69753 curveto +227.11001 160.51524 227.32616 161.15326 227.75845 161.61159 curveto +228.19595 162.06471 228.80532 162.29128 229.58658 162.29128 curveto +230.02407 162.29128 230.43813 162.22617 230.82877 162.09596 curveto +231.22459 161.96576 231.58917 161.77305 231.92252 161.51784 curveto +231.92252 163.8069 lineto +231.48501 163.96836 231.0397 164.08815 230.58658 164.16628 curveto +230.13866 164.24961 229.68813 164.29127 229.23502 164.29128 curveto +227.65689 164.29127 226.42251 163.88763 225.53189 163.08034 curveto +224.64126 162.26784 224.19595 161.14024 224.19595 159.69753 curveto +224.19595 158.25482 224.64126 157.12982 225.53189 156.32253 curveto +226.42251 155.51003 227.65689 155.10378 229.23502 155.10378 curveto +229.69334 155.10378 230.14386 155.14545 230.58658 155.22878 curveto +231.03449 155.30691 231.4798 155.4267 231.92252 155.58815 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +243.14908 158.73659 moveto +243.14908 164.06471 lineto +240.33658 164.06471 lineto +240.33658 163.19753 lineto +240.33658 160.00221 lineto +240.33657 159.23659 240.31834 158.71055 240.28189 158.42409 curveto +240.25063 158.13764 240.19334 157.9267 240.11002 157.79128 curveto +240.00063 157.60899 239.8522 157.46836 239.6647 157.3694 curveto +239.4772 157.26524 239.26366 157.21316 239.02408 157.21315 curveto +238.44074 157.21316 237.98241 157.43972 237.64908 157.89284 curveto +237.31574 158.34076 237.14907 158.96316 237.14908 159.76003 curveto +237.14908 164.06471 lineto +234.3522 164.06471 lineto +234.3522 151.90846 lineto +237.14908 151.90846 lineto +237.14908 156.59596 lineto +237.57095 156.08555 238.01887 155.71055 238.49283 155.47096 curveto +238.96678 155.22618 239.49022 155.10378 240.06314 155.10378 curveto +241.07355 155.10378 241.83917 155.41368 242.36002 156.03346 curveto +242.88605 156.65326 243.14907 157.5543 243.14908 158.73659 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +249.68033 160.12721 moveto +249.09699 160.12722 248.65689 160.22617 248.36002 160.42409 curveto +248.06835 160.62201 247.92251 160.91367 247.92252 161.29909 curveto +247.92251 161.65326 248.0397 161.9319 248.27408 162.13503 curveto +248.51366 162.33294 248.84439 162.4319 249.26627 162.4319 curveto +249.7923 162.4319 250.23501 162.2444 250.59439 161.8694 curveto +250.95376 161.48919 251.13345 161.01524 251.13345 160.44753 curveto +251.13345 160.12721 lineto +249.68033 160.12721 lineto +253.95377 159.07253 moveto +253.95377 164.06471 lineto +251.13345 164.06471 lineto +251.13345 162.76784 lineto +250.75845 163.29909 250.33657 163.68711 249.86783 163.9319 curveto +249.39907 164.17148 248.82876 164.29127 248.15689 164.29128 curveto +247.25064 164.29127 246.51366 164.02825 245.94595 163.50221 curveto +245.38345 162.97096 245.1022 162.28346 245.1022 161.43971 curveto +245.1022 160.41367 245.45376 159.66107 246.15689 159.1819 curveto +246.86522 158.70274 247.9746 158.46316 249.48502 158.46315 curveto +251.13345 158.46315 lineto +251.13345 158.2444 lineto +251.13345 157.8017 250.95897 157.47878 250.61002 157.27565 curveto +250.26105 157.06732 249.71678 156.96316 248.9772 156.96315 curveto +248.37824 156.96316 247.82095 157.02305 247.30533 157.14284 curveto +246.7897 157.26264 246.31053 157.44232 245.86783 157.6819 curveto +245.86783 155.54909 lineto +246.46678 155.40326 247.06835 155.29389 247.67252 155.22096 curveto +248.27668 155.14285 248.88084 155.10378 249.48502 155.10378 curveto +251.06313 155.10378 252.20115 155.41628 252.89908 156.04128 curveto +253.60219 156.66107 253.95376 157.67149 253.95377 159.07253 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +256.57095 155.31471 moveto +259.36783 155.31471 lineto +259.36783 164.06471 lineto +256.57095 164.06471 lineto +256.57095 155.31471 lineto +256.57095 151.90846 moveto +259.36783 151.90846 lineto +259.36783 154.18971 lineto +256.57095 154.18971 lineto +256.57095 151.90846 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +270.86783 158.73659 moveto +270.86783 164.06471 lineto +268.05533 164.06471 lineto +268.05533 163.19753 lineto +268.05533 159.98659 lineto +268.05532 159.23138 268.03709 158.71055 268.00064 158.42409 curveto +267.96938 158.13764 267.91209 157.9267 267.82877 157.79128 curveto +267.71938 157.60899 267.57095 157.46836 267.38345 157.3694 curveto +267.19595 157.26524 266.98241 157.21316 266.74283 157.21315 curveto +266.15949 157.21316 265.70116 157.43972 265.36783 157.89284 curveto +265.03449 158.34076 264.86782 158.96316 264.86783 159.76003 curveto +264.86783 164.06471 lineto +262.07095 164.06471 lineto +262.07095 155.31471 lineto +264.86783 155.31471 lineto +264.86783 156.59596 lineto +265.2897 156.08555 265.73762 155.71055 266.21158 155.47096 curveto +266.68553 155.22618 267.20897 155.10378 267.78189 155.10378 curveto +268.7923 155.10378 269.55792 155.41368 270.07877 156.03346 curveto +270.6048 156.65326 270.86782 157.5543 270.86783 158.73659 curveto +fill +grestore +grestore +0 0 0 setrgbcolor +[] 0 setdash +0.99921262 setlinewidth +0 setlinejoin +0 setlinecap +newpath +424.72469 236.82544 moveto +356.83209 236.82544 lineto +356.83209 236.82544 lineto +stroke +gsave [-0.39968505 4.8945685e-17 -4.8945685e-17 -0.39968505 356.83209 236.82544] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +0 0 0 setrgbcolor +[] 0 setdash +0.99921268 setlinewidth +0 setlinejoin +0 setlinecap +newpath +282.35183 236.82544 moveto +215.11403 236.82544 lineto +215.11403 236.82544 lineto +stroke +gsave [-0.39968507 4.8945688e-17 -4.8945688e-17 -0.39968507 215.11403 236.82544] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +0 0 0 setrgbcolor +[] 0 setdash +0.99999994 setlinewidth +0 setlinejoin +0 setlinecap +newpath +424.69726 192.5341 moveto +215.13005 192.5341 lineto +stroke +gsave [-0.39999998 4.8984251e-17 -4.8984251e-17 -0.39999998 215.13005 192.5341] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +0 0 0 setrgbcolor +[] 0 setdash +1 setlinewidth +0 setlinejoin +0 setlinecap +newpath +211.98429 148.24276 moveto +422.13162 148.24276 lineto +stroke +gsave [0.4 0 0 0.4 422.13162 148.24276] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +gsave [1 0 0 1 70.866146 78.725567] concat +0 0 0 setrgbcolor +[] 0 setdash +0.99921262 setlinewidth +1 setlinejoin +1 setlinecap +newpath +88.044201 42.942394 moveto +123.555 42.942394 lineto +133.34843 42.942394 141.23267 50.826635 141.23267 60.620064 curveto +141.23267 166.99701 lineto +141.23267 176.79044 133.34843 184.67468 123.555 184.67468 curveto +88.044201 184.67468 lineto +78.250772 184.67468 70.366531 176.79044 70.366531 166.99701 curveto +70.366531 60.620064 lineto +70.366531 50.826635 78.250772 42.942394 88.044201 42.942394 curveto +closepath +stroke +gsave +0 0 0 setrgbcolor +newpath +83.823044 115.35931 moveto +83.823044 119.95306 lineto +81.026169 119.95306 lineto +81.026169 107.87494 lineto +83.823044 107.87494 lineto +83.823044 109.15619 lineto +84.208456 108.64578 84.635539 108.27078 85.104294 108.03119 curveto +85.573038 107.78641 86.1121 107.66401 86.721481 107.664 curveto +87.799598 107.66401 88.685014 108.0937 89.377731 108.95306 curveto +90.070429 109.80724 90.416783 110.9088 90.416794 112.25775 curveto +90.416783 113.60671 90.070429 114.71088 89.377731 115.57025 curveto +88.685014 116.42442 87.799598 116.8515 86.721481 116.8515 curveto +86.1121 116.8515 85.573038 116.73171 85.104294 116.49213 curveto +84.635539 116.24734 84.208456 115.86973 83.823044 115.35931 curveto +85.682419 109.69525 moveto +85.083455 109.69526 84.622518 109.91661 84.299606 110.35931 curveto +83.981894 110.79682 83.82304 111.42963 83.823044 112.25775 curveto +83.82304 113.08588 83.981894 113.7213 84.299606 114.164 curveto +84.622518 114.6015 85.083455 114.82025 85.682419 114.82025 curveto +86.281371 114.82025 86.737099 114.6015 87.049606 114.164 curveto +87.367307 113.7265 87.526161 113.09109 87.526169 112.25775 curveto +87.526161 111.42442 87.367307 110.78901 87.049606 110.3515 curveto +86.737099 109.91401 86.281371 109.69526 85.682419 109.69525 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +98.994919 110.25775 moveto +98.75012 110.14317 98.505328 110.05984 98.260544 110.00775 curveto +98.020954 109.95047 97.778766 109.92182 97.533981 109.92181 curveto +96.815226 109.92182 96.260539 110.15359 95.869919 110.61713 curveto +95.484498 111.07547 95.29179 111.73432 95.291794 112.59369 curveto +95.291794 116.62494 lineto +92.494919 116.62494 lineto +92.494919 107.87494 lineto +95.291794 107.87494 lineto +95.291794 109.31244 lineto +95.651164 108.73953 96.062622 108.32286 96.526169 108.06244 curveto +96.994913 107.79682 97.554808 107.66401 98.205856 107.664 curveto +98.299599 107.66401 98.401162 107.66922 98.510544 107.67963 curveto +98.619911 107.68484 98.778765 107.70047 98.987106 107.7265 curveto +98.994919 110.25775 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +104.56523 109.664 moveto +103.94543 109.66401 103.47148 109.88797 103.14336 110.33588 curveto +102.82044 110.77859 102.65898 111.41922 102.65898 112.25775 curveto +102.65898 113.0963 102.82044 113.73953 103.14336 114.18744 curveto +103.47148 114.63015 103.94543 114.8515 104.56523 114.8515 curveto +105.1746 114.8515 105.64075 114.63015 105.96367 114.18744 curveto +106.28658 113.73953 106.44804 113.0963 106.44804 112.25775 curveto +106.44804 111.41922 106.28658 110.77859 105.96367 110.33588 curveto +105.64075 109.88797 105.1746 109.66401 104.56523 109.664 curveto +104.56523 107.664 moveto +106.07043 107.66401 107.24491 108.07026 108.08867 108.88275 curveto +108.93762 109.69526 109.3621 110.82026 109.36211 112.25775 curveto +109.3621 113.69525 108.93762 114.82025 108.08867 115.63275 curveto +107.24491 116.44525 106.07043 116.8515 104.56523 116.8515 curveto +103.05481 116.8515 101.87252 116.44525 101.01836 115.63275 curveto +100.1694 114.82025 99.744918 113.69525 99.744919 112.25775 curveto +99.744918 110.82026 100.1694 109.69526 101.01836 108.88275 curveto +101.87252 108.07026 103.05481 107.66401 104.56523 107.664 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +110.29961 107.87494 moveto +113.09648 107.87494 lineto +115.27617 113.92181 lineto +117.44804 107.87494 lineto +120.25273 107.87494 lineto +116.80742 116.62494 lineto +113.73711 116.62494 lineto +110.29961 107.87494 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +130.57304 112.2265 moveto +130.57304 113.02338 lineto +124.03398 113.02338 lineto +124.10169 113.67963 124.33866 114.17182 124.74492 114.49994 curveto +125.15116 114.82807 125.71887 114.99213 126.44804 114.99213 curveto +127.03658 114.99213 127.63814 114.90619 128.25273 114.73431 curveto +128.87251 114.55723 129.50793 114.29161 130.15898 113.93744 curveto +130.15898 116.09369 lineto +129.49751 116.34369 128.83606 116.53119 128.17461 116.65619 curveto +127.51314 116.7864 126.85168 116.8515 126.19023 116.8515 curveto +124.60689 116.8515 123.37512 116.45046 122.49492 115.64838 curveto +121.61992 114.84109 121.18242 113.71088 121.18242 112.25775 curveto +121.18242 110.83067 121.61211 109.70828 122.47148 108.89056 curveto +123.33606 108.07286 124.52356 107.66401 126.03398 107.664 curveto +127.40897 107.66401 128.50793 108.07807 129.33086 108.90619 curveto +130.15897 109.73432 130.57303 110.84109 130.57304 112.2265 curveto +127.69804 111.29681 moveto +127.69804 110.76557 127.54179 110.33849 127.22929 110.01556 curveto +126.922 109.68745 126.51835 109.52338 126.01836 109.52338 curveto +125.47668 109.52338 125.03658 109.67703 124.69804 109.98431 curveto +124.3595 110.2864 124.14856 110.7239 124.06523 111.29681 curveto +127.69804 111.29681 lineto +fill +grestore +grestore +0 0 0 setrgbcolor +[] 0 setdash +1 setlinewidth +0 setlinejoin +0 setlinecap +newpath +176.66575 92.035445 moveto +176.66575 118.61025 lineto +stroke +gsave [2.4492127e-17 0.4 -0.4 2.4492127e-17 176.66575 118.61025] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +gsave [0.2378166 0 0 -0.2269133 90.621413 253.06251] concat +0 0 0 setrgbcolor +[] 0 setdash +4.3013706 setlinewidth +1 setlinejoin +1 setlinecap +newpath +208.65508 282.05865 moveto +193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto +43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto +45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto +177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto +stroke +gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +grestore +gsave [1 0 0 1 70.866151 78.725565] concat +0 0 0 setrgbcolor +[] 0 setdash +0.99921262 setlinewidth +1 setlinejoin +1 setlinecap +newpath +371.50879 42.942394 moveto +407.01959 42.942394 lineto +416.81302 42.942394 424.69726 50.826635 424.69726 60.620064 curveto +424.69726 166.99701 lineto +424.69726 176.79044 416.81302 184.67468 407.01959 184.67468 curveto +371.50879 184.67468 lineto +361.71536 184.67468 353.83112 176.79044 353.83112 166.99701 curveto +353.83112 60.620064 lineto +353.83112 50.826635 361.71536 42.942394 371.50879 42.942394 curveto +closepath +stroke +gsave +0 0 0 setrgbcolor +newpath +374.16263 110.83588 moveto +374.16263 112.96088 lineto +373.56366 112.71088 372.98554 112.52338 372.42825 112.39838 curveto +371.87096 112.27338 371.34491 112.21088 370.85013 112.21088 curveto +370.31887 112.21088 369.92304 112.27859 369.66263 112.414 curveto +369.40742 112.54422 369.27981 112.74734 369.27982 113.02338 curveto +369.27981 113.24734 369.37617 113.41922 369.56888 113.539 curveto +369.76679 113.6588 370.11835 113.74734 370.62357 113.80463 curveto +371.11575 113.87494 lineto +372.54804 114.05724 373.51158 114.35671 374.00638 114.77338 curveto +374.50116 115.19005 374.74856 115.84369 374.74857 116.73431 curveto +374.74856 117.66661 374.40481 118.36713 373.71732 118.83588 curveto +373.02981 119.30463 372.00377 119.539 370.63919 119.539 curveto +370.06106 119.539 369.4621 119.49213 368.84232 119.39838 curveto +368.22773 119.30983 367.59492 119.17442 366.94388 118.99213 curveto +366.94388 116.86713 lineto +367.50117 117.13796 368.07148 117.34109 368.65482 117.4765 curveto +369.24335 117.61192 369.83971 117.67963 370.44388 117.67963 curveto +370.99075 117.67963 371.40221 117.60411 371.67825 117.45306 curveto +371.95429 117.30202 372.09231 117.07807 372.09232 116.78119 curveto +372.09231 116.53119 371.99596 116.3463 371.80325 116.2265 curveto +371.61575 116.1015 371.23814 116.00515 370.67044 115.93744 curveto +370.17825 115.87494 lineto +368.93346 115.71869 368.06106 115.42963 367.56107 115.00775 curveto +367.06106 114.58588 366.81106 113.94526 366.81107 113.08588 curveto +366.81106 112.1588 367.12877 111.4713 367.76419 111.02338 curveto +368.3996 110.57547 369.37356 110.35151 370.68607 110.3515 curveto +371.20169 110.35151 371.74335 110.39057 372.31107 110.46869 curveto +372.87877 110.54682 373.49595 110.66922 374.16263 110.83588 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +379.91263 108.07806 moveto +379.91263 110.56244 lineto +382.79544 110.56244 lineto +382.79544 112.56244 lineto +379.91263 112.56244 lineto +379.91263 116.27338 lineto +379.91262 116.67963 379.99335 116.95567 380.15482 117.1015 curveto +380.31627 117.24213 380.63658 117.31244 381.11575 117.31244 curveto +382.55325 117.31244 lineto +382.55325 119.31244 lineto +380.15482 119.31244 lineto +379.05065 119.31244 378.26679 119.08327 377.80325 118.62494 curveto +377.34492 118.1614 377.11575 117.37755 377.11575 116.27338 curveto +377.11575 112.56244 lineto +375.72513 112.56244 lineto +375.72513 110.56244 lineto +377.11575 110.56244 lineto +377.11575 108.07806 lineto +379.91263 108.07806 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +388.43607 115.37494 moveto +387.85273 115.37494 387.41262 115.4739 387.11575 115.67181 curveto +386.82408 115.86973 386.67825 116.1614 386.67825 116.54681 curveto +386.67825 116.90098 386.79544 117.17963 387.02982 117.38275 curveto +387.26939 117.58067 387.60012 117.67963 388.022 117.67963 curveto +388.54804 117.67963 388.99075 117.49213 389.35013 117.11713 curveto +389.7095 116.73692 389.88918 116.26296 389.88919 115.69525 curveto +389.88919 115.37494 lineto +388.43607 115.37494 lineto +392.7095 114.32025 moveto +392.7095 119.31244 lineto +389.88919 119.31244 lineto +389.88919 118.01556 lineto +389.51418 118.54681 389.09231 118.93484 388.62357 119.17963 curveto +388.15481 119.41921 387.5845 119.539 386.91263 119.539 curveto +386.00638 119.539 385.2694 119.27598 384.70169 118.74994 curveto +384.13919 118.21869 383.85794 117.53119 383.85794 116.68744 curveto +383.85794 115.6614 384.2095 114.9088 384.91263 114.42963 curveto +385.62096 113.95047 386.73033 113.71088 388.24075 113.71088 curveto +389.88919 113.71088 lineto +389.88919 113.49213 lineto +389.88918 113.04942 389.7147 112.72651 389.36575 112.52338 curveto +389.01679 112.31505 388.47252 112.21088 387.73294 112.21088 curveto +387.13398 112.21088 386.57669 112.27078 386.06107 112.39056 curveto +385.54544 112.51036 385.06627 112.69005 384.62357 112.92963 curveto +384.62357 110.79681 lineto +385.22252 110.65099 385.82408 110.54161 386.42825 110.46869 curveto +387.03242 110.39057 387.63658 110.35151 388.24075 110.3515 curveto +389.81887 110.35151 390.95689 110.66401 391.65482 111.289 curveto +392.35793 111.9088 392.70949 112.91922 392.7095 114.32025 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +398.38138 108.07806 moveto +398.38138 110.56244 lineto +401.26419 110.56244 lineto +401.26419 112.56244 lineto +398.38138 112.56244 lineto +398.38138 116.27338 lineto +398.38137 116.67963 398.4621 116.95567 398.62357 117.1015 curveto +398.78502 117.24213 399.10533 117.31244 399.5845 117.31244 curveto +401.022 117.31244 lineto +401.022 119.31244 lineto +398.62357 119.31244 lineto +397.5194 119.31244 396.73554 119.08327 396.272 118.62494 curveto +395.81367 118.1614 395.5845 117.37755 395.5845 116.27338 curveto +395.5845 112.56244 lineto +394.19388 112.56244 lineto +394.19388 110.56244 lineto +395.5845 110.56244 lineto +395.5845 108.07806 lineto +398.38138 108.07806 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +411.71732 114.914 moveto +411.71732 115.71088 lineto +405.17825 115.71088 lineto +405.24596 116.36713 405.48294 116.85932 405.88919 117.18744 curveto +406.29544 117.51557 406.86314 117.67963 407.59232 117.67963 curveto +408.18085 117.67963 408.78241 117.59369 409.397 117.42181 curveto +410.01679 117.24473 410.6522 116.97911 411.30325 116.62494 curveto +411.30325 118.78119 lineto +410.64179 119.03119 409.98033 119.21869 409.31888 119.34369 curveto +408.65741 119.4739 407.99596 119.539 407.3345 119.539 curveto +405.75117 119.539 404.5194 119.13796 403.63919 118.33588 curveto +402.76419 117.52859 402.32669 116.39838 402.32669 114.94525 curveto +402.32669 113.51817 402.75638 112.39578 403.61575 111.57806 curveto +404.48033 110.76036 405.66783 110.35151 407.17825 110.3515 curveto +408.55325 110.35151 409.6522 110.76557 410.47513 111.59369 curveto +411.30324 112.42182 411.71731 113.52859 411.71732 114.914 curveto +408.84232 113.98431 moveto +408.84231 113.45307 408.68606 113.02599 408.37357 112.70306 curveto +408.06627 112.37495 407.66262 112.21088 407.16263 112.21088 curveto +406.62096 112.21088 406.18085 112.36453 405.84232 112.67181 curveto +405.50377 112.9739 405.29283 113.4114 405.2095 113.98431 curveto +408.84232 113.98431 lineto +fill +grestore +grestore +0 0 0 setrgbcolor +[] 0 setdash +1 setlinewidth +0 setlinejoin +0 setlinecap +newpath +460.13031 263.40024 moveto +460.13031 289.97505 lineto +stroke +gsave [2.4492127e-17 0.4 -0.4 2.4492127e-17 460.13031 289.97505] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +gsave [-0.2378166 0 0 0.2269133 546.17466 132.00569] concat +0 0 0 setrgbcolor +[] 0 setdash +4.3013706 setlinewidth +1 setlinejoin +1 setlinecap +newpath +208.65508 282.05865 moveto +193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto +43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto +45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto +177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto +stroke +gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +grestore +gsave [-0.2378166 0 0 0.2269133 546.17465 87.714359] concat +0 0 0 setrgbcolor +[] 0 setdash +4.3013706 setlinewidth +2 setlinejoin +1 setlinecap +newpath +208.65508 282.05865 moveto +193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto +43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto +45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto +177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto +stroke +gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +grestore +gsave [-0.2378166 0 0 0.2269133 546.17465 176.29703] concat +0 0 0 setrgbcolor +[] 0 setdash +4.3013706 setlinewidth +1 setlinejoin +1 setlinecap +newpath +208.65508 282.05865 moveto +193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto +43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto +45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto +177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto +stroke +gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +grestore +gsave [0 0.2378166 0.2269133 0 399.60191 71.056696] concat +0 0 0 setrgbcolor +[] 0 setdash +4.3013706 setlinewidth +1 setlinejoin +1 setlinecap +newpath +208.65508 282.05865 moveto +193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto +43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto +45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto +177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto +stroke +gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +grestore +gsave [1 0 0 1 17.216929 6.5104864] concat +grestore +gsave +0 0 0 setrgbcolor +newpath +187.35507 103.05839 moveto +187.35507 104.61112 lineto +189.20566 104.61112 lineto +189.20566 105.30936 lineto +187.35507 105.30936 lineto +187.35507 108.27811 lineto +187.35507 108.72408 187.41529 109.01054 187.53574 109.13749 curveto +187.65943 109.26444 187.90846 109.32792 188.28281 109.32792 curveto +189.20566 109.32792 lineto +189.20566 110.07987 lineto +188.28281 110.07987 lineto +187.58944 110.07987 187.11093 109.95129 186.84726 109.69413 curveto +186.58359 109.43371 186.45175 108.96171 186.45175 108.27811 curveto +186.45175 105.30936 lineto +185.79257 105.30936 lineto +185.79257 104.61112 lineto +186.45175 104.61112 lineto +186.45175 103.05839 lineto +187.35507 103.05839 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +194.93808 106.77909 moveto +194.93808 110.07987 lineto +194.03964 110.07987 lineto +194.03964 106.80839 lineto +194.03964 106.29081 193.93873 105.90344 193.73691 105.64628 curveto +193.53508 105.38912 193.23235 105.26054 192.8287 105.26054 curveto +192.34368 105.26054 191.96119 105.41516 191.68124 105.7244 curveto +191.40129 106.03365 191.26132 106.4552 191.26132 106.98905 curveto +191.26132 110.07987 lineto +190.358 110.07987 lineto +190.358 102.48222 lineto +191.26132 102.48222 lineto +191.26132 105.46073 lineto +191.47616 105.13196 191.72844 104.88619 192.01816 104.72343 curveto +192.31112 104.56067 192.64804 104.47929 193.0289 104.47929 curveto +193.65715 104.47929 194.13241 104.6746 194.45468 105.06522 curveto +194.77694 105.4526 194.93807 106.02389 194.93808 106.77909 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +201.41757 107.12089 moveto +201.41757 107.56034 lineto +197.28671 107.56034 lineto +197.32577 108.17883 197.51132 108.65084 197.84335 108.97636 curveto +198.17864 109.29862 198.64413 109.45976 199.23984 109.45975 curveto +199.58489 109.45976 199.91854 109.41744 200.24081 109.3328 curveto +200.56633 109.24817 200.8886 109.12121 201.20761 108.95194 curveto +201.20761 109.80155 lineto +200.88534 109.93827 200.55494 110.04244 200.2164 110.11405 curveto +199.87785 110.18567 199.53443 110.22147 199.18613 110.22147 curveto +198.31373 110.22147 197.622 109.96757 197.11093 109.45975 curveto +196.60312 108.95194 196.34921 108.2651 196.34921 107.39921 curveto +196.34921 106.50403 196.5901 105.79439 197.07187 105.2703 curveto +197.55689 104.74296 198.20956 104.47929 199.02988 104.47929 curveto +199.76555 104.47929 200.3466 104.71692 200.77304 105.19218 curveto +201.20272 105.66419 201.41757 106.30709 201.41757 107.12089 curveto +200.51913 106.85722 moveto +200.51262 106.36568 200.37427 105.97343 200.1041 105.68046 curveto +199.83716 105.38749 199.48235 105.24101 199.03964 105.241 curveto +198.53834 105.24101 198.13632 105.38261 197.83359 105.66581 curveto +197.53411 105.94902 197.36158 106.34778 197.31601 106.8621 curveto +200.51913 106.85722 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +205.01132 105.241 moveto +204.52955 105.24101 204.14869 105.42981 203.86874 105.80741 curveto +203.58879 106.18176 203.44882 106.69609 203.44882 107.35038 curveto +203.44882 108.00468 203.58717 108.52063 203.86386 108.89823 curveto +204.14381 109.27258 204.52629 109.45976 205.01132 109.45975 curveto +205.48983 109.45976 205.86907 109.27095 206.14902 108.89335 curveto +206.42896 108.51575 206.56893 108.00142 206.56894 107.35038 curveto +206.56893 106.7026 206.42896 106.1899 206.14902 105.81229 curveto +205.86907 105.43144 205.48983 105.24101 205.01132 105.241 curveto +205.01132 104.47929 moveto +205.79257 104.47929 206.40617 104.7332 206.85214 105.241 curveto +207.2981 105.74882 207.52108 106.45195 207.52109 107.35038 curveto +207.52108 108.24556 207.2981 108.94869 206.85214 109.45975 curveto +206.40617 109.96757 205.79257 110.22147 205.01132 110.22147 curveto +204.22681 110.22147 203.61158 109.96757 203.16562 109.45975 curveto +202.72291 108.94869 202.50156 108.24556 202.50156 107.35038 curveto +202.50156 106.45195 202.72291 105.74882 203.16562 105.241 curveto +203.61158 104.7332 204.22681 104.47929 205.01132 104.47929 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +212.17441 105.45097 moveto +212.07349 105.39238 211.96282 105.35006 211.84238 105.32401 curveto +211.72519 105.29472 211.59498 105.28007 211.45175 105.28007 curveto +210.94394 105.28007 210.55331 105.44609 210.27988 105.77811 curveto +210.00969 106.10689 209.8746 106.58053 209.8746 107.19901 curveto +209.8746 110.07987 lineto +208.97128 110.07987 lineto +208.97128 104.61112 lineto +209.8746 104.61112 lineto +209.8746 105.46073 lineto +210.0634 105.12871 210.30917 104.88294 210.61191 104.72343 curveto +210.91464 104.56067 211.28248 104.47929 211.71542 104.47929 curveto +211.77727 104.47929 211.84563 104.48417 211.9205 104.49393 curveto +211.99537 104.50045 212.07838 104.51184 212.16953 104.52811 curveto +212.17441 105.45097 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +217.58945 107.12089 moveto +217.58945 107.56034 lineto +213.45859 107.56034 lineto +213.49765 108.17883 213.6832 108.65084 214.01523 108.97636 curveto +214.35051 109.29862 214.81601 109.45976 215.41171 109.45975 curveto +215.75676 109.45976 216.09042 109.41744 216.41269 109.3328 curveto +216.73821 109.24817 217.06047 109.12121 217.37949 108.95194 curveto +217.37949 109.80155 lineto +217.05722 109.93827 216.72681 110.04244 216.38828 110.11405 curveto +216.04973 110.18567 215.70631 110.22147 215.358 110.22147 curveto +214.4856 110.22147 213.79387 109.96757 213.28281 109.45975 curveto +212.77499 108.95194 212.52109 108.2651 212.52109 107.39921 curveto +212.52109 106.50403 212.76197 105.79439 213.24374 105.2703 curveto +213.72877 104.74296 214.38144 104.47929 215.20175 104.47929 curveto +215.93742 104.47929 216.51848 104.71692 216.94492 105.19218 curveto +217.3746 105.66419 217.58944 106.30709 217.58945 107.12089 curveto +216.69101 106.85722 moveto +216.68449 106.36568 216.54615 105.97343 216.27597 105.68046 curveto +216.00904 105.38749 215.65422 105.24101 215.21152 105.241 curveto +214.71021 105.24101 214.30819 105.38261 214.00546 105.66581 curveto +213.70598 105.94902 213.53346 106.34778 213.48788 106.8621 curveto +216.69101 106.85722 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +223.32187 105.66093 moveto +223.54647 105.25729 223.81503 104.95943 224.12753 104.76737 curveto +224.44003 104.57532 224.80786 104.47929 225.23105 104.47929 curveto +225.8007 104.47929 226.24016 104.67949 226.54941 105.07987 curveto +226.85864 105.47701 227.01327 106.04342 227.01328 106.77909 curveto +227.01328 110.07987 lineto +226.10995 110.07987 lineto +226.10995 106.80839 lineto +226.10995 106.2843 226.01717 105.89531 225.83163 105.6414 curveto +225.64608 105.38749 225.36288 105.26054 224.98203 105.26054 curveto +224.51652 105.26054 224.14869 105.41516 223.87851 105.7244 curveto +223.60832 106.03365 223.47323 106.4552 223.47324 106.98905 curveto +223.47324 110.07987 lineto +222.56992 110.07987 lineto +222.56992 106.80839 lineto +222.56991 106.28105 222.47714 105.89205 222.2916 105.6414 curveto +222.10604 105.38749 221.81959 105.26054 221.43222 105.26054 curveto +220.97323 105.26054 220.60865 105.41679 220.33847 105.72929 curveto +220.06829 106.03854 219.9332 106.45846 219.9332 106.98905 curveto +219.9332 110.07987 lineto +219.02988 110.07987 lineto +219.02988 104.61112 lineto +219.9332 104.61112 lineto +219.9332 105.46073 lineto +220.13827 105.12545 220.38404 104.87805 220.6705 104.71854 curveto +220.95696 104.55904 221.29713 104.47929 221.69101 104.47929 curveto +222.08814 104.47929 222.42505 104.5802 222.70175 104.78202 curveto +222.98169 104.98385 223.1884 105.27682 223.32187 105.66093 curveto +fill +grestore +gsave [1 0 0 1 17.216929 6.5104864] concat +grestore +gsave +0 0 0 setrgbcolor +newpath +470.46808 277.74594 moveto +470.46808 278.40675 470.60317 278.92596 470.87335 279.30356 curveto +471.14679 279.67791 471.52114 279.86508 471.9964 279.86508 curveto +472.47166 279.86508 472.846 279.67791 473.11945 279.30356 curveto +473.39288 278.92596 473.5296 278.40675 473.5296 277.74594 curveto +473.5296 277.08514 473.39288 276.56756 473.11945 276.19321 curveto +472.846 275.81561 472.47166 275.62681 471.9964 275.6268 curveto +471.52114 275.62681 471.14679 275.81561 470.87335 276.19321 curveto +470.60317 276.56756 470.46808 277.08514 470.46808 277.74594 curveto +473.5296 279.65512 moveto +473.3408 279.98064 473.10154 280.22315 472.81183 280.38266 curveto +472.52537 280.53891 472.18032 280.61703 471.77667 280.61703 curveto +471.11586 280.61703 470.57713 280.35336 470.16046 279.82602 curveto +469.74705 279.29868 469.54034 278.60532 469.54034 277.74594 curveto +469.54034 276.88657 469.74705 276.19321 470.16046 275.66586 curveto +470.57713 275.13852 471.11586 274.87485 471.77667 274.87485 curveto +472.18032 274.87485 472.52537 274.95461 472.81183 275.11411 curveto +473.10154 275.27036 473.3408 275.51125 473.5296 275.83676 curveto +473.5296 275.00668 lineto +474.42804 275.00668 lineto +474.42804 282.55551 lineto +473.5296 282.55551 lineto +473.5296 279.65512 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +480.95636 277.51645 moveto +480.95636 277.9559 lineto +476.8255 277.9559 lineto +476.86456 278.57439 477.05011 279.0464 477.38214 279.37192 curveto +477.71743 279.69418 478.18292 279.85532 478.77863 279.85532 curveto +479.12367 279.85532 479.45733 279.813 479.7796 279.72836 curveto +480.10512 279.64373 480.42738 279.51678 480.7464 279.3475 curveto +480.7464 280.19711 lineto +480.42413 280.33383 480.09372 280.438 479.75519 280.50961 curveto +479.41664 280.58123 479.07322 280.61703 478.72491 280.61703 curveto +477.85252 280.61703 477.16079 280.36313 476.64972 279.85532 curveto +476.14191 279.3475 475.888 278.66066 475.888 277.79477 curveto +475.888 276.89959 476.12889 276.18996 476.61066 275.66586 curveto +477.09568 275.13852 477.74835 274.87485 478.56866 274.87485 curveto +479.30434 274.87485 479.88539 275.11248 480.31183 275.58774 curveto +480.74151 276.05975 480.95635 276.70265 480.95636 277.51645 curveto +480.05792 277.25278 moveto +480.05141 276.76124 479.91306 276.36899 479.64288 276.07602 curveto +479.37595 275.78306 479.02113 275.63657 478.57843 275.63657 curveto +478.07713 275.63657 477.67511 275.77817 477.37238 276.06137 curveto +477.07289 276.34458 476.90037 276.74334 476.8548 277.25766 curveto +480.05792 277.25278 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +486.0296 275.83676 moveto +486.0296 272.87778 lineto +486.92804 272.87778 lineto +486.92804 280.47543 lineto +486.0296 280.47543 lineto +486.0296 279.65512 lineto +485.8408 279.98064 485.60154 280.22315 485.31183 280.38266 curveto +485.02537 280.53891 484.68032 280.61703 484.27667 280.61703 curveto +483.61586 280.61703 483.07713 280.35336 482.66046 279.82602 curveto +482.24705 279.29868 482.04034 278.60532 482.04034 277.74594 curveto +482.04034 276.88657 482.24705 276.19321 482.66046 275.66586 curveto +483.07713 275.13852 483.61586 274.87485 484.27667 274.87485 curveto +484.68032 274.87485 485.02537 274.95461 485.31183 275.11411 curveto +485.60154 275.27036 485.8408 275.51125 486.0296 275.83676 curveto +482.96808 277.74594 moveto +482.96808 278.40675 483.10317 278.92596 483.37335 279.30356 curveto +483.64679 279.67791 484.02114 279.86508 484.4964 279.86508 curveto +484.97166 279.86508 485.346 279.67791 485.61945 279.30356 curveto +485.89288 278.92596 486.0296 278.40675 486.0296 277.74594 curveto +486.0296 277.08514 485.89288 276.56756 485.61945 276.19321 curveto +485.346 275.81561 484.97166 275.62681 484.4964 275.6268 curveto +484.02114 275.62681 483.64679 275.81561 483.37335 276.19321 curveto +483.10317 276.56756 482.96808 277.08514 482.96808 277.74594 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +550.54895 236.85474 moveto +550.54895 237.51555 550.68404 238.03475 550.95422 238.41235 curveto +551.22766 238.7867 551.60201 238.97388 552.07727 238.97388 curveto +552.55253 238.97388 552.92688 238.7867 553.20032 238.41235 curveto +553.47375 238.03475 553.61047 237.51555 553.61047 236.85474 curveto +553.61047 236.19393 553.47375 235.67635 553.20032 235.302 curveto +552.92688 234.9244 552.55253 234.7356 552.07727 234.7356 curveto +551.60201 234.7356 551.22766 234.9244 550.95422 235.302 curveto +550.68404 235.67635 550.54895 236.19393 550.54895 236.85474 curveto +553.61047 238.76392 moveto +553.42167 239.08944 553.18241 239.33195 552.8927 239.49146 curveto +552.60624 239.64771 552.26119 239.72583 551.85754 239.72583 curveto +551.19673 239.72583 550.658 239.46216 550.24133 238.93481 curveto +549.82792 238.40747 549.62122 237.71411 549.62122 236.85474 curveto +549.62122 235.99536 549.82792 235.30201 550.24133 234.77466 curveto +550.658 234.24732 551.19673 233.98365 551.85754 233.98364 curveto +552.26119 233.98365 552.60624 234.0634 552.8927 234.2229 curveto +553.18241 234.37916 553.42167 234.62004 553.61047 234.94556 curveto +553.61047 234.11548 lineto +554.50891 234.11548 lineto +554.50891 241.66431 lineto +553.61047 241.66431 lineto +553.61047 238.76392 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +561.03723 236.62524 moveto +561.03723 237.0647 lineto +556.90637 237.0647 lineto +556.94543 237.68319 557.13098 238.15519 557.46301 238.48071 curveto +557.7983 238.80298 558.26379 238.96411 558.8595 238.96411 curveto +559.20455 238.96411 559.5382 238.92179 559.86047 238.83716 curveto +560.18599 238.75252 560.50825 238.62557 560.82727 238.4563 curveto +560.82727 239.30591 lineto +560.505 239.44263 560.1746 239.54679 559.83606 239.61841 curveto +559.49751 239.69002 559.15409 239.72583 558.80579 239.72583 curveto +557.93339 239.72583 557.24166 239.47192 556.73059 238.96411 curveto +556.22278 238.4563 555.96887 237.76945 555.96887 236.90356 curveto +555.96887 236.00839 556.20976 235.29875 556.69153 234.77466 curveto +557.17655 234.24732 557.82922 233.98365 558.64954 233.98364 curveto +559.38521 233.98365 559.96626 234.22128 560.3927 234.69653 curveto +560.82238 235.16854 561.03723 235.81145 561.03723 236.62524 curveto +560.13879 236.36157 moveto +560.13228 235.87004 559.99393 235.47779 559.72375 235.18481 curveto +559.45682 234.89185 559.10201 234.74537 558.6593 234.74536 curveto +558.158 234.74537 557.75598 234.88697 557.45325 235.17017 curveto +557.15377 235.45337 556.98124 235.85214 556.93567 236.36646 curveto +560.13879 236.36157 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +566.11047 234.94556 moveto +566.11047 231.98657 lineto +567.00891 231.98657 lineto +567.00891 239.58423 lineto +566.11047 239.58423 lineto +566.11047 238.76392 lineto +565.92167 239.08944 565.68241 239.33195 565.3927 239.49146 curveto +565.10624 239.64771 564.76119 239.72583 564.35754 239.72583 curveto +563.69673 239.72583 563.158 239.46216 562.74133 238.93481 curveto +562.32792 238.40747 562.12122 237.71411 562.12122 236.85474 curveto +562.12122 235.99536 562.32792 235.30201 562.74133 234.77466 curveto +563.158 234.24732 563.69673 233.98365 564.35754 233.98364 curveto +564.76119 233.98365 565.10624 234.0634 565.3927 234.2229 curveto +565.68241 234.37916 565.92167 234.62004 566.11047 234.94556 curveto +563.04895 236.85474 moveto +563.04895 237.51555 563.18404 238.03475 563.45422 238.41235 curveto +563.72766 238.7867 564.10201 238.97388 564.57727 238.97388 curveto +565.05253 238.97388 565.42688 238.7867 565.70032 238.41235 curveto +565.97375 238.03475 566.11047 237.51555 566.11047 236.85474 curveto +566.11047 236.19393 565.97375 235.67635 565.70032 235.302 curveto +565.42688 234.9244 565.05253 234.7356 564.57727 234.7356 curveto +564.10201 234.7356 563.72766 234.9244 563.45422 235.302 curveto +563.18404 235.67635 563.04895 236.19393 563.04895 236.85474 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +553.10266 183.66447 moveto +553.10266 184.41154 lineto +552.24329 184.41154 lineto +551.92102 184.41155 551.69641 184.47666 551.56946 184.60686 curveto +551.44576 184.73707 551.38391 184.97145 551.38391 185.30998 curveto +551.38391 185.79338 lineto +552.8634 185.79338 lineto +552.8634 186.49162 lineto +551.38391 186.49162 lineto +551.38391 191.26213 lineto +550.48059 191.26213 lineto +550.48059 186.49162 lineto +549.62122 186.49162 lineto +549.62122 185.79338 lineto +550.48059 185.79338 lineto +550.48059 185.41252 lineto +550.48059 184.8038 550.62219 184.3611 550.9054 184.0844 curveto +551.1886 183.80446 551.63782 183.66448 552.25305 183.66447 curveto +553.10266 183.66447 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +553.84973 185.79338 moveto +554.74817 185.79338 lineto +554.74817 191.26213 lineto +553.84973 191.26213 lineto +553.84973 185.79338 lineto +553.84973 183.66447 moveto +554.74817 183.66447 lineto +554.74817 184.80217 lineto +553.84973 184.80217 lineto +553.84973 183.66447 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +561.16907 185.79338 moveto +559.19153 188.45451 lineto +561.27161 191.26213 lineto +560.21204 191.26213 lineto +558.62024 189.11369 lineto +557.02844 191.26213 lineto +555.96887 191.26213 lineto +558.0929 188.4008 lineto +556.14954 185.79338 lineto +557.20911 185.79338 lineto +558.6593 187.74162 lineto +560.1095 185.79338 lineto +561.16907 185.79338 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +552.81946 198.51311 moveto +552.09354 198.51311 551.59061 198.59612 551.31067 198.76213 curveto +551.03072 198.92815 550.89075 199.21135 550.89075 199.61174 curveto +550.89075 199.93075 550.99491 200.18466 551.20325 200.37346 curveto +551.41483 200.55901 551.70129 200.65178 552.06262 200.65178 curveto +552.56067 200.65178 552.95943 200.476 553.25891 200.12444 curveto +553.56164 199.76962 553.71301 199.29924 553.71301 198.7133 curveto +553.71301 198.51311 lineto +552.81946 198.51311 lineto +554.61145 198.14201 moveto +554.61145 201.26213 lineto +553.71301 201.26213 lineto +553.71301 200.43205 lineto +553.50793 200.76408 553.2524 201.00985 552.94641 201.16936 curveto +552.64042 201.32561 552.26607 201.40373 551.82336 201.40373 curveto +551.26347 201.40373 550.8175 201.24748 550.48547 200.93498 curveto +550.1567 200.61923 549.99231 200.19768 549.99231 199.67033 curveto +549.99231 199.0551 550.19739 198.59123 550.60754 198.27873 curveto +551.02095 197.96624 551.63619 197.80999 552.45325 197.80998 curveto +553.71301 197.80998 lineto +553.71301 197.72209 lineto +553.71301 197.30868 553.57629 196.98967 553.30286 196.76506 curveto +553.03267 196.5372 552.65181 196.42327 552.16028 196.42326 curveto +551.84778 196.42327 551.54341 196.4607 551.24719 196.53557 curveto +550.95097 196.61044 550.66614 196.72275 550.3927 196.87248 curveto +550.3927 196.0424 lineto +550.72147 195.91546 551.04049 195.82106 551.34973 195.7592 curveto +551.65897 195.6941 551.96008 195.66155 552.25305 195.66154 curveto +553.04406 195.66155 553.63488 195.86663 554.02551 196.27678 curveto +554.41613 196.68694 554.61144 197.30868 554.61145 198.14201 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +559.95325 195.95451 moveto +559.95325 196.80412 lineto +559.69934 196.67392 559.43567 196.57626 559.16223 196.51115 curveto +558.88879 196.44605 558.60559 196.4135 558.31262 196.4135 curveto +557.86666 196.4135 557.53137 196.48186 557.30676 196.61858 curveto +557.08541 196.7553 556.97473 196.96038 556.97473 197.23381 curveto +556.97473 197.44215 557.05448 197.60654 557.21399 197.72697 curveto +557.37349 197.84417 557.69413 197.95647 558.1759 198.06389 curveto +558.48352 198.13225 lineto +559.12154 198.26897 559.57401 198.46265 559.84094 198.7133 curveto +560.11112 198.9607 560.24621 199.30738 560.24622 199.75334 curveto +560.24621 200.26116 560.04439 200.66317 559.64075 200.9594 curveto +559.24035 201.25562 558.6886 201.40373 557.98547 201.40373 curveto +557.6925 201.40373 557.38651 201.37444 557.0675 201.31584 curveto +556.75175 201.2605 556.41809 201.17587 556.06653 201.06194 curveto +556.06653 200.1342 lineto +556.39856 200.30673 556.72571 200.43694 557.04797 200.52483 curveto +557.37024 200.60946 557.68925 200.65178 558.005 200.65178 curveto +558.42818 200.65178 558.7537 200.58017 558.98157 200.43694 curveto +559.20943 200.29045 559.32336 200.08537 559.32336 199.8217 curveto +559.32336 199.57756 559.24035 199.39039 559.07434 199.26018 curveto +558.91158 199.12997 558.55188 199.00465 557.99524 198.8842 curveto +557.68274 198.81096 lineto +557.1261 198.69377 556.72408 198.51474 556.47668 198.27385 curveto +556.22929 198.02971 556.10559 197.69605 556.10559 197.27287 curveto +556.10559 196.75855 556.28788 196.36142 556.65247 196.08147 curveto +557.01705 195.80152 557.53463 195.66155 558.2052 195.66154 curveto +558.53723 195.66155 558.84973 195.68596 559.1427 195.73479 curveto +559.43567 195.78362 559.70585 195.85686 559.95325 195.95451 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +565.16809 195.95451 moveto +565.16809 196.80412 lineto +564.91418 196.67392 564.65051 196.57626 564.37708 196.51115 curveto +564.10363 196.44605 563.82043 196.4135 563.52747 196.4135 curveto +563.0815 196.4135 562.74621 196.48186 562.52161 196.61858 curveto +562.30025 196.7553 562.18957 196.96038 562.18958 197.23381 curveto +562.18957 197.44215 562.26933 197.60654 562.42883 197.72697 curveto +562.58834 197.84417 562.90897 197.95647 563.39075 198.06389 curveto +563.69836 198.13225 lineto +564.33638 198.26897 564.78886 198.46265 565.05579 198.7133 curveto +565.32596 198.9607 565.46105 199.30738 565.46106 199.75334 curveto +565.46105 200.26116 565.25923 200.66317 564.85559 200.9594 curveto +564.4552 201.25562 563.90344 201.40373 563.20032 201.40373 curveto +562.90735 201.40373 562.60136 201.37444 562.28235 201.31584 curveto +561.96659 201.2605 561.63293 201.17587 561.28137 201.06194 curveto +561.28137 200.1342 lineto +561.6134 200.30673 561.94055 200.43694 562.26282 200.52483 curveto +562.58508 200.60946 562.90409 200.65178 563.21985 200.65178 curveto +563.64302 200.65178 563.96854 200.58017 564.19641 200.43694 curveto +564.42427 200.29045 564.5382 200.08537 564.53821 199.8217 curveto +564.5382 199.57756 564.4552 199.39039 564.28918 199.26018 curveto +564.12642 199.12997 563.76672 199.00465 563.21008 198.8842 curveto +562.89758 198.81096 lineto +562.34094 198.69377 561.93892 198.51474 561.69153 198.27385 curveto +561.44413 198.02971 561.32043 197.69605 561.32043 197.27287 curveto +561.32043 196.75855 561.50273 196.36142 561.86731 196.08147 curveto +562.23189 195.80152 562.74947 195.66155 563.42004 195.66154 curveto +563.75207 195.66155 564.06457 195.68596 564.35754 195.73479 curveto +564.65051 195.78362 564.92069 195.85686 565.16809 195.95451 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +566.80383 199.10393 moveto +566.80383 195.79338 lineto +567.70227 195.79338 lineto +567.70227 199.06975 lineto +567.70227 199.58733 567.80318 199.97632 568.005 200.23674 curveto +568.20683 200.4939 568.50956 200.62248 568.91321 200.62248 curveto +569.39823 200.62248 569.78072 200.46786 570.06067 200.15862 curveto +570.34387 199.84937 570.48547 199.42782 570.48547 198.89397 curveto +570.48547 195.79338 lineto +571.38391 195.79338 lineto +571.38391 201.26213 lineto +570.48547 201.26213 lineto +570.48547 200.42229 lineto +570.26737 200.75432 570.01346 201.00171 569.72375 201.16447 curveto +569.43729 201.32398 569.10363 201.40373 568.72278 201.40373 curveto +568.09452 201.40373 567.61763 201.20842 567.29211 200.81779 curveto +566.96659 200.42717 566.80383 199.85588 566.80383 199.10393 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +577.50208 196.84319 moveto +577.72668 196.43954 577.99523 196.14169 578.30774 195.94963 curveto +578.62023 195.75758 578.98807 195.66155 579.41125 195.66154 curveto +579.98091 195.66155 580.42036 195.86175 580.72961 196.26213 curveto +581.03885 196.65927 581.19347 197.22568 581.19348 197.96135 curveto +581.19348 201.26213 lineto +580.29016 201.26213 lineto +580.29016 197.99065 lineto +580.29015 197.46656 580.19738 197.07756 580.01184 196.82365 curveto +579.82629 196.56975 579.54308 196.4428 579.16223 196.44279 curveto +578.69673 196.4428 578.32889 196.59742 578.05872 196.90666 curveto +577.78853 197.21591 577.65344 197.63746 577.65344 198.17131 curveto +577.65344 201.26213 lineto +576.75012 201.26213 lineto +576.75012 197.99065 lineto +576.75012 197.46331 576.65734 197.07431 576.4718 196.82365 curveto +576.28625 196.56975 575.99979 196.4428 575.61243 196.44279 curveto +575.15344 196.4428 574.78886 196.59905 574.51868 196.91154 curveto +574.24849 197.22079 574.1134 197.64072 574.1134 198.17131 curveto +574.1134 201.26213 lineto +573.21008 201.26213 lineto +573.21008 195.79338 lineto +574.1134 195.79338 lineto +574.1134 196.64299 lineto +574.31848 196.30771 574.56425 196.06031 574.85071 195.9008 curveto +575.13716 195.7413 575.47733 195.66155 575.87122 195.66154 curveto +576.26835 195.66155 576.60526 195.76246 576.88196 195.96428 curveto +577.1619 196.16611 577.36861 196.45908 577.50208 196.84319 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +587.66809 198.30315 moveto +587.66809 198.7426 lineto +583.53723 198.7426 lineto +583.57629 199.36109 583.76184 199.8331 584.09387 200.15862 curveto +584.42916 200.48088 584.89465 200.64201 585.49036 200.64201 curveto +585.8354 200.64201 586.16906 200.5997 586.49133 200.51506 curveto +586.81685 200.43043 587.13911 200.30347 587.45813 200.1342 curveto +587.45813 200.98381 lineto +587.13586 201.12053 586.80546 201.2247 586.46692 201.29631 curveto +586.12837 201.36792 585.78495 201.40373 585.43665 201.40373 curveto +584.56425 201.40373 583.87252 201.14983 583.36145 200.64201 curveto +582.85364 200.1342 582.59973 199.44735 582.59973 198.58147 curveto +582.59973 197.68629 582.84062 196.97665 583.32239 196.45256 curveto +583.80741 195.92522 584.46008 195.66155 585.2804 195.66154 curveto +586.01607 195.66155 586.59712 195.89918 587.02356 196.37444 curveto +587.45324 196.84645 587.66809 197.48935 587.66809 198.30315 curveto +586.76965 198.03947 moveto +586.76314 197.54794 586.62479 197.15569 586.35461 196.86272 curveto +586.08768 196.56975 585.73287 196.42327 585.29016 196.42326 curveto +584.78886 196.42327 584.38684 196.56487 584.08411 196.84807 curveto +583.78463 197.13128 583.6121 197.53004 583.56653 198.04436 curveto +586.76965 198.03947 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +553.82532 147.89853 moveto +553.82532 148.60165 lineto +553.52258 148.60165 lineto +552.71203 148.60165 552.16841 148.48121 551.89172 148.24033 curveto +551.61828 147.99944 551.48156 147.5193 551.48157 146.7999 curveto +551.48157 145.6329 lineto +551.48156 145.14137 551.39367 144.8012 551.2179 144.6124 curveto +551.04211 144.4236 550.7231 144.3292 550.26086 144.32919 curveto +549.96301 144.32919 lineto +549.96301 143.63095 lineto +550.26086 143.63095 lineto +550.72636 143.63095 551.04537 143.53818 551.2179 143.35263 curveto +551.39367 143.16383 551.48156 142.82692 551.48157 142.34189 curveto +551.48157 141.17001 lineto +551.48156 140.45062 551.61828 139.9721 551.89172 139.73447 curveto +552.16841 139.49359 552.71203 139.37315 553.52258 139.37314 curveto +553.82532 139.37314 lineto +553.82532 140.07138 lineto +553.49329 140.07138 lineto +553.0343 140.07139 552.73482 140.143 552.59485 140.28622 curveto +552.45487 140.42946 552.38488 140.73057 552.38489 141.18954 curveto +552.38489 142.40048 lineto +552.38488 142.91155 552.31001 143.28265 552.16028 143.51376 curveto +552.01379 143.74489 551.76151 143.90114 551.40344 143.98251 curveto +551.76477 144.07041 552.01867 144.22991 552.16516 144.46103 curveto +552.31164 144.69215 552.38488 145.06162 552.38489 145.56943 curveto +552.38489 146.78036 lineto +552.38488 147.23935 552.45487 147.54046 552.59485 147.68369 curveto +552.73482 147.82691 553.0343 147.89853 553.49329 147.89853 curveto +553.82532 147.89853 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +559.51379 147.89853 moveto +559.85559 147.89853 lineto +560.31132 147.89853 560.60754 147.82854 560.74426 147.68857 curveto +560.88423 147.54859 560.95422 147.24586 560.95422 146.78036 curveto +560.95422 145.56943 lineto +560.95422 145.06162 561.02746 144.69215 561.17395 144.46103 curveto +561.32043 144.22991 561.57434 144.07041 561.93567 143.98251 curveto +561.57434 143.90114 561.32043 143.74489 561.17395 143.51376 curveto +561.02746 143.28265 560.95422 142.91155 560.95422 142.40048 curveto +560.95422 141.18954 lineto +560.95422 140.72731 560.88423 140.4262 560.74426 140.28622 curveto +560.60754 140.143 560.31132 140.07139 559.85559 140.07138 curveto +559.51379 140.07138 lineto +559.51379 139.37314 lineto +559.82141 139.37314 lineto +560.63196 139.37315 561.17232 139.49359 561.4425 139.73447 curveto +561.71594 139.9721 561.85266 140.45062 561.85266 141.17001 curveto +561.85266 142.34189 lineto +561.85266 142.82692 561.94055 143.16383 562.11633 143.35263 curveto +562.29211 143.53818 562.61112 143.63095 563.07336 143.63095 curveto +563.3761 143.63095 lineto +563.3761 144.32919 lineto +563.07336 144.32919 lineto +562.61112 144.3292 562.29211 144.4236 562.11633 144.6124 curveto +561.94055 144.8012 561.85266 145.14137 561.85266 145.6329 curveto +561.85266 146.7999 lineto +561.85266 147.5193 561.71594 147.99944 561.4425 148.24033 curveto +561.17232 148.48121 560.63196 148.60165 559.82141 148.60165 curveto +559.51379 148.60165 lineto +559.51379 147.89853 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +554.20129 153.67001 moveto +554.20129 156.97079 lineto +553.30286 156.97079 lineto +553.30286 153.69931 lineto +553.30285 153.18174 553.20194 152.79437 553.00012 152.5372 curveto +552.7983 152.28004 552.49556 152.15146 552.09192 152.15146 curveto +551.60689 152.15146 551.2244 152.30609 550.94446 152.61533 curveto +550.66451 152.92457 550.52453 153.34612 550.52454 153.87997 curveto +550.52454 156.97079 lineto +549.62122 156.97079 lineto +549.62122 151.50204 lineto +550.52454 151.50204 lineto +550.52454 152.35165 lineto +550.73938 152.02288 550.99166 151.77711 551.28137 151.61435 curveto +551.57434 151.45159 551.91125 151.37021 552.29211 151.37021 curveto +552.92037 151.37021 553.39563 151.56553 553.7179 151.95615 curveto +554.04016 152.34352 554.20129 152.91481 554.20129 153.67001 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +560.68079 154.01181 moveto +560.68079 154.45126 lineto +556.54993 154.45126 lineto +556.58899 155.06975 556.77453 155.54176 557.10657 155.86728 curveto +557.44185 156.18955 557.90735 156.35068 558.50305 156.35068 curveto +558.8481 156.35068 559.18176 156.30836 559.50403 156.22372 curveto +559.82954 156.13909 560.15181 156.01214 560.47083 155.84286 curveto +560.47083 156.69247 lineto +560.14855 156.82919 559.81815 156.93336 559.47961 157.00497 curveto +559.14107 157.07659 558.79764 157.1124 558.44934 157.1124 curveto +557.57694 157.1124 556.88521 156.85849 556.37415 156.35068 curveto +555.86633 155.84287 555.61243 155.15602 555.61243 154.29013 curveto +555.61243 153.39495 555.85331 152.68532 556.33508 152.16122 curveto +556.82011 151.63389 557.47278 151.37021 558.29309 151.37021 curveto +559.02876 151.37021 559.60982 151.60784 560.03625 152.0831 curveto +560.46594 152.55511 560.68078 153.19801 560.68079 154.01181 curveto +559.78235 153.74814 moveto +559.77583 153.25661 559.63749 152.86435 559.36731 152.57138 curveto +559.10038 152.27842 558.74556 152.13193 558.30286 152.13193 curveto +557.80155 152.13193 557.39953 152.27353 557.0968 152.55673 curveto +556.79732 152.83994 556.62479 153.2387 556.57922 153.75302 curveto +559.78235 153.74814 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +566.52551 151.50204 moveto +564.54797 154.16318 lineto +566.62805 156.97079 lineto +565.56848 156.97079 lineto +563.97668 154.82236 lineto +562.38489 156.97079 lineto +561.32532 156.97079 lineto +563.44934 154.10947 lineto +561.50598 151.50204 lineto +562.56555 151.50204 lineto +564.01575 153.45029 lineto +565.46594 151.50204 lineto +566.52551 151.50204 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +568.78625 149.94931 moveto +568.78625 151.50204 lineto +570.63684 151.50204 lineto +570.63684 152.20029 lineto +568.78625 152.20029 lineto +568.78625 155.16904 lineto +568.78625 155.615 568.84647 155.90146 568.96692 156.02841 curveto +569.09061 156.15537 569.33964 156.21884 569.71399 156.21884 curveto +570.63684 156.21884 lineto +570.63684 156.97079 lineto +569.71399 156.97079 lineto +569.02063 156.97079 568.54211 156.84221 568.27844 156.58505 curveto +568.01477 156.32464 567.88293 155.85263 567.88293 155.16904 curveto +567.88293 152.20029 lineto +567.22375 152.20029 lineto +567.22375 151.50204 lineto +567.88293 151.50204 lineto +567.88293 149.94931 lineto +568.78625 149.94931 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +483.33514 94.963516 moveto +483.33514 98.264297 lineto +482.43671 98.264297 lineto +482.43671 94.992813 lineto +482.4367 94.475239 482.33579 94.087869 482.13397 93.830704 curveto +481.93215 93.573547 481.62941 93.444966 481.22577 93.444962 curveto +480.74074 93.444966 480.35825 93.599589 480.07831 93.908829 curveto +479.79836 94.218078 479.65838 94.639627 479.65839 95.173477 curveto +479.65839 98.264297 lineto +478.75507 98.264297 lineto +478.75507 92.795547 lineto +479.65839 92.795547 lineto +479.65839 93.645157 lineto +479.87323 93.316386 480.12551 93.070618 480.41522 92.907852 curveto +480.70819 92.745097 481.0451 92.663717 481.42596 92.663712 curveto +482.05422 92.663717 482.52948 92.859029 482.85175 93.249649 curveto +483.17401 93.637023 483.33514 94.208312 483.33514 94.963516 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +487.25604 93.42543 moveto +486.77427 93.425435 486.39341 93.614237 486.11346 93.991837 curveto +485.83351 94.366189 485.69354 94.880512 485.69354 95.534805 curveto +485.69354 96.189104 485.83189 96.705054 486.10858 97.082657 curveto +486.38853 97.457007 486.77101 97.644181 487.25604 97.64418 curveto +487.73455 97.644181 488.11379 97.455379 488.39374 97.077774 curveto +488.67368 96.700171 488.81366 96.185849 488.81366 95.534805 curveto +488.81366 94.887022 488.67368 94.374327 488.39374 93.996719 curveto +488.11379 93.615865 487.73455 93.425435 487.25604 93.42543 curveto +487.25604 92.663712 moveto +488.03729 92.663717 488.65089 92.917623 489.09686 93.42543 curveto +489.54282 93.933247 489.7658 94.636371 489.76581 95.534805 curveto +489.7658 96.429989 489.54282 97.133114 489.09686 97.64418 curveto +488.65089 98.151993 488.03729 98.405899 487.25604 98.405899 curveto +486.47153 98.405899 485.8563 98.151993 485.41034 97.64418 curveto +484.96763 97.133114 484.74628 96.429989 484.74628 95.534805 curveto +484.74628 94.636371 484.96763 93.933247 485.41034 93.42543 curveto +485.8563 92.917623 486.47153 92.663717 487.25604 92.663712 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +492.13885 91.242813 moveto +492.13885 92.795547 lineto +493.98944 92.795547 lineto +493.98944 93.49379 lineto +492.13885 93.49379 lineto +492.13885 96.46254 lineto +492.13885 96.908505 492.19907 97.194963 492.31952 97.321915 curveto +492.44321 97.448869 492.69224 97.512345 493.06659 97.512344 curveto +493.98944 97.512344 lineto +493.98944 98.264297 lineto +493.06659 98.264297 lineto +492.37323 98.264297 491.89471 98.135717 491.63104 97.878555 curveto +491.36737 97.618139 491.23553 97.146135 491.23553 96.46254 curveto +491.23553 93.49379 lineto +490.57635 93.49379 lineto +490.57635 92.795547 lineto +491.23553 92.795547 lineto +491.23553 91.242813 lineto +492.13885 91.242813 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +499.8537 95.305313 moveto +499.8537 95.744766 lineto +495.72284 95.744766 lineto +495.7619 96.363258 495.94745 96.835262 496.27948 97.160782 curveto +496.61476 97.483048 497.08026 97.644181 497.67596 97.64418 curveto +498.02101 97.644181 498.35467 97.601863 498.67694 97.517227 curveto +499.00246 97.432593 499.32472 97.30564 499.64374 97.136368 curveto +499.64374 97.985977 lineto +499.32147 98.122696 498.99106 98.226863 498.65253 98.298477 curveto +498.31398 98.370092 497.97056 98.405899 497.62225 98.405899 curveto +496.74986 98.405899 496.05812 98.151993 495.54706 97.64418 curveto +495.03924 97.136369 494.78534 96.449521 494.78534 95.583633 curveto +494.78534 94.688455 495.02622 93.97882 495.508 93.454727 curveto +495.99302 92.927389 496.64569 92.663717 497.466 92.663712 curveto +498.20168 92.663717 498.78273 92.901347 499.20917 93.376602 curveto +499.63885 93.848612 499.85369 94.491515 499.8537 95.305313 curveto +498.95526 95.041641 moveto +498.94875 94.550108 498.8104 94.157856 498.54022 93.864883 curveto +498.27329 93.571919 497.91847 93.425435 497.47577 93.42543 curveto +496.97446 93.425435 496.57245 93.567037 496.26971 93.850235 curveto +495.97023 94.133442 495.79771 94.532205 495.75214 95.046524 curveto +498.95526 95.041641 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +478.78925 100.66664 moveto +479.68768 100.66664 lineto +479.68768 108.2643 lineto +478.78925 108.2643 lineto +478.78925 100.66664 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +486.24042 105.30531 moveto +486.24042 105.74477 lineto +482.10956 105.74477 lineto +482.14862 106.36326 482.33417 106.83526 482.6662 107.16078 curveto +483.00148 107.48305 483.46698 107.64418 484.06268 107.64418 curveto +484.40773 107.64418 484.74139 107.60186 485.06366 107.51723 curveto +485.38918 107.43259 485.71144 107.30564 486.03046 107.13637 curveto +486.03046 107.98598 lineto +485.70819 108.1227 485.37778 108.22686 485.03925 108.29848 curveto +484.7007 108.37009 484.35728 108.4059 484.00897 108.4059 curveto +483.13657 108.4059 482.44484 108.15199 481.93378 107.64418 curveto +481.42596 107.13637 481.17206 106.44952 481.17206 105.58363 curveto +481.17206 104.68845 481.41294 103.97882 481.89471 103.45473 curveto +482.37974 102.92739 483.03241 102.66372 483.85272 102.66371 curveto +484.5884 102.66372 485.16945 102.90135 485.59589 103.3766 curveto +486.02557 103.84861 486.24041 104.49151 486.24042 105.30531 curveto +485.34198 105.04164 moveto +485.33546 104.55011 485.19712 104.15786 484.92694 103.86488 curveto +484.66001 103.57192 484.30519 103.42544 483.86249 103.42543 curveto +483.36118 103.42544 482.95917 103.56704 482.65643 103.85023 curveto +482.35695 104.13344 482.18443 104.5322 482.13885 105.04652 curveto +485.34198 105.04164 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +488.6037 101.24281 moveto +488.6037 102.79555 lineto +490.45428 102.79555 lineto +490.45428 103.49379 lineto +488.6037 103.49379 lineto +488.6037 106.46254 lineto +488.6037 106.9085 488.66392 107.19496 488.78436 107.32191 curveto +488.90806 107.44887 489.15708 107.51235 489.53143 107.51234 curveto +490.45428 107.51234 lineto +490.45428 108.2643 lineto +489.53143 108.2643 lineto +488.83807 108.2643 488.35956 108.13572 488.09589 107.87856 curveto +487.83221 107.61814 487.70038 107.14613 487.70038 106.46254 curveto +487.70038 103.49379 lineto +487.0412 103.49379 lineto +487.0412 102.79555 lineto +487.70038 102.79555 lineto +487.70038 101.24281 lineto +488.6037 101.24281 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +44.641342 188.13469 moveto +44.641342 184.82414 lineto +45.53978 184.82414 lineto +45.53978 188.10051 lineto +45.539778 188.61809 45.640689 189.00709 45.842514 189.2675 curveto +46.044335 189.52466 46.347069 189.65324 46.750717 189.65324 curveto +47.23574 189.65324 47.618226 189.49862 47.898178 189.18938 curveto +48.181377 188.88013 48.322978 188.45858 48.322983 187.92473 curveto +48.322983 184.82414 lineto +49.22142 184.82414 lineto +49.22142 190.29289 lineto +48.322983 190.29289 lineto +48.322983 189.45305 lineto +48.10488 189.78508 47.850974 190.03248 47.561264 190.19524 curveto +47.274802 190.35474 46.941144 190.43449 46.560287 190.43449 curveto +45.93203 190.43449 45.455143 190.23918 45.129623 189.84856 curveto +44.804102 189.45793 44.641341 188.88664 44.641342 188.13469 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +54.5681 184.98528 moveto +54.5681 185.83488 lineto +54.31419 185.70468 54.050518 185.60702 53.777084 185.54192 curveto +53.503643 185.47682 53.220441 185.44426 52.927475 185.44426 curveto +52.481509 185.44426 52.146223 185.51262 51.921616 185.64934 curveto +51.70026 185.78606 51.589583 185.99114 51.589584 186.26457 curveto +51.589583 186.47291 51.669335 186.6373 51.828842 186.75774 curveto +51.988346 186.87493 52.308983 186.98723 52.790756 187.09465 curveto +53.098373 187.16301 lineto +53.736391 187.29973 54.188864 187.49342 54.455795 187.74406 curveto +54.725973 187.99146 54.861064 188.33814 54.861069 188.7841 curveto +54.861064 189.29192 54.659241 189.69393 54.2556 189.99016 curveto +53.855206 190.28638 53.303448 190.43449 52.600327 190.43449 curveto +52.307356 190.43449 52.001366 190.4052 51.682358 190.3466 curveto +51.366601 190.29126 51.032943 190.20663 50.681381 190.0927 curveto +50.681381 189.16496 lineto +51.013412 189.33749 51.34056 189.4677 51.662827 189.55559 curveto +51.98509 189.64022 52.3041 189.68254 52.619858 189.68254 curveto +53.043032 189.68254 53.368552 189.61093 53.59642 189.4677 curveto +53.824281 189.32121 53.938213 189.11614 53.938217 188.85246 curveto +53.938213 188.60832 53.855206 188.42115 53.689194 188.29094 curveto +53.52643 188.16073 53.16673 188.03541 52.610092 187.91496 curveto +52.297592 187.84172 lineto +51.74095 187.72454 51.338932 187.5455 51.091537 187.30461 curveto +50.844141 187.06047 50.720443 186.72682 50.720444 186.30363 curveto +50.720443 185.78932 50.902735 185.39218 51.267319 185.11223 curveto +51.631901 184.83229 52.149478 184.69231 52.820053 184.69231 curveto +53.152081 184.69231 53.464581 184.71673 53.757553 184.76555 curveto +54.050518 184.81438 54.3207 184.88762 54.5681 184.98528 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +56.296616 184.82414 moveto +57.195053 184.82414 lineto +57.195053 190.29289 lineto +56.296616 190.29289 lineto +56.296616 184.82414 lineto +56.296616 182.69524 moveto +57.195053 182.69524 lineto +57.195053 183.83293 lineto +56.296616 183.83293 lineto +56.296616 182.69524 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +63.615952 186.99211 moveto +63.615952 190.29289 lineto +62.717514 190.29289 lineto +62.717514 187.02141 lineto +62.717509 186.50383 62.616598 186.11646 62.41478 185.8593 curveto +62.212953 185.60214 61.910219 185.47356 61.506577 185.47356 curveto +61.021548 185.47356 60.639061 185.62818 60.359116 185.93742 curveto +60.079166 186.24667 59.939192 186.66822 59.939194 187.20207 curveto +59.939194 190.29289 lineto +59.035873 190.29289 lineto +59.035873 184.82414 lineto +59.939194 184.82414 lineto +59.939194 185.67375 lineto +60.154035 185.34498 60.406314 185.09921 60.69603 184.93645 curveto +60.988996 184.77369 61.325909 184.69231 61.706772 184.69231 curveto +62.335023 184.69231 62.810283 184.88762 63.132553 185.27824 curveto +63.454813 185.66562 63.615946 186.23691 63.615952 186.99211 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +69.016342 187.49504 moveto +69.016338 186.844 68.881247 186.33945 68.611069 185.98137 curveto +68.344138 185.6233 67.968162 185.44426 67.483139 185.44426 curveto +67.001366 185.44426 66.625389 185.6233 66.355209 185.98137 curveto +66.088281 186.33945 65.954817 186.844 65.954819 187.49504 curveto +65.954817 188.14283 66.088281 188.64576 66.355209 189.00383 curveto +66.625389 189.3619 67.001366 189.54094 67.483139 189.54094 curveto +67.968162 189.54094 68.344138 189.3619 68.611069 189.00383 curveto +68.881247 188.64576 69.016338 188.14283 69.016342 187.49504 curveto +69.91478 189.61418 moveto +69.914774 190.54517 69.708069 191.2369 69.294662 191.68938 curveto +68.881247 192.1451 68.248109 192.37297 67.395248 192.37297 curveto +67.079491 192.37297 66.781639 192.34855 66.501694 192.29973 curveto +66.221744 192.25415 65.949934 192.18254 65.686264 192.08488 curveto +65.686264 191.21086 lineto +65.949934 191.35409 66.210351 191.45988 66.467514 191.52824 curveto +66.724673 191.5966 66.986717 191.63078 67.253647 191.63078 curveto +67.842836 191.63078 68.283916 191.47616 68.576889 191.16692 curveto +68.869853 190.86093 69.016338 190.39706 69.016342 189.77531 curveto +69.016342 189.33098 lineto +68.830791 189.65324 68.593161 189.89413 68.303452 190.05363 curveto +68.013734 190.21314 67.667055 190.29289 67.263412 190.29289 curveto +66.592837 190.29289 66.052473 190.03736 65.642319 189.52629 curveto +65.232162 189.01522 65.027084 188.33814 65.027084 187.49504 curveto +65.027084 186.64869 65.232162 185.96998 65.642319 185.45891 curveto +66.052473 184.94785 66.592837 184.69231 67.263412 184.69231 curveto +67.667055 184.69231 68.013734 184.77206 68.303452 184.93156 curveto +68.593161 185.09107 68.830791 185.33196 69.016342 185.65422 curveto +69.016342 184.82414 lineto +69.91478 184.82414 lineto +69.91478 189.61418 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +44.641342 198.13469 moveto +44.641342 194.82414 lineto +45.53978 194.82414 lineto +45.53978 198.10051 lineto +45.539778 198.61809 45.640689 199.00709 45.842514 199.2675 curveto +46.044335 199.52466 46.347069 199.65324 46.750717 199.65324 curveto +47.23574 199.65324 47.618226 199.49862 47.898178 199.18938 curveto +48.181377 198.88013 48.322978 198.45858 48.322983 197.92473 curveto +48.322983 194.82414 lineto +49.22142 194.82414 lineto +49.22142 200.29289 lineto +48.322983 200.29289 lineto +48.322983 199.45305 lineto +48.10488 199.78508 47.850974 200.03248 47.561264 200.19524 curveto +47.274802 200.35474 46.941144 200.43449 46.560287 200.43449 curveto +45.93203 200.43449 45.455143 200.23918 45.129623 199.84856 curveto +44.804102 199.45793 44.641341 198.88664 44.641342 198.13469 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +55.62767 196.99211 moveto +55.62767 200.29289 lineto +54.729233 200.29289 lineto +54.729233 197.02141 lineto +54.729228 196.50383 54.628317 196.11646 54.426498 195.8593 curveto +54.224671 195.60214 53.921937 195.47356 53.518295 195.47356 curveto +53.033266 195.47356 52.65078 195.62818 52.370834 195.93742 curveto +52.090884 196.24667 51.950911 196.66822 51.950912 197.20207 curveto +51.950912 200.29289 lineto +51.047592 200.29289 lineto +51.047592 194.82414 lineto +51.950912 194.82414 lineto +51.950912 195.67375 lineto +52.165754 195.34498 52.418033 195.09921 52.707748 194.93645 curveto +53.000714 194.77369 53.337628 194.69231 53.718491 194.69231 curveto +54.346742 194.69231 54.822002 194.88762 55.144272 195.27824 curveto +55.466532 195.66562 55.627665 196.23691 55.62767 196.99211 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +60.197983 192.69524 moveto +60.197983 193.44231 lineto +59.338608 193.44231 lineto +59.01634 193.44231 58.79173 193.50742 58.66478 193.63762 curveto +58.54108 193.76783 58.479231 194.00221 58.479233 194.34074 curveto +58.479233 194.82414 lineto +59.958725 194.82414 lineto +59.958725 195.52238 lineto +58.479233 195.52238 lineto +58.479233 200.29289 lineto +57.575912 200.29289 lineto +57.575912 195.52238 lineto +56.716537 195.52238 lineto +56.716537 194.82414 lineto +57.575912 194.82414 lineto +57.575912 194.44328 lineto +57.575911 193.83457 57.717513 193.39186 58.000717 193.11516 curveto +58.283918 192.83522 58.733137 192.69524 59.348373 192.69524 curveto +60.197983 192.69524 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +63.064194 195.45403 moveto +62.58242 195.45403 62.201561 195.64283 61.921616 196.02043 curveto +61.641666 196.39478 61.501692 196.90911 61.501694 197.5634 curveto +61.501692 198.2177 61.640038 198.73365 61.916733 199.11125 curveto +62.196679 199.4856 62.579165 199.67278 63.064194 199.67278 curveto +63.542706 199.67278 63.921937 199.48397 64.201889 199.10637 curveto +64.481832 198.72877 64.621806 198.21444 64.621811 197.5634 curveto +64.621806 196.91562 64.481832 196.40292 64.201889 196.02531 curveto +63.921937 195.64446 63.542706 195.45403 63.064194 195.45403 curveto +63.064194 194.69231 moveto +63.84544 194.69231 64.459046 194.94622 64.905014 195.45403 curveto +65.350972 195.96184 65.573954 196.66497 65.573959 197.5634 curveto +65.573954 198.45858 65.350972 199.16171 64.905014 199.67278 curveto +64.459046 200.18059 63.84544 200.43449 63.064194 200.43449 curveto +62.279686 200.43449 61.664452 200.18059 61.218491 199.67278 curveto +60.775781 199.16171 60.554428 198.45858 60.554428 197.5634 curveto +60.554428 196.66497 60.775781 195.96184 61.218491 195.45403 curveto +61.664452 194.94622 62.279686 194.69231 63.064194 194.69231 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +67.058334 192.69524 moveto +67.956772 192.69524 lineto +67.956772 200.29289 lineto +67.058334 200.29289 lineto +67.058334 192.69524 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +73.430405 195.65422 moveto +73.430405 192.69524 lineto +74.328842 192.69524 lineto +74.328842 200.29289 lineto +73.430405 200.29289 lineto +73.430405 199.47258 lineto +73.241598 199.7981 73.002341 200.04061 72.712631 200.20012 curveto +72.426169 200.35637 72.081118 200.43449 71.677475 200.43449 curveto +71.016666 200.43449 70.477929 200.17082 70.061264 199.64348 curveto +69.647852 199.11614 69.441146 198.42278 69.441147 197.5634 curveto +69.441146 196.70403 69.647852 196.01067 70.061264 195.48332 curveto +70.477929 194.95598 71.016666 194.69231 71.677475 194.69231 curveto +72.081118 194.69231 72.426169 194.77206 72.712631 194.93156 curveto +73.002341 195.08782 73.241598 195.3287 73.430405 195.65422 curveto +70.368881 197.5634 moveto +70.36888 198.22421 70.503971 198.74341 70.774155 199.12102 curveto +71.04759 199.49537 71.421939 199.68254 71.897202 199.68254 curveto +72.372458 199.68254 72.746807 199.49537 73.020248 199.12102 curveto +73.293682 198.74341 73.4304 198.22421 73.430405 197.5634 curveto +73.4304 196.9026 73.293682 196.38502 73.020248 196.01067 curveto +72.746807 195.63307 72.372458 195.44426 71.897202 195.44426 curveto +71.421939 195.44426 71.04759 195.63307 70.774155 196.01067 curveto +70.503971 196.38502 70.36888 196.9026 70.368881 197.5634 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +76.179428 194.82414 moveto +77.077866 194.82414 lineto +77.077866 200.29289 lineto +76.179428 200.29289 lineto +76.179428 194.82414 lineto +76.179428 192.69524 moveto +77.077866 192.69524 lineto +77.077866 193.83293 lineto +76.179428 193.83293 lineto +76.179428 192.69524 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +83.498764 196.99211 moveto +83.498764 200.29289 lineto +82.600327 200.29289 lineto +82.600327 197.02141 lineto +82.600322 196.50383 82.499411 196.11646 82.297592 195.8593 curveto +82.095765 195.60214 81.793031 195.47356 81.389389 195.47356 curveto +80.90436 195.47356 80.521874 195.62818 80.241928 195.93742 curveto +79.961978 196.24667 79.822004 196.66822 79.822006 197.20207 curveto +79.822006 200.29289 lineto +78.918686 200.29289 lineto +78.918686 194.82414 lineto +79.822006 194.82414 lineto +79.822006 195.67375 lineto +80.036848 195.34498 80.289126 195.09921 80.578842 194.93645 curveto +80.871808 194.77369 81.208722 194.69231 81.589584 194.69231 curveto +82.217835 194.69231 82.693095 194.88762 83.015366 195.27824 curveto +83.337626 195.66562 83.498759 196.23691 83.498764 196.99211 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +88.899155 197.49504 moveto +88.89915 196.844 88.764059 196.33945 88.493881 195.98137 curveto +88.22695 195.6233 87.850974 195.44426 87.365952 195.44426 curveto +86.884178 195.44426 86.508202 195.6233 86.238022 195.98137 curveto +85.971093 196.33945 85.83763 196.844 85.837631 197.49504 curveto +85.83763 198.14283 85.971093 198.64576 86.238022 199.00383 curveto +86.508202 199.3619 86.884178 199.54094 87.365952 199.54094 curveto +87.850974 199.54094 88.22695 199.3619 88.493881 199.00383 curveto +88.764059 198.64576 88.89915 198.14283 88.899155 197.49504 curveto +89.797592 199.61418 moveto +89.797587 200.54517 89.590881 201.2369 89.177475 201.68938 curveto +88.764059 202.1451 88.130922 202.37297 87.278061 202.37297 curveto +86.962303 202.37297 86.664452 202.34855 86.384506 202.29973 curveto +86.104557 202.25415 85.832747 202.18254 85.569077 202.08488 curveto +85.569077 201.21086 lineto +85.832747 201.35409 86.093163 201.45988 86.350327 201.52824 curveto +86.607486 201.5966 86.86953 201.63078 87.136459 201.63078 curveto +87.725649 201.63078 88.166729 201.47616 88.459702 201.16692 curveto +88.752666 200.86093 88.89915 200.39706 88.899155 199.77531 curveto +88.899155 199.33098 lineto +88.713603 199.65324 88.475973 199.89413 88.186264 200.05363 curveto +87.896547 200.21314 87.549868 200.29289 87.146225 200.29289 curveto +86.47565 200.29289 85.935286 200.03736 85.525131 199.52629 curveto +85.114974 199.01522 84.909896 198.33814 84.909897 197.49504 curveto +84.909896 196.64869 85.114974 195.96998 85.525131 195.45891 curveto +85.935286 194.94785 86.47565 194.69231 87.146225 194.69231 curveto +87.549868 194.69231 87.896547 194.77206 88.186264 194.93156 curveto +88.475973 195.09107 88.713603 195.33196 88.899155 195.65422 curveto +88.899155 194.82414 lineto +89.797592 194.82414 lineto +89.797592 199.61418 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +380.48996 223.50369 moveto +380.48996 225.05643 lineto +382.34055 225.05643 lineto +382.34055 225.75467 lineto +380.48996 225.75467 lineto +380.48996 228.72342 lineto +380.48996 229.16938 380.55018 229.45584 380.67062 229.58279 curveto +380.79432 229.70975 381.04334 229.77322 381.41769 229.77322 curveto +382.34055 229.77322 lineto +382.34055 230.52518 lineto +381.41769 230.52518 lineto +380.72433 230.52518 380.24582 230.3966 379.98215 230.13943 curveto +379.71847 229.87902 379.58664 229.40701 379.58664 228.72342 curveto +379.58664 225.75467 lineto +378.92746 225.75467 lineto +378.92746 225.05643 lineto +379.58664 225.05643 lineto +379.58664 223.50369 lineto +380.48996 223.50369 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +388.07297 227.2244 moveto +388.07297 230.52518 lineto +387.17453 230.52518 lineto +387.17453 227.25369 lineto +387.17453 226.73612 387.07361 226.34875 386.8718 226.09158 curveto +386.66997 225.83443 386.36723 225.70585 385.96359 225.70584 curveto +385.47856 225.70585 385.09608 225.86047 384.81613 226.16971 curveto +384.53618 226.47896 384.39621 226.90051 384.39621 227.43436 curveto +384.39621 230.52518 lineto +383.49289 230.52518 lineto +383.49289 222.92752 lineto +384.39621 222.92752 lineto +384.39621 225.90604 lineto +384.61105 225.57727 384.86333 225.3315 385.15305 225.16873 curveto +385.44601 225.00598 385.78293 224.9246 386.16379 224.92459 curveto +386.79204 224.9246 387.2673 225.11991 387.58957 225.51053 curveto +387.91183 225.8979 388.07296 226.46919 388.07297 227.2244 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +394.55246 227.56619 moveto +394.55246 228.00565 lineto +390.4216 228.00565 lineto +390.46066 228.62414 390.64621 229.09614 390.97824 229.42166 curveto +391.31353 229.74393 391.77902 229.90506 392.37473 229.90506 curveto +392.71977 229.90506 393.05343 229.86274 393.3757 229.77811 curveto +393.70122 229.69347 394.02348 229.56652 394.3425 229.39725 curveto +394.3425 230.24686 lineto +394.02023 230.38358 393.68982 230.48774 393.35129 230.55936 curveto +393.01274 230.63097 392.66932 230.66678 392.32101 230.66678 curveto +391.44862 230.66678 390.75688 230.41287 390.24582 229.90506 curveto +389.73801 229.39725 389.4841 228.7104 389.4841 227.84451 curveto +389.4841 226.94933 389.72498 226.2397 390.20676 225.71561 curveto +390.69178 225.18827 391.34445 224.9246 392.16476 224.92459 curveto +392.90044 224.9246 393.48149 225.16223 393.90793 225.63748 curveto +394.33761 226.10949 394.55245 226.75239 394.55246 227.56619 curveto +393.65402 227.30252 moveto +393.64751 226.81099 393.50916 226.41874 393.23898 226.12576 curveto +392.97205 225.8328 392.61723 225.68631 392.17453 225.68631 curveto +391.67323 225.68631 391.27121 225.82792 390.96848 226.11111 curveto +390.66899 226.39432 390.49647 226.79308 390.4509 227.3074 curveto +393.65402 227.30252 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +400.57297 227.2244 moveto +400.57297 230.52518 lineto +399.67453 230.52518 lineto +399.67453 227.25369 lineto +399.67453 226.73612 399.57361 226.34875 399.3718 226.09158 curveto +399.16997 225.83443 398.86723 225.70585 398.46359 225.70584 curveto +397.97856 225.70585 397.59608 225.86047 397.31613 226.16971 curveto +397.03618 226.47896 396.89621 226.90051 396.89621 227.43436 curveto +396.89621 230.52518 lineto +395.99289 230.52518 lineto +395.99289 225.05643 lineto +396.89621 225.05643 lineto +396.89621 225.90604 lineto +397.11105 225.57727 397.36333 225.3315 397.65305 225.16873 curveto +397.94601 225.00598 398.28293 224.9246 398.66379 224.92459 curveto +399.29204 224.9246 399.7673 225.11991 400.08957 225.51053 curveto +400.41183 225.8979 400.57296 226.46919 400.57297 227.2244 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +239.47623 229.75269 moveto +239.47623 233.05347 lineto +238.57779 233.05347 lineto +238.57779 229.78198 lineto +238.57778 229.26441 238.47687 228.87704 238.27505 228.61987 curveto +238.07323 228.36272 237.77049 228.23414 237.36685 228.23413 curveto +236.88182 228.23414 236.49934 228.38876 236.21939 228.698 curveto +235.93944 229.00725 235.79947 229.4288 235.79947 229.96265 curveto +235.79947 233.05347 lineto +234.89615 233.05347 lineto +234.89615 225.45581 lineto +235.79947 225.45581 lineto +235.79947 228.43433 lineto +236.01431 228.10556 236.26659 227.85979 236.5563 227.69702 curveto +236.84927 227.53427 237.18618 227.45289 237.56705 227.45288 curveto +238.1953 227.45289 238.67056 227.6482 238.99283 228.03882 curveto +239.31509 228.42619 239.47622 228.99748 239.47623 229.75269 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +243.76334 230.30444 moveto +243.03742 230.30445 242.53449 230.38745 242.25455 230.55347 curveto +241.9746 230.71948 241.83462 231.00269 241.83463 231.40308 curveto +241.83462 231.72209 241.93879 231.97599 242.14713 232.16479 curveto +242.35871 232.35034 242.64517 232.44312 243.0065 232.44312 curveto +243.50454 232.44312 243.90331 232.26733 244.20279 231.91577 curveto +244.50552 231.56096 244.65689 231.09058 244.65689 230.50464 curveto +244.65689 230.30444 lineto +243.76334 230.30444 lineto +245.55533 229.93335 moveto +245.55533 233.05347 lineto +244.65689 233.05347 lineto +244.65689 232.22339 lineto +244.45181 232.55542 244.19628 232.80119 243.89029 232.96069 curveto +243.5843 233.11694 243.20995 233.19507 242.76724 233.19507 curveto +242.20734 233.19507 241.76138 233.03882 241.42935 232.72632 curveto +241.10057 232.41056 240.93619 231.98901 240.93619 231.46167 curveto +240.93619 230.84644 241.14127 230.38257 241.55142 230.07007 curveto +241.96483 229.75757 242.58007 229.60132 243.39713 229.60132 curveto +244.65689 229.60132 lineto +244.65689 229.51343 lineto +244.65689 229.10002 244.52017 228.78101 244.24673 228.5564 curveto +243.97655 228.32854 243.59569 228.2146 243.10416 228.2146 curveto +242.79165 228.2146 242.48729 228.25204 242.19107 228.3269 curveto +241.89485 228.40178 241.61001 228.51408 241.33658 228.66382 curveto +241.33658 227.83374 lineto +241.66535 227.70679 241.98436 227.61239 242.29361 227.55054 curveto +242.60285 227.48544 242.90396 227.45289 243.19693 227.45288 curveto +243.98794 227.45289 244.57876 227.65796 244.96939 228.06812 curveto +245.36001 228.47828 245.55532 229.10002 245.55533 229.93335 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +246.76627 227.58472 moveto +247.71841 227.58472 lineto +249.4274 232.17456 lineto +251.13638 227.58472 lineto +252.08853 227.58472 lineto +250.03775 233.05347 lineto +248.81705 233.05347 lineto +246.76627 227.58472 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +258.0065 230.09448 moveto +258.0065 230.53394 lineto +253.87564 230.53394 lineto +253.9147 231.15243 254.10025 231.62443 254.43228 231.94995 curveto +254.76757 232.27222 255.23306 232.43335 255.82877 232.43335 curveto +256.17381 232.43335 256.50747 232.39103 256.82974 232.3064 curveto +257.15526 232.22176 257.47752 232.09481 257.79654 231.92554 curveto +257.79654 232.77515 lineto +257.47427 232.91187 257.14387 233.01603 256.80533 233.08765 curveto +256.46678 233.15926 256.12336 233.19507 255.77505 233.19507 curveto +254.90266 233.19507 254.21093 232.94116 253.69986 232.43335 curveto +253.19205 231.92554 252.93814 231.23869 252.93814 230.3728 curveto +252.93814 229.47762 253.17903 228.76799 253.6608 228.2439 curveto +254.14582 227.71656 254.79849 227.45289 255.6188 227.45288 curveto +256.35448 227.45289 256.93553 227.69052 257.36197 228.16577 curveto +257.79165 228.63778 258.00649 229.28068 258.0065 230.09448 curveto +257.10806 229.83081 moveto +257.10155 229.33928 256.9632 228.94703 256.69302 228.65405 curveto +256.42609 228.36109 256.07128 228.2146 255.62857 228.2146 curveto +255.12727 228.2146 254.72525 228.35621 254.42252 228.6394 curveto +254.12303 228.92261 253.95051 229.32137 253.90494 229.83569 curveto +257.10806 229.83081 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +238.41666 242.74585 moveto +238.41666 243.59546 lineto +238.16275 243.46526 237.89907 243.3676 237.62564 243.30249 curveto +237.3522 243.23739 237.069 243.20484 236.77603 243.20483 curveto +236.33007 243.20484 235.99478 243.2732 235.77017 243.40991 curveto +235.54882 243.54664 235.43814 243.75171 235.43814 244.02515 curveto +235.43814 244.23348 235.51789 244.39787 235.6774 244.51831 curveto +235.8369 244.6355 236.15754 244.74781 236.63931 244.85522 curveto +236.94693 244.92358 lineto +237.58495 245.06031 238.03742 245.25399 238.30435 245.50464 curveto +238.57453 245.75204 238.70962 246.09872 238.70963 246.54468 curveto +238.70962 247.05249 238.5078 247.45451 238.10416 247.75073 curveto +237.70376 248.04696 237.152 248.19507 236.44888 248.19507 curveto +236.15591 248.19507 235.84992 248.16577 235.53091 248.10718 curveto +235.21516 248.05184 234.8815 247.9672 234.52994 247.85327 curveto +234.52994 246.92554 lineto +234.86197 247.09806 235.18912 247.22827 235.51138 247.31616 curveto +235.83365 247.4008 236.15266 247.44312 236.46841 247.44312 curveto +236.89159 247.44312 237.21711 247.3715 237.44498 247.22827 curveto +237.67284 247.08179 237.78677 246.87671 237.78677 246.61304 curveto +237.78677 246.3689 237.70376 246.18172 237.53775 246.05151 curveto +237.37499 245.92131 237.01529 245.79598 236.45865 245.67554 curveto +236.14615 245.60229 lineto +235.58951 245.48511 235.18749 245.30607 234.94009 245.06519 curveto +234.6927 244.82105 234.569 244.48739 234.569 244.06421 curveto +234.569 243.54989 234.75129 243.15276 235.11588 242.8728 curveto +235.48046 242.59286 235.99803 242.45289 236.66861 242.45288 curveto +237.00064 242.45289 237.31314 242.4773 237.60611 242.52612 curveto +237.89907 242.57496 238.16926 242.6482 238.41666 242.74585 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +244.69107 244.75269 moveto +244.69107 248.05347 lineto +243.79263 248.05347 lineto +243.79263 244.78198 lineto +243.79263 244.26441 243.69172 243.87704 243.4899 243.61987 curveto +243.28807 243.36272 242.98534 243.23414 242.5817 243.23413 curveto +242.09667 243.23414 241.71418 243.38876 241.43423 243.698 curveto +241.15428 244.00725 241.01431 244.4288 241.01431 244.96265 curveto +241.01431 248.05347 lineto +240.11099 248.05347 lineto +240.11099 240.45581 lineto +241.01431 240.45581 lineto +241.01431 243.43433 lineto +241.22915 243.10556 241.48143 242.85979 241.77115 242.69702 curveto +242.06411 242.53427 242.40103 242.45289 242.78189 242.45288 curveto +243.41014 242.45289 243.8854 242.6482 244.20767 243.03882 curveto +244.52993 243.42619 244.69107 243.99748 244.69107 244.75269 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +248.61197 243.2146 moveto +248.1302 243.2146 247.74934 243.40341 247.46939 243.78101 curveto +247.18944 244.15536 247.04947 244.66968 247.04947 245.32397 curveto +247.04947 245.97827 247.18781 246.49422 247.46451 246.87183 curveto +247.74445 247.24618 248.12694 247.43335 248.61197 247.43335 curveto +249.09048 247.43335 249.46971 247.24455 249.74966 246.86694 curveto +250.02961 246.48934 250.16958 245.97502 250.16959 245.32397 curveto +250.16958 244.67619 250.02961 244.1635 249.74966 243.78589 curveto +249.46971 243.40503 249.09048 243.2146 248.61197 243.2146 curveto +248.61197 242.45288 moveto +249.39322 242.45289 250.00682 242.70679 250.45279 243.2146 curveto +250.89875 243.72242 251.12173 244.42554 251.12173 245.32397 curveto +251.12173 246.21916 250.89875 246.92228 250.45279 247.43335 curveto +250.00682 247.94116 249.39322 248.19507 248.61197 248.19507 curveto +247.82746 248.19507 247.21223 247.94116 246.76627 247.43335 curveto +246.32356 246.92228 246.1022 246.21916 246.1022 245.32397 curveto +246.1022 244.42554 246.32356 243.72242 246.76627 243.2146 curveto +247.21223 242.70679 247.82746 242.45289 248.61197 242.45288 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +252.08365 242.58472 moveto +252.98209 242.58472 lineto +254.10513 246.85229 lineto +255.2233 242.58472 lineto +256.28287 242.58472 lineto +257.40591 246.85229 lineto +258.52408 242.58472 lineto +259.42252 242.58472 lineto +257.99185 248.05347 lineto +256.93228 248.05347 lineto +255.75552 243.57104 lineto +254.57388 248.05347 lineto +253.51431 248.05347 lineto +252.08365 242.58472 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +311.38464 185.46135 moveto +311.38464 188.76213 lineto +310.48621 188.76213 lineto +310.48621 185.49065 lineto +310.4862 184.97307 310.38529 184.5857 310.18347 184.32854 curveto +309.98164 184.07138 309.67891 183.9428 309.27527 183.94279 curveto +308.79024 183.9428 308.40775 184.09742 308.12781 184.40666 curveto +307.84786 184.71591 307.70788 185.13746 307.70789 185.67131 curveto +307.70789 188.76213 lineto +306.80457 188.76213 lineto +306.80457 181.16447 lineto +307.70789 181.16447 lineto +307.70789 184.14299 lineto +307.92273 183.81422 308.17501 183.56845 308.46472 183.40569 curveto +308.75769 183.24293 309.0946 183.16155 309.47546 183.16154 curveto +310.10371 183.16155 310.57897 183.35686 310.90125 183.74748 curveto +311.22351 184.13486 311.38464 184.70615 311.38464 185.46135 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +315.67175 186.01311 moveto +314.94584 186.01311 314.44291 186.09612 314.16296 186.26213 curveto +313.88301 186.42815 313.74304 186.71135 313.74304 187.11174 curveto +313.74304 187.43075 313.84721 187.68466 314.05554 187.87346 curveto +314.26713 188.05901 314.55359 188.15178 314.91492 188.15178 curveto +315.41296 188.15178 315.81172 187.976 316.11121 187.62444 curveto +316.41394 187.26962 316.5653 186.79924 316.56531 186.2133 curveto +316.56531 186.01311 lineto +315.67175 186.01311 lineto +317.46375 185.64201 moveto +317.46375 188.76213 lineto +316.56531 188.76213 lineto +316.56531 187.93205 lineto +316.36023 188.26408 316.10469 188.50985 315.79871 188.66936 curveto +315.49271 188.82561 315.11836 188.90373 314.67566 188.90373 curveto +314.11576 188.90373 313.6698 188.74748 313.33777 188.43498 curveto +313.00899 188.11923 312.8446 187.69768 312.8446 187.17033 curveto +312.8446 186.5551 313.04968 186.09123 313.45984 185.77873 curveto +313.87325 185.46624 314.48848 185.30999 315.30554 185.30998 curveto +316.56531 185.30998 lineto +316.56531 185.22209 lineto +316.5653 184.80868 316.42858 184.48967 316.15515 184.26506 curveto +315.88497 184.0372 315.50411 183.92327 315.01257 183.92326 curveto +314.70007 183.92327 314.39571 183.9607 314.09949 184.03557 curveto +313.80326 184.11044 313.51843 184.22275 313.245 184.37248 curveto +313.245 183.5424 lineto +313.57377 183.41546 313.89278 183.32106 314.20203 183.2592 curveto +314.51127 183.1941 314.81238 183.16155 315.10535 183.16154 curveto +315.89636 183.16155 316.48718 183.36663 316.87781 183.77678 curveto +317.26843 184.18694 317.46374 184.80868 317.46375 185.64201 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +318.67468 183.29338 moveto +319.62683 183.29338 lineto +321.33582 187.88322 lineto +323.0448 183.29338 lineto +323.99695 183.29338 lineto +321.94617 188.76213 lineto +320.72546 188.76213 lineto +318.67468 183.29338 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +329.91492 185.80315 moveto +329.91492 186.2426 lineto +325.78406 186.2426 lineto +325.82312 186.86109 326.00867 187.3331 326.3407 187.65862 curveto +326.67598 187.98088 327.14148 188.14201 327.73718 188.14201 curveto +328.08223 188.14201 328.41589 188.0997 328.73816 188.01506 curveto +329.06368 187.93043 329.38594 187.80347 329.70496 187.6342 curveto +329.70496 188.48381 lineto +329.38269 188.62053 329.05228 188.7247 328.71375 188.79631 curveto +328.3752 188.86792 328.03178 188.90373 327.68347 188.90373 curveto +326.81107 188.90373 326.11934 188.64983 325.60828 188.14201 curveto +325.10046 187.6342 324.84656 186.94735 324.84656 186.08147 curveto +324.84656 185.18629 325.08744 184.47665 325.56921 183.95256 curveto +326.05424 183.42522 326.70691 183.16155 327.52722 183.16154 curveto +328.26289 183.16155 328.84395 183.39918 329.27039 183.87444 curveto +329.70007 184.34645 329.91491 184.98935 329.91492 185.80315 curveto +329.01648 185.53947 moveto +329.00996 185.04794 328.87162 184.65569 328.60144 184.36272 curveto +328.33451 184.06975 327.97969 183.92327 327.53699 183.92326 curveto +327.03568 183.92327 326.63366 184.06487 326.33093 184.34807 curveto +326.03145 184.63128 325.85893 185.03004 325.81335 185.54436 curveto +329.01648 185.53947 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +310.32507 198.45451 moveto +310.32507 199.30412 lineto +310.07116 199.17392 309.80749 199.07626 309.53406 199.01115 curveto +309.26062 198.94605 308.97741 198.9135 308.68445 198.9135 curveto +308.23848 198.9135 307.9032 198.98186 307.67859 199.11858 curveto +307.45723 199.2553 307.34656 199.46038 307.34656 199.73381 curveto +307.34656 199.94215 307.42631 200.10654 307.58582 200.22697 curveto +307.74532 200.34417 308.06596 200.45647 308.54773 200.56389 curveto +308.85535 200.63225 lineto +309.49336 200.76897 309.94584 200.96265 310.21277 201.2133 curveto +310.48295 201.4607 310.61804 201.80738 310.61804 202.25334 curveto +310.61804 202.76116 310.41621 203.16317 310.01257 203.4594 curveto +309.61218 203.75562 309.06042 203.90373 308.3573 203.90373 curveto +308.06433 203.90373 307.75834 203.87444 307.43933 203.81584 curveto +307.12357 203.7605 306.78992 203.67587 306.43835 203.56194 curveto +306.43835 202.6342 lineto +306.77038 202.80673 307.09753 202.93694 307.4198 203.02483 curveto +307.74206 203.10946 308.06107 203.15178 308.37683 203.15178 curveto +308.80001 203.15178 309.12553 203.08017 309.35339 202.93694 curveto +309.58125 202.79045 309.69519 202.58537 309.69519 202.3217 curveto +309.69519 202.07756 309.61218 201.89039 309.44617 201.76018 curveto +309.2834 201.62997 308.9237 201.50465 308.36707 201.3842 curveto +308.05457 201.31096 lineto +307.49792 201.19377 307.09591 201.01474 306.84851 200.77385 curveto +306.60111 200.52971 306.47742 200.19605 306.47742 199.77287 curveto +306.47742 199.25855 306.65971 198.86142 307.02429 198.58147 curveto +307.38887 198.30152 307.90645 198.16155 308.57703 198.16154 curveto +308.90905 198.16155 309.22155 198.18596 309.51453 198.23479 curveto +309.80749 198.28362 310.07767 198.35686 310.32507 198.45451 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +316.59949 200.46135 moveto +316.59949 203.76213 lineto +315.70105 203.76213 lineto +315.70105 200.49065 lineto +315.70105 199.97307 315.60013 199.5857 315.39832 199.32854 curveto +315.19649 199.07138 314.89375 198.9428 314.49011 198.94279 curveto +314.00508 198.9428 313.6226 199.09742 313.34265 199.40666 curveto +313.0627 199.71591 312.92273 200.13746 312.92273 200.67131 curveto +312.92273 203.76213 lineto +312.01941 203.76213 lineto +312.01941 196.16447 lineto +312.92273 196.16447 lineto +312.92273 199.14299 lineto +313.13757 198.81422 313.38985 198.56845 313.67957 198.40569 curveto +313.97253 198.24293 314.30945 198.16155 314.69031 198.16154 curveto +315.31856 198.16155 315.79382 198.35686 316.11609 198.74748 curveto +316.43835 199.13486 316.59948 199.70615 316.59949 200.46135 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +320.52039 198.92326 moveto +320.03861 198.92327 319.65775 199.11207 319.37781 199.48967 curveto +319.09786 199.86402 318.95788 200.37835 318.95789 201.03264 curveto +318.95788 201.68694 319.09623 202.20289 319.37292 202.58049 curveto +319.65287 202.95484 320.03536 203.14201 320.52039 203.14201 curveto +320.9989 203.14201 321.37813 202.95321 321.65808 202.57561 curveto +321.93802 202.198 322.078 201.68368 322.078 201.03264 curveto +322.078 200.38486 321.93802 199.87216 321.65808 199.49455 curveto +321.37813 199.1137 320.9989 198.92327 320.52039 198.92326 curveto +320.52039 198.16154 moveto +321.30163 198.16155 321.91524 198.41546 322.36121 198.92326 curveto +322.80716 199.43108 323.03015 200.1342 323.03015 201.03264 curveto +323.03015 201.92782 322.80716 202.63095 322.36121 203.14201 curveto +321.91524 203.64983 321.30163 203.90373 320.52039 203.90373 curveto +319.73588 203.90373 319.12064 203.64983 318.67468 203.14201 curveto +318.23197 202.63095 318.01062 201.92782 318.01062 201.03264 curveto +318.01062 200.1342 318.23197 199.43108 318.67468 198.92326 curveto +319.12064 198.41546 319.73588 198.16155 320.52039 198.16154 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +323.99207 198.29338 moveto +324.8905 198.29338 lineto +326.01355 202.56096 lineto +327.13171 198.29338 lineto +328.19128 198.29338 lineto +329.31433 202.56096 lineto +330.4325 198.29338 lineto +331.33093 198.29338 lineto +329.90027 203.76213 lineto +328.8407 203.76213 lineto +327.66394 199.27971 lineto +326.4823 203.76213 lineto +325.42273 203.76213 lineto +323.99207 198.29338 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +305.63477 140.25864 moveto +305.63477 143.15903 lineto +304.73145 143.15903 lineto +304.73145 135.6102 lineto +305.63477 135.6102 lineto +305.63477 136.44028 lineto +305.82357 136.11476 306.0612 135.87388 306.34766 135.71762 curveto +306.63737 135.55812 306.98242 135.47837 307.38281 135.47836 curveto +308.04687 135.47837 308.58561 135.74204 308.99902 136.26938 curveto +309.41568 136.79673 309.62402 137.49009 309.62402 138.34946 curveto +309.62402 139.20883 309.41568 139.90219 308.99902 140.42953 curveto +308.58561 140.95688 308.04687 141.22055 307.38281 141.22055 curveto +306.98242 141.22055 306.63737 141.14243 306.34766 140.98618 curveto +306.0612 140.82667 305.82357 140.58416 305.63477 140.25864 curveto +308.69141 138.34946 moveto +308.6914 137.68865 308.55468 137.17108 308.28125 136.79672 curveto +308.01106 136.41912 307.63834 136.23032 307.16309 136.23032 curveto +306.68782 136.23032 306.31347 136.41912 306.04004 136.79672 curveto +305.76985 137.17108 305.63476 137.68865 305.63477 138.34946 curveto +305.63476 139.01027 305.76985 139.52947 306.04004 139.90707 curveto +306.31347 140.28142 306.68782 140.4686 307.16309 140.4686 curveto +307.63834 140.4686 308.01106 140.28142 308.28125 139.90707 curveto +308.55468 139.52947 308.6914 139.01027 308.69141 138.34946 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +314.28223 136.45004 moveto +314.18131 136.39145 314.07063 136.34914 313.9502 136.32309 curveto +313.833 136.2938 313.7028 136.27915 313.55957 136.27914 curveto +313.05175 136.27915 312.66113 136.44516 312.3877 136.77719 curveto +312.11751 137.10597 311.98242 137.5796 311.98242 138.19809 curveto +311.98242 141.07895 lineto +311.0791 141.07895 lineto +311.0791 135.6102 lineto +311.98242 135.6102 lineto +311.98242 136.45981 lineto +312.17122 136.12778 312.41699 135.88201 312.71973 135.7225 curveto +313.02246 135.55975 313.3903 135.47837 313.82324 135.47836 curveto +313.88509 135.47837 313.95345 135.48325 314.02832 135.49301 curveto +314.10319 135.49953 314.18619 135.51092 314.27734 135.52719 curveto +314.28223 136.45004 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +317.13867 136.24008 moveto +316.6569 136.24009 316.27604 136.42889 315.99609 136.80649 curveto +315.71614 137.18084 315.57617 137.69516 315.57617 138.34946 curveto +315.57617 139.00376 315.71452 139.51971 315.99121 139.89731 curveto +316.27116 140.27166 316.65364 140.45883 317.13867 140.45883 curveto +317.61718 140.45883 317.99642 140.27003 318.27637 139.89243 curveto +318.55631 139.51482 318.69628 139.0005 318.69629 138.34946 curveto +318.69628 137.70167 318.55631 137.18898 318.27637 136.81137 curveto +317.99642 136.43052 317.61718 136.24009 317.13867 136.24008 curveto +317.13867 135.47836 moveto +317.91992 135.47837 318.53352 135.73227 318.97949 136.24008 curveto +319.42545 136.7479 319.64843 137.45102 319.64844 138.34946 curveto +319.64843 139.24464 319.42545 139.94777 318.97949 140.45883 curveto +318.53352 140.96664 317.91992 141.22055 317.13867 141.22055 curveto +316.35416 141.22055 315.73893 140.96664 315.29297 140.45883 curveto +314.85026 139.94777 314.62891 139.24464 314.62891 138.34946 curveto +314.62891 137.45102 314.85026 136.7479 315.29297 136.24008 curveto +315.73893 135.73227 316.35416 135.47837 317.13867 135.47836 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +323.25195 136.24008 moveto +322.77018 136.24009 322.38932 136.42889 322.10938 136.80649 curveto +321.82943 137.18084 321.68945 137.69516 321.68945 138.34946 curveto +321.68945 139.00376 321.8278 139.51971 322.10449 139.89731 curveto +322.38444 140.27166 322.76692 140.45883 323.25195 140.45883 curveto +323.73047 140.45883 324.1097 140.27003 324.38965 139.89243 curveto +324.66959 139.51482 324.80957 139.0005 324.80957 138.34946 curveto +324.80957 137.70167 324.66959 137.18898 324.38965 136.81137 curveto +324.1097 136.43052 323.73047 136.24009 323.25195 136.24008 curveto +323.25195 135.47836 moveto +324.0332 135.47837 324.64681 135.73227 325.09277 136.24008 curveto +325.53873 136.7479 325.76171 137.45102 325.76172 138.34946 curveto +325.76171 139.24464 325.53873 139.94777 325.09277 140.45883 curveto +324.64681 140.96664 324.0332 141.22055 323.25195 141.22055 curveto +322.46745 141.22055 321.85221 140.96664 321.40625 140.45883 curveto +320.96354 139.94777 320.74219 139.24464 320.74219 138.34946 curveto +320.74219 137.45102 320.96354 136.7479 321.40625 136.24008 curveto +321.85221 135.73227 322.46745 135.47837 323.25195 135.47836 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +330.01465 133.48129 moveto +330.01465 134.22836 lineto +329.15527 134.22836 lineto +328.83301 134.22837 328.6084 134.29347 328.48145 134.42368 curveto +328.35775 134.55389 328.2959 134.78827 328.2959 135.1268 curveto +328.2959 135.6102 lineto +329.77539 135.6102 lineto +329.77539 136.30844 lineto +328.2959 136.30844 lineto +328.2959 141.07895 lineto +327.39258 141.07895 lineto +327.39258 136.30844 lineto +326.5332 136.30844 lineto +326.5332 135.6102 lineto +327.39258 135.6102 lineto +327.39258 135.22934 lineto +327.39258 134.62062 327.53418 134.17791 327.81738 133.90121 curveto +328.10058 133.62127 328.5498 133.4813 329.16504 133.48129 curveto +330.01465 133.48129 lineto +fill +grestore +grestore +grestore +showpage +%%EOF diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/document/isar-vm.pdf Binary file doc-src/IsarRef/Thy/document/isar-vm.pdf has changed diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/Thy/document/isar-vm.svg --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Thy/document/isar-vm.svg Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,460 @@ + + + + + + + + + + + + + + + + + + + + + + + + + image/svg+xml + + + + + + + + chain + + + + + + + + prove + + + + + + state + + + + + + + + theorem + + qed + qed + fixassume + { }next + notelet + usingunfolding + then + haveshow + haveshow + proof + + diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/isar-ref.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/IsarRef/style.sty diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Locales/.cvsignore --- a/doc-src/Locales/.cvsignore Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2 +0,0 @@ -locales.out -locales.pdf diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Ref/goals.tex --- a/doc-src/Ref/goals.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,722 +0,0 @@ -%% $Id$ -\chapter{Proof Management: The Subgoal Module} -\index{proofs|(} -\index{subgoal module|(} -\index{reading!goals|see{proofs, starting}} - -The subgoal module stores the current proof state\index{proof state} and -many previous states; commands can produce new states or return to previous -ones. The {\em state list\/} at level $n$ is a list of pairs -\[ [(\psi@n,\Psi@n),\; (\psi@{n-1},\Psi@{n-1}),\; \ldots,\; (\psi@0,[])] \] -where $\psi@n$ is the current proof state, $\psi@{n-1}$ is the previous -one, \ldots, and $\psi@0$ is the initial proof state. The $\Psi@i$ are -sequences (lazy lists) of proof states, storing branch points where a -tactic returned a list longer than one. The state lists permit various -forms of backtracking. - -Chopping elements from the state list reverts to previous proof states. -Besides this, the \ttindex{undo} command keeps a list of state lists. The -module actually maintains a stack of state lists, to support several -proofs at the same time. - -The subgoal module always contains some proof state. At the start of the -Isabelle session, this state consists of a dummy formula. - - -\section{Basic commands} -Most proofs begin with \texttt{Goal} or \texttt{Goalw} and require no other -commands than \texttt{by}, \texttt{chop} and \texttt{undo}. They typically end -with a call to \texttt{qed}. -\subsection{Starting a backward proof} -\index{proofs!starting} -\begin{ttbox} -Goal : string -> thm list -Goalw : thm list -> string -> thm list -goal : theory -> string -> thm list -goalw : theory -> thm list -> string -> thm list -goalw_cterm : thm list -> cterm -> thm list -premises : unit -> thm list -\end{ttbox} - -The goal commands start a new proof by setting the goal. They replace -the current state list by a new one consisting of the initial proof -state. They also empty the \ttindex{undo} list; this command cannot -be undone! - -They all return a list of meta-hypotheses taken from the main goal. If -this list is non-empty, bind its value to an \ML{} identifier by typing -something like -\begin{ttbox} -val prems = goal{\it theory\/ formula}; -\end{ttbox}\index{assumptions!of main goal} -These assumptions typically serve as the premises when you are -deriving a rule. They are also stored internally and can be retrieved -later by the function \texttt{premises}. When the proof is finished, -\texttt{qed} compares the stored assumptions with the actual -assumptions in the proof state. - -The capital versions of \texttt{Goal} are the basic user level -commands and should be preferred over the more primitive lowercase -\texttt{goal} commands. Apart from taking the current theory context -as implicit argument, the former ones try to be smart in handling -meta-hypotheses when deriving rules. Thus \texttt{prems} have to be -seldom bound explicitly, the effect is as if \texttt{cut_facts_tac} -had been called automatically. - -\index{definitions!unfolding} -Some of the commands unfold definitions using meta-rewrite rules. This -expansion affects both the initial subgoal and the premises, which would -otherwise require use of \texttt{rewrite_goals_tac} and -\texttt{rewrite_rule}. - -\begin{ttdescription} -\item[\ttindexbold{Goal} {\it formula};] begins a new proof, where - {\it formula\/} is written as an \ML\ string. - -\item[\ttindexbold{Goalw} {\it defs} {\it formula};] is like - \texttt{Goal} but also applies the list of {\it defs\/} as - meta-rewrite rules to the first subgoal and the premises. - \index{meta-rewriting} - -\item[\ttindexbold{goal} {\it theory} {\it formula};] -begins a new proof, where {\it theory} is usually an \ML\ identifier -and the {\it formula\/} is written as an \ML\ string. - -\item[\ttindexbold{goalw} {\it theory} {\it defs} {\it formula};] -is like \texttt{goal} but also applies the list of {\it defs\/} as -meta-rewrite rules to the first subgoal and the premises. -\index{meta-rewriting} - -\item[\ttindexbold{goalw_cterm} {\it defs} {\it ct};] is - a version of \texttt{goalw} intended for programming. The main - goal is supplied as a cterm, not as a string. See also - \texttt{prove_goalw_cterm}, \S\ref{sec:executing-batch}. - -\item[\ttindexbold{premises}()] -returns the list of meta-hypotheses associated with the current proof (in -case you forgot to bind them to an \ML{} identifier). -\end{ttdescription} - - -\subsection{Applying a tactic} -\index{tactics!commands for applying} -\begin{ttbox} -by : tactic -> unit -byev : tactic list -> unit -\end{ttbox} -These commands extend the state list. They apply a tactic to the current -proof state. If the tactic succeeds, it returns a non-empty sequence of -next states. The head of the sequence becomes the next state, while the -tail is retained for backtracking (see~\texttt{back}). -\begin{ttdescription} \item[\ttindexbold{by} {\it tactic};] -applies the {\it tactic\/} to the proof state. - -\item[\ttindexbold{byev} {\it tactics};] -applies the list of {\it tactics}, one at a time. It is useful for testing -calls to \texttt{prove_goal}, and abbreviates \texttt{by (EVERY {\it -tactics})}. -\end{ttdescription} - -\noindent{\it Error indications:}\nobreak -\begin{itemize} -\item {\footnotesize\tt "by:\ tactic failed"} means that the tactic - returned an empty sequence when applied to the current proof state. -\item {\footnotesize\tt "Warning:\ same as previous level"} means that the - new proof state is identical to the previous state. -\item{\footnotesize\tt "Warning:\ signature of proof state has changed"} - means that some rule was applied whose theory is outside the theory of - the initial proof state. This could signify a mistake such as expressing - the goal in intuitionistic logic and proving it using classical logic. -\end{itemize} - -\subsection{Extracting and storing the proved theorem} -\label{ExtractingAndStoringTheProvedTheorem} -\index{theorems!extracting proved} -\begin{ttbox} -qed : string -> unit -no_qed : unit -> unit -result : unit -> thm -uresult : unit -> thm -bind_thm : string * thm -> unit -bind_thms : string * thm list -> unit -store_thm : string * thm -> thm -store_thms : string * thm list -> thm list -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{qed} $name$;] is the usual way of ending a proof. - It combines \texttt{result} and \texttt{bind_thm}: it gets the theorem - using \texttt{result()} and stores it the theorem database associated - with its theory. See below for details. - -\item[\ttindexbold{no_qed}();] indicates that the proof is not concluded by a - proper \texttt{qed} command. This is a do-nothing operation, it merely - helps user interfaces such as Proof~General \cite{proofgeneral} to figure - out the structure of the {\ML} text. - -\item[\ttindexbold{result}()]\index{assumptions!of main goal} - returns the final theorem, after converting the free variables to - schematics. It discharges the assumptions supplied to the matching - \texttt{goal} command. - - It raises an exception unless the proof state passes certain checks. There - must be no assumptions other than those supplied to \texttt{goal}. There - must be no subgoals. The theorem proved must be a (first-order) instance - of the original goal, as stated in the \texttt{goal} command. This allows - {\bf answer extraction} --- instantiation of variables --- but no other - changes to the main goal. The theorem proved must have the same signature - as the initial proof state. - - These checks are needed because an Isabelle tactic can return any proof - state at all. - -\item[\ttindexbold{uresult}()] is like \texttt{result()} but omits the checks. - It is needed when the initial goal contains function unknowns, when - definitions are unfolded in the main goal (by calling - \ttindex{rewrite_tac}),\index{definitions!unfolding} or when - \ttindex{assume_ax} has been used. - -\item[\ttindexbold{bind_thm} ($name$, $thm$);]\index{theorems!storing} - stores \texttt{standard $thm$} (see \S\ref{MiscellaneousForwardRules}) - in the theorem database associated with its theory and in the {\ML} - variable $name$. The theorem can be retrieved from the database - using \texttt{get_thm} (see \S\ref{BasicOperationsOnTheories}). - -\item[\ttindexbold{store_thm} ($name$, $thm$)]\index{theorems!storing} - stores $thm$ in the theorem database associated with its theory and - returns that theorem. - -\item[\ttindexbold{bind_thms} \textrm{and} \ttindexbold{store_thms}] are similar to - \texttt{bind_thm} and \texttt{store_thm}, but store lists of theorems. - -\end{ttdescription} - -The name argument of \texttt{qed}, \texttt{bind_thm} etc.\ may be the empty -string as well; in that case the result is \emph{not} stored, but proper -checks and presentation of the result still apply. - - -\subsection{Extracting axioms and stored theorems} -\index{theories!axioms of}\index{axioms!extracting} -\index{theories!theorems of}\index{theorems!extracting} -\begin{ttbox} -thm : xstring -> thm -thms : xstring -> thm list -get_axiom : theory -> xstring -> thm -get_thm : theory -> xstring -> thm -get_thms : theory -> xstring -> thm list -axioms_of : theory -> (string * thm) list -thms_of : theory -> (string * thm) list -assume_ax : theory -> string -> thm -\end{ttbox} -\begin{ttdescription} - -\item[\ttindexbold{thm} $name$] is the basic user function for - extracting stored theorems from the current theory context. - -\item[\ttindexbold{thms} $name$] is like \texttt{thm}, but returns a - whole list associated with $name$ rather than a single theorem. - Typically this will be collections stored by packages, e.g.\ - \verb|list.simps|. - -\item[\ttindexbold{get_axiom} $thy$ $name$] returns an axiom with the - given $name$ from $thy$ or any of its ancestors, raising exception - \xdx{THEORY} if none exists. Merging theories can cause several - axioms to have the same name; {\tt get_axiom} returns an arbitrary - one. Usually, axioms are also stored as theorems and may be - retrieved via \texttt{get_thm} as well. - -\item[\ttindexbold{get_thm} $thy$ $name$] is analogous to {\tt - get_axiom}, but looks for a theorem stored in the theory's - database. Like {\tt get_axiom} it searches all parents of a theory - if the theorem is not found directly in $thy$. - -\item[\ttindexbold{get_thms} $thy$ $name$] is like \texttt{get_thm} - for retrieving theorem lists stored within the theory. It returns a - singleton list if $name$ actually refers to a theorem rather than a - theorem list. - -\item[\ttindexbold{axioms_of} $thy$] returns the axioms of this theory - node, not including the ones of its ancestors. - -\item[\ttindexbold{thms_of} $thy$] returns all theorems stored within - the database of this theory node, not including the ones of its - ancestors. - -\item[\ttindexbold{assume_ax} $thy$ $formula$] reads the {\it formula} - using the syntax of $thy$, following the same conventions as axioms - in a theory definition. You can thus pretend that {\it formula} is - an axiom and use the resulting theorem like an axiom. Actually {\tt - assume_ax} returns an assumption; \ttindex{qed} and - \ttindex{result} complain about additional assumptions, but - \ttindex{uresult} does not. - -For example, if {\it formula} is -\hbox{\tt a=b ==> b=a} then the resulting theorem has the form -\hbox{\verb'?a=?b ==> ?b=?a [!!a b. a=b ==> b=a]'} -\end{ttdescription} - - -\subsection{Retrieving theorems} -\index{theorems!retrieving} - -The following functions retrieve theorems (together with their names) -from the theorem database that is associated with the current proof -state's theory. They can only find theorems that have explicitly been -stored in the database using \ttindex{qed}, \ttindex{bind_thm} or -related functions. -\begin{ttbox} -findI : int -> (string * thm) list -findE : int -> int -> (string * thm) list -findEs : int -> (string * thm) list -thms_containing : xstring list -> (string * thm) list -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{findI} $i$]\index{assumptions!of main goal} - returns all ``introduction rules'' applicable to subgoal $i$ --- all - theorems whose conclusion matches (rather than unifies with) subgoal - $i$. Useful in connection with \texttt{resolve_tac}. - -\item[\ttindexbold{findE} $n$ $i$] returns all ``elimination rules'' - applicable to premise $n$ of subgoal $i$ --- all those theorems whose - first premise matches premise $n$ of subgoal $i$. Useful in connection with - \texttt{eresolve_tac} and \texttt{dresolve_tac}. - -\item[\ttindexbold{findEs} $i$] returns all ``elimination rules'' applicable - to subgoal $i$ --- all those theorems whose first premise matches some - premise of subgoal $i$. Useful in connection with \texttt{eresolve_tac} and - \texttt{dresolve_tac}. - -\item[\ttindexbold{thms_containing} $consts$] returns all theorems that - contain \emph{all} of the given constants. -\end{ttdescription} - - -\subsection{Undoing and backtracking} -\begin{ttbox} -chop : unit -> unit -choplev : int -> unit -back : unit -> unit -undo : unit -> unit -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{chop}();] -deletes the top level of the state list, cancelling the last \ttindex{by} -command. It provides a limited undo facility, and the \texttt{undo} command -can cancel it. - -\item[\ttindexbold{choplev} {\it n};] -truncates the state list to level~{\it n}, if $n\geq0$. A negative value -of~$n$ refers to the $n$th previous level: -\hbox{\verb|choplev ~1|} has the same effect as \texttt{chop}. - -\item[\ttindexbold{back}();] -searches the state list for a non-empty branch point, starting from the top -level. The first one found becomes the current proof state --- the most -recent alternative branch is taken. This is a form of interactive -backtracking. - -\item[\ttindexbold{undo}();] -cancels the most recent change to the proof state by the commands \ttindex{by}, -\texttt{chop}, \texttt{choplev}, and~\texttt{back}. It {\bf cannot} -cancel \texttt{goal} or \texttt{undo} itself. It can be repeated to -cancel a series of commands. -\end{ttdescription} - -\goodbreak -\noindent{\it Error indications for {\tt back}:}\par\nobreak -\begin{itemize} -\item{\footnotesize\tt"Warning:\ same as previous choice at this level"} - means \texttt{back} found a non-empty branch point, but that it contained - the same proof state as the current one. -\item{\footnotesize\tt "Warning:\ signature of proof state has changed"} - means the signature of the alternative proof state differs from that of - the current state. -\item {\footnotesize\tt "back:\ no alternatives"} means \texttt{back} could - find no alternative proof state. -\end{itemize} - -\subsection{Printing the proof state}\label{sec:goals-printing} -\index{proof state!printing of} -\begin{ttbox} -pr : unit -> unit -prlev : int -> unit -prlim : int -> unit -goals_limit: int ref \hfill{\bf initially 10} -\end{ttbox} -See also the printing control options described -in~\S\ref{sec:printing-control}. -\begin{ttdescription} -\item[\ttindexbold{pr}();] -prints the current proof state. - -\item[\ttindexbold{prlev} {\it n};] -prints the proof state at level {\it n}, if $n\geq0$. A negative value -of~$n$ refers to the $n$th previous level. This command allows -you to review earlier stages of the proof. - -\item[\ttindexbold{prlim} {\it k};] -prints the current proof state, limiting the number of subgoals to~$k$. It -updates \texttt{goals_limit} (see below) and is helpful when there are many -subgoals. - -\item[\ttindexbold{goals_limit} := {\it k};] -specifies~$k$ as the maximum number of subgoals to print. -\end{ttdescription} - - -\subsection{Timing} -\index{timing statistics}\index{proofs!timing} -\begin{ttbox} -timing: bool ref \hfill{\bf initially false} -\end{ttbox} - -\begin{ttdescription} -\item[set \ttindexbold{timing};] enables global timing in Isabelle. In - particular, this makes the \ttindex{by} and \ttindex{prove_goal} commands - display how much processor time was spent. This information is - compiler-dependent. -\end{ttdescription} - - -\section{Shortcuts for applying tactics} -\index{shortcuts!for \texttt{by} commands} -These commands call \ttindex{by} with common tactics. Their chief purpose -is to minimise typing, although the scanning shortcuts are useful in their -own right. Chapter~\ref{tactics} explains the tactics themselves. - -\subsection{Refining a given subgoal} -\begin{ttbox} -ba : int -> unit -br : thm -> int -> unit -be : thm -> int -> unit -bd : thm -> int -> unit -brs : thm list -> int -> unit -bes : thm list -> int -> unit -bds : thm list -> int -> unit -\end{ttbox} - -\begin{ttdescription} -\item[\ttindexbold{ba} {\it i};] -performs \texttt{by (assume_tac {\it i});} - -\item[\ttindexbold{br} {\it thm} {\it i};] -performs \texttt{by (resolve_tac [{\it thm}] {\it i});} - -\item[\ttindexbold{be} {\it thm} {\it i};] -performs \texttt{by (eresolve_tac [{\it thm}] {\it i});} - -\item[\ttindexbold{bd} {\it thm} {\it i};] -performs \texttt{by (dresolve_tac [{\it thm}] {\it i});} - -\item[\ttindexbold{brs} {\it thms} {\it i};] -performs \texttt{by (resolve_tac {\it thms} {\it i});} - -\item[\ttindexbold{bes} {\it thms} {\it i};] -performs \texttt{by (eresolve_tac {\it thms} {\it i});} - -\item[\ttindexbold{bds} {\it thms} {\it i};] -performs \texttt{by (dresolve_tac {\it thms} {\it i});} -\end{ttdescription} - - -\subsection{Scanning shortcuts} -These shortcuts scan for a suitable subgoal (starting from subgoal~1). -They refine the first subgoal for which the tactic succeeds. Thus, they -require less typing than \texttt{br}, etc. They display the selected -subgoal's number; please watch this, for it may not be what you expect! - -\begin{ttbox} -fa : unit -> unit -fr : thm -> unit -fe : thm -> unit -fd : thm -> unit -frs : thm list -> unit -fes : thm list -> unit -fds : thm list -> unit -\end{ttbox} - -\begin{ttdescription} -\item[\ttindexbold{fa}();] -solves some subgoal by assumption. - -\item[\ttindexbold{fr} {\it thm};] -refines some subgoal using \texttt{resolve_tac [{\it thm}]} - -\item[\ttindexbold{fe} {\it thm};] -refines some subgoal using \texttt{eresolve_tac [{\it thm}]} - -\item[\ttindexbold{fd} {\it thm};] -refines some subgoal using \texttt{dresolve_tac [{\it thm}]} - -\item[\ttindexbold{frs} {\it thms};] -refines some subgoal using \texttt{resolve_tac {\it thms}} - -\item[\ttindexbold{fes} {\it thms};] -refines some subgoal using \texttt{eresolve_tac {\it thms}} - -\item[\ttindexbold{fds} {\it thms};] -refines some subgoal using \texttt{dresolve_tac {\it thms}} -\end{ttdescription} - -\subsection{Other shortcuts} -\begin{ttbox} -bw : thm -> unit -bws : thm list -> unit -ren : string -> int -> unit -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{bw} {\it def};] performs \texttt{by - (rewrite_goals_tac [{\it def}]);} It unfolds definitions in the - subgoals (but not the main goal), by meta-rewriting with the given - definition (see also \S\ref{sec:rewrite_goals}). - \index{meta-rewriting} - -\item[\ttindexbold{bws}] -is like \texttt{bw} but takes a list of definitions. - -\item[\ttindexbold{ren} {\it names} {\it i};] -performs \texttt{by (rename_tac {\it names} {\it i});} it renames -parameters in subgoal~$i$. (Ignore the message {\footnotesize\tt Warning:\ - same as previous level}.) -\index{parameters!renaming} -\end{ttdescription} - - -\section{Executing batch proofs}\label{sec:executing-batch} -\index{batch execution}% -To save space below, let type \texttt{tacfn} abbreviate \texttt{thm list -> - tactic list}, which is the type of a tactical proof. -\begin{ttbox} -prove_goal : theory -> string -> tacfn -> thm -qed_goal : string -> theory -> string -> tacfn -> unit -prove_goalw: theory -> thm list -> string -> tacfn -> thm -qed_goalw : string -> theory -> thm list -> string -> tacfn -> unit -prove_goalw_cterm: thm list -> cterm -> tacfn -> thm -\end{ttbox} -These batch functions create an initial proof state, then apply a tactic to -it, yielding a sequence of final proof states. The head of the sequence is -returned, provided it is an instance of the theorem originally proposed. -The forms \texttt{prove_goal}, \texttt{prove_goalw} and -\texttt{prove_goalw_cterm} are analogous to \texttt{goal}, \texttt{goalw} and -\texttt{goalw_cterm}. - -The tactic is specified by a function from theorem lists to tactic lists. -The function is applied to the list of meta-assumptions taken from -the main goal. The resulting tactics are applied in sequence (using {\tt - EVERY}). For example, a proof consisting of the commands -\begin{ttbox} -val prems = goal {\it theory} {\it formula}; -by \(tac@1\); \ldots by \(tac@n\); -qed "my_thm"; -\end{ttbox} -can be transformed to an expression as follows: -\begin{ttbox} -qed_goal "my_thm" {\it theory} {\it formula} - (fn prems=> [ \(tac@1\), \ldots, \(tac@n\) ]); -\end{ttbox} -The methods perform identical processing of the initial {\it formula} and -the final proof state. But \texttt{prove_goal} executes the tactic as a -atomic operation, bypassing the subgoal module; the current interactive -proof is unaffected. -% -\begin{ttdescription} -\item[\ttindexbold{prove_goal} {\it theory} {\it formula} {\it tacsf};] -executes a proof of the {\it formula\/} in the given {\it theory}, using -the given tactic function. - -\item[\ttindexbold{qed_goal} $name$ $theory$ $formula$ $tacsf$;] acts - like \texttt{prove_goal} but it also stores the resulting theorem in the - theorem database associated with its theory and in the {\ML} - variable $name$ (see \S\ref{ExtractingAndStoringTheProvedTheorem}). - -\item[\ttindexbold{prove_goalw} {\it theory} {\it defs} {\it formula} - {\it tacsf};]\index{meta-rewriting} -is like \texttt{prove_goal} but also applies the list of {\it defs\/} as -meta-rewrite rules to the first subgoal and the premises. - -\item[\ttindexbold{qed_goalw} $name$ $theory$ $defs$ $formula$ $tacsf$;] -is analogous to \texttt{qed_goal}. - -\item[\ttindexbold{prove_goalw_cterm} {\it defs} {\it ct} - {\it tacsf};] -is a version of \texttt{prove_goalw} intended for programming. The main -goal is supplied as a cterm, not as a string. A cterm carries a theory with - it, and can be created from a term~$t$ by -\begin{ttbox} -cterm_of (sign_of thy) \(t\) -\end{ttbox} - \index{*cterm_of}\index{*sign_of} -\end{ttdescription} - - -\section{Internal proofs} - -\begin{ttbox} -Tactic.prove: Sign.sg -> string list -> term list -> term -> - (thm list -> tactic) -> thm -Tactic.prove_standard: Sign.sg -> string list -> term list -> term -> - (thm list -> tactic) -> thm -\end{ttbox} - -These functions provide a clean internal interface for programmed proofs. The -special overhead of top-level statements (cf.\ \verb,prove_goalw_cterm,) is -omitted. Statements may be established within a local contexts of fixed -variables and assumptions, which are the only hypotheses to be discharged in -the result. - -\begin{ttdescription} -\item[\ttindexbold{Tactic.prove}~$sg~xs~As~C~tacf$] establishes the result - $\Forall xs. As \Imp C$ via the given tactic (which is a function taking the - assumptions as local premises). - -\item[\ttindexbold{Tactic.prove_standard}] is simular to \verb,Tactic.prove,, - but performs the \verb,standard, operation on the result, essentially - turning it into a top-level statement. Never do this for local proofs - within other proof tools! - -\end{ttdescription} - - -\section{Managing multiple proofs} -\index{proofs!managing multiple} -You may save the current state of the subgoal module and resume work on it -later. This serves two purposes. -\begin{enumerate} -\item At some point, you may be uncertain of the next step, and -wish to experiment. - -\item During a proof, you may see that a lemma should be proved first. -\end{enumerate} -Each saved proof state consists of a list of levels; \ttindex{chop} behaves -independently for each of the saved proofs. In addition, each saved state -carries a separate \ttindex{undo} list. - -\subsection{The stack of proof states} -\index{proofs!stacking} -\begin{ttbox} -push_proof : unit -> unit -pop_proof : unit -> thm list -rotate_proof : unit -> thm list -\end{ttbox} -The subgoal module maintains a stack of proof states. Most subgoal -commands affect only the top of the stack. The \ttindex{Goal} command {\em -replaces\/} the top of the stack; the only command that pushes a proof on the -stack is \texttt{push_proof}. - -To save some point of the proof, call \texttt{push_proof}. You may now -state a lemma using \texttt{goal}, or simply continue to apply tactics. -Later, you can return to the saved point by calling \texttt{pop_proof} or -\texttt{rotate_proof}. - -To view the entire stack, call \texttt{rotate_proof} repeatedly; as it rotates -the stack, it prints the new top element. - -\begin{ttdescription} -\item[\ttindexbold{push_proof}();] -duplicates the top element of the stack, pushing a copy of the current -proof state on to the stack. - -\item[\ttindexbold{pop_proof}();] -discards the top element of the stack. It returns the list of -assumptions associated with the new proof; you should bind these to an -\ML\ identifier. They can also be obtained by calling \ttindex{premises}. - -\item[\ttindexbold{rotate_proof}();] -\index{assumptions!of main goal} -rotates the stack, moving the top element to the bottom. It returns the -list of assumptions associated with the new proof. -\end{ttdescription} - - -\subsection{Saving and restoring proof states} -\index{proofs!saving and restoring} -\begin{ttbox} -save_proof : unit -> proof -restore_proof : proof -> thm list -\end{ttbox} -States of the subgoal module may be saved as \ML\ values of -type~\mltydx{proof}, and later restored. - -\begin{ttdescription} -\item[\ttindexbold{save_proof}();] -returns the current state, which is on top of the stack. - -\item[\ttindexbold{restore_proof} {\it prf};]\index{assumptions!of main goal} - replaces the top of the stack by~{\it prf}. It returns the list of - assumptions associated with the new proof. -\end{ttdescription} - - -\section{*Debugging and inspecting} -\index{tactics!debugging} -These functions can be useful when you are debugging a tactic. They refer -to the current proof state stored in the subgoal module. A tactic -should never call them; it should operate on the proof state supplied as its -argument. - -\subsection{Reading and printing terms} -\index{terms!reading of}\index{terms!printing of}\index{types!printing of} -\begin{ttbox} -read : string -> term -prin : term -> unit -printyp : typ -> unit -\end{ttbox} -These read and print terms (or types) using the syntax associated with the -proof state. - -\begin{ttdescription} -\item[\ttindexbold{read} {\it string}] -reads the {\it string} as a term, without type-checking. - -\item[\ttindexbold{prin} {\it t};] -prints the term~$t$ at the terminal. - -\item[\ttindexbold{printyp} {\it T};] -prints the type~$T$ at the terminal. -\end{ttdescription} - -\subsection{Inspecting the proof state} -\index{proofs!inspecting the state} -\begin{ttbox} -topthm : unit -> thm -getgoal : int -> term -gethyps : int -> thm list -\end{ttbox} - -\begin{ttdescription} -\item[\ttindexbold{topthm}()] -returns the proof state as an Isabelle theorem. This is what \ttindex{by} -would supply to a tactic at this point. It omits the post-processing of -\ttindex{result} and \ttindex{uresult}. - -\item[\ttindexbold{getgoal} {\it i}] -returns subgoal~$i$ of the proof state, as a term. You may print -this using \texttt{prin}, though you may have to examine the internal -data structure in order to locate the problem! - -\item[\ttindexbold{gethyps} {\it i}] - returns the hypotheses of subgoal~$i$ as meta-level assumptions. In - these theorems, the subgoal's parameters become free variables. This - command is supplied for debugging uses of \ttindex{METAHYPS}. -\end{ttdescription} - - -\subsection{Filtering lists of rules} -\begin{ttbox} -filter_goal: (term*term->bool) -> thm list -> int -> thm list -\end{ttbox} - -\begin{ttdescription} -\item[\ttindexbold{filter_goal} {\it could} {\it ths} {\it i}] -applies \texttt{filter_thms {\it could}} to subgoal~$i$ of the proof -state and returns the list of theorems that survive the filtering. -\end{ttdescription} - -\index{subgoal module|)} -\index{proofs|)} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "ref" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Ref/ref.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/Ref/theory-syntax.tex --- a/doc-src/Ref/theory-syntax.tex Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,178 +0,0 @@ -%% $Id$ - -\appendix -\newlinechar=-1 %mathsing.sty sets \newlinechar=`\|, which would cause mayhem - -\chapter{Syntax of Isabelle Theories}\label{app:TheorySyntax} - -Below we present the full syntax of theory definition files as provided by -Pure Isabelle --- object-logics may add their own sections. -\S\ref{sec:ref-defining-theories} explains the meanings of these constructs. -The syntax obeys the following conventions: -\begin{itemize} -\item {\tt Typewriter font} denotes terminal symbols. - -\item $id$, $tid$, $nat$, $string$ and $longident$ are the lexical - classes of identifiers, type identifiers, natural numbers, quoted - strings (without the need for \verb$\$\dots\verb$\$ between lines) - and long qualified \ML{} identifiers. - The categories $id$, $tid$, $nat$ are fully defined in \iflabelundefined{Defining-Logics}% - {{\it The Isabelle Reference Manual}, chapter `Defining Logics'}% - {\S\ref{Defining-Logics}}. - -\item $text$ is all text from the current position to the end of file, - $verbatim$ is any text enclosed in \verb.{|.\dots\verb.|}. - -\item Comments in theories take the form {\tt (*}\dots{\tt*)} and may - be nested, just as in \ML. -\end{itemize} - -\begin{rail} - -theoryDef : id '=' (name + '+') ('+' extension | ()) - ; - -name : id | string - ; - -extension : (section +) 'end' ( () | ml ) - ; - -section : classes - | default - | types - | arities - | nonterminals - | consts - | syntax - | trans - | defs - | constdefs - | rules - | axclass - | instance - | oracle - | locale - | local - | global - | setup - ; - -classes : 'classes' ( classDecl + ) - ; - -classDecl : (id (() | '<' (id + ','))) - ; - -default : 'default' sort - ; - -sort : id - | lbrace (id * ',') rbrace - ; - -types : 'types' ( ( typeDecl ( () | '(' infix ')' ) ) + ) - ; - -infix : ( 'infixr' | 'infixl' ) (() | string) nat - ; - -typeDecl : typevarlist name - ( () | '=' ( string | type ) ); - -typevarlist : () | tid | '(' ( tid + ',' ) ')'; - -type : simpleType | '(' type ')' | type '=>' type | - '[' ( type + "," ) ']' '=>' type; - -simpleType: id | ( tid ( () | '::' id ) ) | - '(' ( type + "," ) ')' id | simpleType id - ; - -arities : 'arities' ((name + ',') '::' arity +) - ; - -arity : ( () | '(' (sort + ',') ')' ) sort - ; - -nonterminals : 'nonterminals' (name+) - ; - -consts : 'consts' ( mixfixConstDecl + ) - ; - -syntax : 'syntax' (() | mode) ( mixfixConstDecl + ); - -mode : '(' name (() | 'output') ')' - ; - -mixfixConstDecl : constDecl (() | ( '(' mixfix ')' )) - ; - -constDecl : ( name + ',') '::' (string | type); - -mixfix : string ( () | ( () | ('[' (nat + ',') ']')) nat ) - | infix - | 'binder' string nat ; - -trans : 'translations' ( pat ( '==' | '=>' | '<=' ) pat + ) - ; - -pat : ( () | ( '(' id ')' ) ) string; - -rules : 'rules' (( id string ) + ) - ; - -defs : 'defs' (( id string ) + ) - ; - -constdefs : 'constdefs' (name '::' (string | type) (() | mixfix) string +) - ; - -axclass : 'axclass' classDecl (() | ( id string ) +) - ; - -instance : 'instance' ( name '<' name | name '::' arity) witness - ; - -witness : (() | '(' ((string | id | longident) + ',') ')') (() | verbatim) - ; - -locale : 'locale' name '=' ( () | name '+' ) localeBody - ; - -localeBody : localeConsts ( () | localeAsms ) ( () | localeDefs ) - ; - -localeConsts: ( 'fixes' ( ( (name '::' ( string | type )) ( () | '(' mixfix ')' ) ) + ) ) - ; - - -localeAsms: ( 'assumes' ( ( id string ) + ) ) - ; - -localeDefs: ( 'defines' ( ( id string ) +) ) - ; - -oracle : 'oracle' name '=' name - ; - -local : 'local' - ; - -global : 'global' - ; - -setup : 'setup' (id | longident) - ; - -ml : 'ML' text - ; - -\end{rail} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "ref" -%%% End: diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/TutorialI/Types/Numbers.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/TutorialI/Types/document/Numbers.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/TutorialI/Types/numerics.tex diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/antiquote_setup.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc-src/manual.bib diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 doc/Contents diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 etc/settings diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 lib/Tools/codegen diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 lib/browser/.cvsignore --- a/lib/browser/.cvsignore Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1 +0,0 @@ -GraphBrowser.jar diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 lib/browser/GraphBrowser/.cvsignore --- a/lib/browser/GraphBrowser/.cvsignore Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1 +0,0 @@ -*.class diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 lib/browser/awtUtilities/.cvsignore --- a/lib/browser/awtUtilities/.cvsignore Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1 +0,0 @@ -*.class diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/FOL/IFOL.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/FOL/IsaMakefile diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/FOL/ex/IffOracle.thy --- a/src/FOL/ex/IffOracle.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,77 +0,0 @@ -(* Title: FOL/ex/IffOracle.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1996 University of Cambridge -*) - -header {* Example of Declaring an Oracle *} - -theory IffOracle -imports FOL -begin - -subsection {* Oracle declaration *} - -text {* - This oracle makes tautologies of the form @{text "P <-> P <-> P <-> P"}. - The length is specified by an integer, which is checked to be even - and positive. -*} - -oracle iff_oracle = {* - let - fun mk_iff 1 = Var (("P", 0), @{typ o}) - | mk_iff n = FOLogic.iff $ Var (("P", 0), @{typ o}) $ mk_iff (n - 1); - in - fn (thy, n) => - if n > 0 andalso n mod 2 = 0 - then Thm.cterm_of thy (FOLogic.mk_Trueprop (mk_iff n)) - else raise Fail ("iff_oracle: " ^ string_of_int n) - end -*} - - -subsection {* Oracle as low-level rule *} - -ML {* iff_oracle (@{theory}, 2) *} -ML {* iff_oracle (@{theory}, 10) *} -ML {* Thm.proof_of (iff_oracle (@{theory}, 10)) *} - -text {* These oracle calls had better fail. *} - -ML {* - (iff_oracle (@{theory}, 5); error "?") - handle Fail _ => warning "Oracle failed, as expected" -*} - -ML {* - (iff_oracle (@{theory}, 1); error "?") - handle Fail _ => warning "Oracle failed, as expected" -*} - - -subsection {* Oracle as proof method *} - -method_setup iff = {* - Method.simple_args OuterParse.nat (fn n => fn ctxt => - Method.SIMPLE_METHOD - (HEADGOAL (Tactic.rtac (iff_oracle (ProofContext.theory_of ctxt, n))) - handle Fail _ => no_tac)) -*} "iff oracle" - - -lemma "A <-> A" - by (iff 2) - -lemma "A <-> A <-> A <-> A <-> A <-> A <-> A <-> A <-> A <-> A" - by (iff 10) - -lemma "A <-> A <-> A <-> A <-> A" - apply (iff 5)? - oops - -lemma A - apply (iff 1)? - oops - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/FOL/ex/Iff_Oracle.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/Iff_Oracle.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,76 @@ +(* Title: FOL/ex/Iff_Oracle.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1996 University of Cambridge +*) + +header {* Example of Declaring an Oracle *} + +theory Iff_Oracle +imports FOL +begin + +subsection {* Oracle declaration *} + +text {* + This oracle makes tautologies of the form @{text "P <-> P <-> P <-> P"}. + The length is specified by an integer, which is checked to be even + and positive. +*} + +oracle iff_oracle = {* + let + fun mk_iff 1 = Var (("P", 0), @{typ o}) + | mk_iff n = FOLogic.iff $ Var (("P", 0), @{typ o}) $ mk_iff (n - 1); + in + fn (thy, n) => + if n > 0 andalso n mod 2 = 0 + then Thm.cterm_of thy (FOLogic.mk_Trueprop (mk_iff n)) + else raise Fail ("iff_oracle: " ^ string_of_int n) + end +*} + + +subsection {* Oracle as low-level rule *} + +ML {* iff_oracle (@{theory}, 2) *} +ML {* iff_oracle (@{theory}, 10) *} +ML {* Thm.proof_of (iff_oracle (@{theory}, 10)) *} + +text {* These oracle calls had better fail. *} + +ML {* + (iff_oracle (@{theory}, 5); error "?") + handle Fail _ => warning "Oracle failed, as expected" +*} + +ML {* + (iff_oracle (@{theory}, 1); error "?") + handle Fail _ => warning "Oracle failed, as expected" +*} + + +subsection {* Oracle as proof method *} + +method_setup iff = {* + Method.simple_args OuterParse.nat (fn n => fn ctxt => + Method.SIMPLE_METHOD + (HEADGOAL (Tactic.rtac (iff_oracle (ProofContext.theory_of ctxt, n))) + handle Fail _ => no_tac)) +*} "iff oracle" + + +lemma "A <-> A" + by (iff 2) + +lemma "A <-> A <-> A <-> A <-> A <-> A <-> A <-> A <-> A <-> A" + by (iff 10) + +lemma "A <-> A <-> A <-> A <-> A" + apply (iff 5)? + oops + +lemma A + apply (iff 1)? + oops + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/FOL/ex/NatClass.thy --- a/src/FOL/ex/NatClass.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,90 +0,0 @@ -(* Title: FOL/ex/NatClass.thy - ID: $Id$ - Author: Markus Wenzel, TU Muenchen -*) - -theory NatClass -imports FOL -begin - -text {* - This is an abstract version of theory @{text "Nat"}. Instead of - axiomatizing a single type @{text nat} we define the class of all - these types (up to isomorphism). - - Note: The @{text rec} operator had to be made \emph{monomorphic}, - because class axioms may not contain more than one type variable. -*} - -consts - 0 :: 'a ("0") - Suc :: "'a => 'a" - rec :: "['a, 'a, ['a, 'a] => 'a] => 'a" - -axclass - nat < "term" - induct: "[| P(0); !!x. P(x) ==> P(Suc(x)) |] ==> P(n)" - Suc_inject: "Suc(m) = Suc(n) ==> m = n" - Suc_neq_0: "Suc(m) = 0 ==> R" - rec_0: "rec(0, a, f) = a" - rec_Suc: "rec(Suc(m), a, f) = f(m, rec(m, a, f))" - -definition - add :: "['a::nat, 'a] => 'a" (infixl "+" 60) where - "m + n = rec(m, n, %x y. Suc(y))" - -lemma Suc_n_not_n: "Suc(k) ~= (k::'a::nat)" -apply (rule_tac n = k in induct) -apply (rule notI) -apply (erule Suc_neq_0) -apply (rule notI) -apply (erule notE) -apply (erule Suc_inject) -done - -lemma "(k+m)+n = k+(m+n)" -apply (rule induct) -back -back -back -back -back -back -oops - -lemma add_0 [simp]: "0+n = n" -apply (unfold add_def) -apply (rule rec_0) -done - -lemma add_Suc [simp]: "Suc(m)+n = Suc(m+n)" -apply (unfold add_def) -apply (rule rec_Suc) -done - -lemma add_assoc: "(k+m)+n = k+(m+n)" -apply (rule_tac n = k in induct) -apply simp -apply simp -done - -lemma add_0_right: "m+0 = m" -apply (rule_tac n = m in induct) -apply simp -apply simp -done - -lemma add_Suc_right: "m+Suc(n) = Suc(m+n)" -apply (rule_tac n = m in induct) -apply simp_all -done - -lemma - assumes prem: "!!n. f(Suc(n)) = Suc(f(n))" - shows "f(i+j) = i+f(j)" -apply (rule_tac n = i in induct) -apply simp -apply (simp add: prem) -done - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/FOL/ex/Nat_Class.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/Nat_Class.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,88 @@ +(* Title: FOL/ex/Nat_Class.thy + Author: Markus Wenzel, TU Muenchen +*) + +theory Nat_Class +imports FOL +begin + +text {* + This is an abstract version of theory @{text Nat}. Instead of + axiomatizing a single type @{text nat} we define the class of all + these types (up to isomorphism). + + Note: The @{text rec} operator had to be made \emph{monomorphic}, + because class axioms may not contain more than one type variable. +*} + +class nat = + fixes Zero :: 'a ("0") + and Suc :: "'a \ 'a" + and rec :: "'a \ 'a \ ('a \ 'a \ 'a) \ 'a" + assumes induct: "P(0) \ (\x. P(x) \ P(Suc(x))) \ P(n)" + and Suc_inject: "Suc(m) = Suc(n) \ m = n" + and Suc_neq_Zero: "Suc(m) = 0 \ R" + and rec_Zero: "rec(0, a, f) = a" + and rec_Suc: "rec(Suc(m), a, f) = f(m, rec(m, a, f))" +begin + +definition + add :: "'a \ 'a \ 'a" (infixl "+" 60) where + "m + n = rec(m, n, \x y. Suc(y))" + +lemma Suc_n_not_n: "Suc(k) \ (k::'a)" + apply (rule_tac n = k in induct) + apply (rule notI) + apply (erule Suc_neq_Zero) + apply (rule notI) + apply (erule notE) + apply (erule Suc_inject) + done + +lemma "(k + m) + n = k + (m + n)" + apply (rule induct) + back + back + back + back + back + oops + +lemma add_Zero [simp]: "0 + n = n" + apply (unfold add_def) + apply (rule rec_Zero) + done + +lemma add_Suc [simp]: "Suc(m) + n = Suc(m + n)" + apply (unfold add_def) + apply (rule rec_Suc) + done + +lemma add_assoc: "(k + m) + n = k + (m + n)" + apply (rule_tac n = k in induct) + apply simp + apply simp + done + +lemma add_Zero_right: "m + 0 = m" + apply (rule_tac n = m in induct) + apply simp + apply simp + done + +lemma add_Suc_right: "m + Suc(n) = Suc(m + n)" + apply (rule_tac n = m in induct) + apply simp_all + done + +lemma + assumes prem: "\n. f(Suc(n)) = Suc(f(n))" + shows "f(i + j) = i + f(j)" + apply (rule_tac n = i in induct) + apply simp + apply (simp add: prem) + done + +end + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Algebra/Exponent.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Archimedean_Field.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Archimedean_Field.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,400 @@ +(* Title: Archimedean_Field.thy + Author: Brian Huffman +*) + +header {* Archimedean Fields, Floor and Ceiling Functions *} + +theory Archimedean_Field +imports Main +begin + +subsection {* Class of Archimedean fields *} + +text {* Archimedean fields have no infinite elements. *} + +class archimedean_field = ordered_field + number_ring + + assumes ex_le_of_int: "\z. x \ of_int z" + +lemma ex_less_of_int: + fixes x :: "'a::archimedean_field" shows "\z. x < of_int z" +proof - + from ex_le_of_int obtain z where "x \ of_int z" .. + then have "x < of_int (z + 1)" by simp + then show ?thesis .. +qed + +lemma ex_of_int_less: + fixes x :: "'a::archimedean_field" shows "\z. of_int z < x" +proof - + from ex_less_of_int obtain z where "- x < of_int z" .. + then have "of_int (- z) < x" by simp + then show ?thesis .. +qed + +lemma ex_less_of_nat: + fixes x :: "'a::archimedean_field" shows "\n. x < of_nat n" +proof - + obtain z where "x < of_int z" using ex_less_of_int .. + also have "\ \ of_int (int (nat z))" by simp + also have "\ = of_nat (nat z)" by (simp only: of_int_of_nat_eq) + finally show ?thesis .. +qed + +lemma ex_le_of_nat: + fixes x :: "'a::archimedean_field" shows "\n. x \ of_nat n" +proof - + obtain n where "x < of_nat n" using ex_less_of_nat .. + then have "x \ of_nat n" by simp + then show ?thesis .. +qed + +text {* Archimedean fields have no infinitesimal elements. *} + +lemma ex_inverse_of_nat_Suc_less: + fixes x :: "'a::archimedean_field" + assumes "0 < x" shows "\n. inverse (of_nat (Suc n)) < x" +proof - + from `0 < x` have "0 < inverse x" + by (rule positive_imp_inverse_positive) + obtain n where "inverse x < of_nat n" + using ex_less_of_nat .. + then obtain m where "inverse x < of_nat (Suc m)" + using `0 < inverse x` by (cases n) (simp_all del: of_nat_Suc) + then have "inverse (of_nat (Suc m)) < inverse (inverse x)" + using `0 < inverse x` by (rule less_imp_inverse_less) + then have "inverse (of_nat (Suc m)) < x" + using `0 < x` by (simp add: nonzero_inverse_inverse_eq) + then show ?thesis .. +qed + +lemma ex_inverse_of_nat_less: + fixes x :: "'a::archimedean_field" + assumes "0 < x" shows "\n>0. inverse (of_nat n) < x" + using ex_inverse_of_nat_Suc_less [OF `0 < x`] by auto + +lemma ex_less_of_nat_mult: + fixes x :: "'a::archimedean_field" + assumes "0 < x" shows "\n. y < of_nat n * x" +proof - + obtain n where "y / x < of_nat n" using ex_less_of_nat .. + with `0 < x` have "y < of_nat n * x" by (simp add: pos_divide_less_eq) + then show ?thesis .. +qed + + +subsection {* Existence and uniqueness of floor function *} + +lemma exists_least_lemma: + assumes "\ P 0" and "\n. P n" + shows "\n. \ P n \ P (Suc n)" +proof - + from `\n. P n` have "P (Least P)" by (rule LeastI_ex) + with `\ P 0` obtain n where "Least P = Suc n" + by (cases "Least P") auto + then have "n < Least P" by simp + then have "\ P n" by (rule not_less_Least) + then have "\ P n \ P (Suc n)" + using `P (Least P)` `Least P = Suc n` by simp + then show ?thesis .. +qed + +lemma floor_exists: + fixes x :: "'a::archimedean_field" + shows "\z. of_int z \ x \ x < of_int (z + 1)" +proof (cases) + assume "0 \ x" + then have "\ x < of_nat 0" by simp + then have "\n. \ x < of_nat n \ x < of_nat (Suc n)" + using ex_less_of_nat by (rule exists_least_lemma) + then obtain n where "\ x < of_nat n \ x < of_nat (Suc n)" .. + then have "of_int (int n) \ x \ x < of_int (int n + 1)" by simp + then show ?thesis .. +next + assume "\ 0 \ x" + then have "\ - x \ of_nat 0" by simp + then have "\n. \ - x \ of_nat n \ - x \ of_nat (Suc n)" + using ex_le_of_nat by (rule exists_least_lemma) + then obtain n where "\ - x \ of_nat n \ - x \ of_nat (Suc n)" .. + then have "of_int (- int n - 1) \ x \ x < of_int (- int n - 1 + 1)" by simp + then show ?thesis .. +qed + +lemma floor_exists1: + fixes x :: "'a::archimedean_field" + shows "\!z. of_int z \ x \ x < of_int (z + 1)" +proof (rule ex_ex1I) + show "\z. of_int z \ x \ x < of_int (z + 1)" + by (rule floor_exists) +next + fix y z assume + "of_int y \ x \ x < of_int (y + 1)" + "of_int z \ x \ x < of_int (z + 1)" + then have + "of_int y \ x" "x < of_int (y + 1)" + "of_int z \ x" "x < of_int (z + 1)" + by simp_all + from le_less_trans [OF `of_int y \ x` `x < of_int (z + 1)`] + le_less_trans [OF `of_int z \ x` `x < of_int (y + 1)`] + show "y = z" by (simp del: of_int_add) +qed + + +subsection {* Floor function *} + +definition + floor :: "'a::archimedean_field \ int" where + [code del]: "floor x = (THE z. of_int z \ x \ x < of_int (z + 1))" + +notation (xsymbols) + floor ("\_\") + +notation (HTML output) + floor ("\_\") + +lemma floor_correct: "of_int (floor x) \ x \ x < of_int (floor x + 1)" + unfolding floor_def using floor_exists1 by (rule theI') + +lemma floor_unique: "\of_int z \ x; x < of_int z + 1\ \ floor x = z" + using floor_correct [of x] floor_exists1 [of x] by auto + +lemma of_int_floor_le: "of_int (floor x) \ x" + using floor_correct .. + +lemma le_floor_iff: "z \ floor x \ of_int z \ x" +proof + assume "z \ floor x" + then have "(of_int z :: 'a) \ of_int (floor x)" by simp + also have "of_int (floor x) \ x" by (rule of_int_floor_le) + finally show "of_int z \ x" . +next + assume "of_int z \ x" + also have "x < of_int (floor x + 1)" using floor_correct .. + finally show "z \ floor x" by (simp del: of_int_add) +qed + +lemma floor_less_iff: "floor x < z \ x < of_int z" + by (simp add: not_le [symmetric] le_floor_iff) + +lemma less_floor_iff: "z < floor x \ of_int z + 1 \ x" + using le_floor_iff [of "z + 1" x] by auto + +lemma floor_le_iff: "floor x \ z \ x < of_int z + 1" + by (simp add: not_less [symmetric] less_floor_iff) + +lemma floor_mono: assumes "x \ y" shows "floor x \ floor y" +proof - + have "of_int (floor x) \ x" by (rule of_int_floor_le) + also note `x \ y` + finally show ?thesis by (simp add: le_floor_iff) +qed + +lemma floor_less_cancel: "floor x < floor y \ x < y" + by (auto simp add: not_le [symmetric] floor_mono) + +lemma floor_of_int [simp]: "floor (of_int z) = z" + by (rule floor_unique) simp_all + +lemma floor_of_nat [simp]: "floor (of_nat n) = int n" + using floor_of_int [of "of_nat n"] by simp + +text {* Floor with numerals *} + +lemma floor_zero [simp]: "floor 0 = 0" + using floor_of_int [of 0] by simp + +lemma floor_one [simp]: "floor 1 = 1" + using floor_of_int [of 1] by simp + +lemma floor_number_of [simp]: "floor (number_of v) = number_of v" + using floor_of_int [of "number_of v"] by simp + +lemma zero_le_floor [simp]: "0 \ floor x \ 0 \ x" + by (simp add: le_floor_iff) + +lemma one_le_floor [simp]: "1 \ floor x \ 1 \ x" + by (simp add: le_floor_iff) + +lemma number_of_le_floor [simp]: "number_of v \ floor x \ number_of v \ x" + by (simp add: le_floor_iff) + +lemma zero_less_floor [simp]: "0 < floor x \ 1 \ x" + by (simp add: less_floor_iff) + +lemma one_less_floor [simp]: "1 < floor x \ 2 \ x" + by (simp add: less_floor_iff) + +lemma number_of_less_floor [simp]: + "number_of v < floor x \ number_of v + 1 \ x" + by (simp add: less_floor_iff) + +lemma floor_le_zero [simp]: "floor x \ 0 \ x < 1" + by (simp add: floor_le_iff) + +lemma floor_le_one [simp]: "floor x \ 1 \ x < 2" + by (simp add: floor_le_iff) + +lemma floor_le_number_of [simp]: + "floor x \ number_of v \ x < number_of v + 1" + by (simp add: floor_le_iff) + +lemma floor_less_zero [simp]: "floor x < 0 \ x < 0" + by (simp add: floor_less_iff) + +lemma floor_less_one [simp]: "floor x < 1 \ x < 1" + by (simp add: floor_less_iff) + +lemma floor_less_number_of [simp]: + "floor x < number_of v \ x < number_of v" + by (simp add: floor_less_iff) + +text {* Addition and subtraction of integers *} + +lemma floor_add_of_int [simp]: "floor (x + of_int z) = floor x + z" + using floor_correct [of x] by (simp add: floor_unique) + +lemma floor_add_number_of [simp]: + "floor (x + number_of v) = floor x + number_of v" + using floor_add_of_int [of x "number_of v"] by simp + +lemma floor_add_one [simp]: "floor (x + 1) = floor x + 1" + using floor_add_of_int [of x 1] by simp + +lemma floor_diff_of_int [simp]: "floor (x - of_int z) = floor x - z" + using floor_add_of_int [of x "- z"] by (simp add: algebra_simps) + +lemma floor_diff_number_of [simp]: + "floor (x - number_of v) = floor x - number_of v" + using floor_diff_of_int [of x "number_of v"] by simp + +lemma floor_diff_one [simp]: "floor (x - 1) = floor x - 1" + using floor_diff_of_int [of x 1] by simp + + +subsection {* Ceiling function *} + +definition + ceiling :: "'a::archimedean_field \ int" where + [code del]: "ceiling x = - floor (- x)" + +notation (xsymbols) + ceiling ("\_\") + +notation (HTML output) + ceiling ("\_\") + +lemma ceiling_correct: "of_int (ceiling x) - 1 < x \ x \ of_int (ceiling x)" + unfolding ceiling_def using floor_correct [of "- x"] by simp + +lemma ceiling_unique: "\of_int z - 1 < x; x \ of_int z\ \ ceiling x = z" + unfolding ceiling_def using floor_unique [of "- z" "- x"] by simp + +lemma le_of_int_ceiling: "x \ of_int (ceiling x)" + using ceiling_correct .. + +lemma ceiling_le_iff: "ceiling x \ z \ x \ of_int z" + unfolding ceiling_def using le_floor_iff [of "- z" "- x"] by auto + +lemma less_ceiling_iff: "z < ceiling x \ of_int z < x" + by (simp add: not_le [symmetric] ceiling_le_iff) + +lemma ceiling_less_iff: "ceiling x < z \ x \ of_int z - 1" + using ceiling_le_iff [of x "z - 1"] by simp + +lemma le_ceiling_iff: "z \ ceiling x \ of_int z - 1 < x" + by (simp add: not_less [symmetric] ceiling_less_iff) + +lemma ceiling_mono: "x \ y \ ceiling x \ ceiling y" + unfolding ceiling_def by (simp add: floor_mono) + +lemma ceiling_less_cancel: "ceiling x < ceiling y \ x < y" + by (auto simp add: not_le [symmetric] ceiling_mono) + +lemma ceiling_of_int [simp]: "ceiling (of_int z) = z" + by (rule ceiling_unique) simp_all + +lemma ceiling_of_nat [simp]: "ceiling (of_nat n) = int n" + using ceiling_of_int [of "of_nat n"] by simp + +text {* Ceiling with numerals *} + +lemma ceiling_zero [simp]: "ceiling 0 = 0" + using ceiling_of_int [of 0] by simp + +lemma ceiling_one [simp]: "ceiling 1 = 1" + using ceiling_of_int [of 1] by simp + +lemma ceiling_number_of [simp]: "ceiling (number_of v) = number_of v" + using ceiling_of_int [of "number_of v"] by simp + +lemma ceiling_le_zero [simp]: "ceiling x \ 0 \ x \ 0" + by (simp add: ceiling_le_iff) + +lemma ceiling_le_one [simp]: "ceiling x \ 1 \ x \ 1" + by (simp add: ceiling_le_iff) + +lemma ceiling_le_number_of [simp]: + "ceiling x \ number_of v \ x \ number_of v" + by (simp add: ceiling_le_iff) + +lemma ceiling_less_zero [simp]: "ceiling x < 0 \ x \ -1" + by (simp add: ceiling_less_iff) + +lemma ceiling_less_one [simp]: "ceiling x < 1 \ x \ 0" + by (simp add: ceiling_less_iff) + +lemma ceiling_less_number_of [simp]: + "ceiling x < number_of v \ x \ number_of v - 1" + by (simp add: ceiling_less_iff) + +lemma zero_le_ceiling [simp]: "0 \ ceiling x \ -1 < x" + by (simp add: le_ceiling_iff) + +lemma one_le_ceiling [simp]: "1 \ ceiling x \ 0 < x" + by (simp add: le_ceiling_iff) + +lemma number_of_le_ceiling [simp]: + "number_of v \ ceiling x\ number_of v - 1 < x" + by (simp add: le_ceiling_iff) + +lemma zero_less_ceiling [simp]: "0 < ceiling x \ 0 < x" + by (simp add: less_ceiling_iff) + +lemma one_less_ceiling [simp]: "1 < ceiling x \ 1 < x" + by (simp add: less_ceiling_iff) + +lemma number_of_less_ceiling [simp]: + "number_of v < ceiling x \ number_of v < x" + by (simp add: less_ceiling_iff) + +text {* Addition and subtraction of integers *} + +lemma ceiling_add_of_int [simp]: "ceiling (x + of_int z) = ceiling x + z" + using ceiling_correct [of x] by (simp add: ceiling_unique) + +lemma ceiling_add_number_of [simp]: + "ceiling (x + number_of v) = ceiling x + number_of v" + using ceiling_add_of_int [of x "number_of v"] by simp + +lemma ceiling_add_one [simp]: "ceiling (x + 1) = ceiling x + 1" + using ceiling_add_of_int [of x 1] by simp + +lemma ceiling_diff_of_int [simp]: "ceiling (x - of_int z) = ceiling x - z" + using ceiling_add_of_int [of x "- z"] by (simp add: algebra_simps) + +lemma ceiling_diff_number_of [simp]: + "ceiling (x - number_of v) = ceiling x - number_of v" + using ceiling_diff_of_int [of x "number_of v"] by simp + +lemma ceiling_diff_one [simp]: "ceiling (x - 1) = ceiling x - 1" + using ceiling_diff_of_int [of x 1] by simp + + +subsection {* Negation *} + +lemma floor_minus: "floor (- x) = - ceiling x" + unfolding ceiling_def by simp + +lemma ceiling_minus: "ceiling (- x) = - floor x" + unfolding ceiling_def by simp + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/AxClasses/Group.thy --- a/src/HOL/AxClasses/Group.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,124 +0,0 @@ -(* Title: HOL/AxClasses/Group.thy - ID: $Id$ - Author: Markus Wenzel, TU Muenchen -*) - -theory Group imports Main begin - -subsection {* Monoids and Groups *} - -consts - times :: "'a => 'a => 'a" (infixl "[*]" 70) - invers :: "'a => 'a" - one :: 'a - - -axclass monoid < type - assoc: "(x [*] y) [*] z = x [*] (y [*] z)" - left_unit: "one [*] x = x" - right_unit: "x [*] one = x" - -axclass semigroup < type - assoc: "(x [*] y) [*] z = x [*] (y [*] z)" - -axclass group < semigroup - left_unit: "one [*] x = x" - left_inverse: "invers x [*] x = one" - -axclass agroup < group - commute: "x [*] y = y [*] x" - - -subsection {* Abstract reasoning *} - -theorem group_right_inverse: "x [*] invers x = (one::'a::group)" -proof - - have "x [*] invers x = one [*] (x [*] invers x)" - by (simp only: group_class.left_unit) - also have "... = one [*] x [*] invers x" - by (simp only: semigroup_class.assoc) - also have "... = invers (invers x) [*] invers x [*] x [*] invers x" - by (simp only: group_class.left_inverse) - also have "... = invers (invers x) [*] (invers x [*] x) [*] invers x" - by (simp only: semigroup_class.assoc) - also have "... = invers (invers x) [*] one [*] invers x" - by (simp only: group_class.left_inverse) - also have "... = invers (invers x) [*] (one [*] invers x)" - by (simp only: semigroup_class.assoc) - also have "... = invers (invers x) [*] invers x" - by (simp only: group_class.left_unit) - also have "... = one" - by (simp only: group_class.left_inverse) - finally show ?thesis . -qed - -theorem group_right_unit: "x [*] one = (x::'a::group)" -proof - - have "x [*] one = x [*] (invers x [*] x)" - by (simp only: group_class.left_inverse) - also have "... = x [*] invers x [*] x" - by (simp only: semigroup_class.assoc) - also have "... = one [*] x" - by (simp only: group_right_inverse) - also have "... = x" - by (simp only: group_class.left_unit) - finally show ?thesis . -qed - - -subsection {* Abstract instantiation *} - -instance monoid < semigroup -proof intro_classes - fix x y z :: "'a::monoid" - show "x [*] y [*] z = x [*] (y [*] z)" - by (rule monoid_class.assoc) -qed - -instance group < monoid -proof intro_classes - fix x y z :: "'a::group" - show "x [*] y [*] z = x [*] (y [*] z)" - by (rule semigroup_class.assoc) - show "one [*] x = x" - by (rule group_class.left_unit) - show "x [*] one = x" - by (rule group_right_unit) -qed - - -subsection {* Concrete instantiation *} - -defs (overloaded) - times_bool_def: "x [*] y == x ~= (y::bool)" - inverse_bool_def: "invers x == x::bool" - unit_bool_def: "one == False" - -instance bool :: agroup -proof (intro_classes, - unfold times_bool_def inverse_bool_def unit_bool_def) - fix x y z - show "((x ~= y) ~= z) = (x ~= (y ~= z))" by blast - show "(False ~= x) = x" by blast - show "(x ~= x) = False" by blast - show "(x ~= y) = (y ~= x)" by blast -qed - - -subsection {* Lifting and Functors *} - -defs (overloaded) - times_prod_def: "p [*] q == (fst p [*] fst q, snd p [*] snd q)" - -instance * :: (semigroup, semigroup) semigroup -proof (intro_classes, unfold times_prod_def) - fix p q r :: "'a::semigroup * 'b::semigroup" - show - "(fst (fst p [*] fst q, snd p [*] snd q) [*] fst r, - snd (fst p [*] fst q, snd p [*] snd q) [*] snd r) = - (fst p [*] fst (fst q [*] fst r, snd q [*] snd r), - snd p [*] snd (fst q [*] fst r, snd q [*] snd r))" - by (simp add: semigroup_class.assoc) -qed - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/AxClasses/Lattice/OrdInsts.thy --- a/src/HOL/AxClasses/Lattice/OrdInsts.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,43 +0,0 @@ -(* Title: OrdInsts.thy - ID: $Id$ - Author: Markus Wenzel, TU Muenchen - -Some order instantiations. -*) - -OrdInsts = OrdDefs + - - -(* binary / general products of quasi_orders / orders *) - -instance - "*" :: (quasi_order, quasi_order) quasi_order (le_prod_refl, le_prod_trans) - -instance - "*" :: (partial_order, partial_order) partial_order (le_prod_antisym) - - -instance - fun :: (term, quasi_order) quasi_order (le_fun_refl, le_fun_trans) - -instance - fun :: (term, partial_order) partial_order (le_fun_antisym) - - -(* duals of quasi orders / partial orders / linear orders *) - -instance - dual :: (quasi_order) quasi_order (le_dual_refl, le_dual_trans) - -instance - dual :: (partial_order) partial_order (le_dual_antisym) - - -(*FIXME: had to be moved to LatInsts.thy due to some unpleasant - 'feature' in Pure/type.ML - -instance - dual :: (linear_order) linear_order (le_dual_lin) -*) - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/AxClasses/Product.thy --- a/src/HOL/AxClasses/Product.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,20 +0,0 @@ -(* Title: HOL/AxClasses/Product.thy - ID: $Id$ - Author: Markus Wenzel, TU Muenchen -*) - -theory Product imports Main begin - -axclass product < type - -consts - product :: "'a::product => 'a => 'a" (infixl "[*]" 70) - - -instance bool :: product - by intro_classes - -defs (overloaded) - product_bool_def: "x [*] y == x & y" - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/AxClasses/README.html --- a/src/HOL/AxClasses/README.html Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,20 +0,0 @@ - - - - - - - - - HOL/AxClasses - - - -

HOL/AxClasses

- -These are the HOL examples of the tutorial
Using Axiomatic Type -Classes in Isabelle. See also FOL/ex/NatClass for the natural -number example. - - diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/AxClasses/ROOT.ML --- a/src/HOL/AxClasses/ROOT.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,3 +0,0 @@ -(* $Id$ *) - -use_thys ["Semigroups", "Group", "Product"]; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/AxClasses/Semigroups.thy --- a/src/HOL/AxClasses/Semigroups.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,21 +0,0 @@ -(* Title: HOL/AxClasses/Semigroups.thy - ID: $Id$ - Author: Markus Wenzel, TU Muenchen -*) - -theory Semigroups imports Main begin - -consts - times :: "'a => 'a => 'a" (infixl "[*]" 70) - -axclass semigroup < type - assoc: "(x [*] y) [*] z = x [*] (y [*] z)" - - -consts - plus :: "'a => 'a => 'a" (infixl "[+]" 70) - -axclass plus_semigroup < type - assoc: "(x [+] y) [+] z = x [+] (y [+] z)" - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Bali/Basis.thy --- a/src/HOL/Bali/Basis.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Bali/Basis.thy Wed Mar 04 11:05:29 2009 +0100 @@ -251,8 +251,8 @@ Oex :: "[pttrn, 'a option, bool] => bool" ("(3\_\_:/ _)" [0,0,10] 10) translations - "! x:A: P" == "! x:o2s A. P" - "? x:A: P" == "? x:o2s A. P" + "! x:A: P" == "! x:CONST Option.set A. P" + "? x:A: P" == "? x:CONST Option.set A. P" section "Special map update" diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Bali/Conform.thy --- a/src/HOL/Bali/Conform.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Bali/Conform.thy Wed Mar 04 11:05:29 2009 +0100 @@ -102,7 +102,7 @@ constdefs conf :: "prog \ st \ val \ ty \ bool" ("_,_\_\\_" [71,71,71,71] 70) - "G,s\v\\T \ \T'\typeof (\a. option_map obj_ty (heap s a)) v:G\T'\T" + "G,s\v\\T \ \T'\typeof (\a. Option.map obj_ty (heap s a)) v:G\T'\T" lemma conf_cong [simp]: "G,set_locals l s\v\\T = G,s\v\\T" by (auto simp: conf_def) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Bali/Decl.thy --- a/src/HOL/Bali/Decl.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Bali/Decl.thy Wed Mar 04 11:05:29 2009 +0100 @@ -801,7 +801,7 @@ "imethds G I \ iface_rec (G,I) (\I i ts. (Un_tables ts) \\ - (o2s \ table_of (map (\(s,m). (s,I,m)) (imethods i))))" + (Option.set \ table_of (map (\(s,m). (s,I,m)) (imethods i))))" diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Bali/DeclConcepts.thy --- a/src/HOL/Bali/DeclConcepts.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Bali/DeclConcepts.thy Wed Mar 04 11:05:29 2009 +0100 @@ -1385,7 +1385,7 @@ "imethds G I \ iface_rec (G,I) (\I i ts. (Un_tables ts) \\ - (o2s \ table_of (map (\(s,m). (s,I,m)) (imethods i))))" + (Option.set \ table_of (map (\(s,m). (s,I,m)) (imethods i))))" text {* methods of an interface, with overriding and inheritance, cf. 9.2 *} constdefs @@ -1528,7 +1528,7 @@ lemma imethds_rec: "\iface G I = Some i; ws_prog G\ \ imethds G I = Un_tables ((\J. imethds G J)`set (isuperIfs i)) \\ - (o2s \ table_of (map (\(s,mh). (s,I,mh)) (imethods i)))" + (Option.set \ table_of (map (\(s,mh). (s,I,mh)) (imethods i)))" apply (unfold imethds_def) apply (rule iface_rec [THEN trans]) apply auto diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Bali/Example.thy --- a/src/HOL/Bali/Example.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Bali/Example.thy Wed Mar 04 11:05:29 2009 +0100 @@ -458,7 +458,7 @@ lemmas methd_rec' = methd_rec [OF _ ws_tprg] lemma imethds_HasFoo [simp]: - "imethds tprg HasFoo = o2s \ empty(foo_sig\(HasFoo, foo_mhead))" + "imethds tprg HasFoo = Option.set \ empty(foo_sig\(HasFoo, foo_mhead))" apply (rule trans) apply (rule imethds_rec') apply (auto simp add: HasFooInt_def) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Bali/State.thy --- a/src/HOL/Bali/State.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Bali/State.thy Wed Mar 04 11:05:29 2009 +0100 @@ -146,7 +146,7 @@ fields_table:: "prog \ qtname \ (fspec \ field \ bool) \ (fspec, ty) table" "fields_table G C P - \ option_map type \ table_of (filter (split P) (DeclConcepts.fields G C))" + \ Option.map type \ table_of (filter (split P) (DeclConcepts.fields G C))" lemma fields_table_SomeI: "\table_of (DeclConcepts.fields G C) n = Some f; P n f\ @@ -258,8 +258,8 @@ lookup_obj :: "st \ val \ obj" translations - "val_this s" == "the (locals s This)" - "lookup_obj s a'" == "the (heap s (the_Addr a'))" + "val_this s" == "CONST the (locals s This)" + "lookup_obj s a'" == "CONST the (heap s (the_Addr a'))" subsection "memory allocation" @@ -290,7 +290,7 @@ init_vals :: "('a, ty) table \ ('a, val) table" translations - "init_vals vs" == "CONST option_map default_val \ vs" + "init_vals vs" == "CONST Option.map default_val \ vs" lemma init_arr_comps_base [simp]: "init_vals (arr_comps T 0) = empty" apply (unfold arr_comps_def in_bounds_def) @@ -315,12 +315,12 @@ lupd :: "lname \ val \ st \ st" ("lupd'(_\_')"[10,10]1000) "lupd vn v \ st_case (\g l. st g (l(vn\v)))" - upd_gobj :: "oref \ vn \ val \ st \ st" + upd_gobj :: "oref \ vn \ val \ st \ st" "upd_gobj r n v \ st_case (\g l. st (chg_map (upd_obj n v) r g) l)" set_locals :: "locals \ st \ st" "set_locals l \ st_case (\g l'. st g l)" - + init_obj :: "prog \ obj_tag \ oref \ st \ st" "init_obj G oi r \ gupd(r\\tag=oi, values=init_vals (var_tys G oi r)\)" diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Bali/Table.thy --- a/src/HOL/Bali/Table.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Bali/Table.thy Wed Mar 04 11:05:29 2009 +0100 @@ -194,7 +194,7 @@ done lemma Ball_set_tableD: - "\(\ (x,y)\ set l. P x y); x \ o2s (table_of l xa)\ \ P xa x" + "\(\ (x,y)\ set l. P x y); x \ Option.set (table_of l xa)\ \ P xa x" apply (frule Ball_set_table) by auto diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Bali/WellForm.thy --- a/src/HOL/Bali/WellForm.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Bali/WellForm.thy Wed Mar 04 11:05:29 2009 +0100 @@ -236,7 +236,7 @@ under (\ new old. accmodi old \ Private) entails (\new old. G\resTy new\resTy old \ is_static new = is_static old)) \ - (o2s \ table_of (imethods i) + (Option.set \ table_of (imethods i) hidings Un_tables((\J.(imethds G J))`set (isuperIfs i)) entails (\new old. G\resTy new\resTy old))" @@ -248,7 +248,7 @@ lemma wf_idecl_hidings: "wf_idecl G (I, i) \ - (\s. o2s (table_of (imethods i) s)) + (\s. Option.set (table_of (imethods i) s)) hidings Un_tables ((\J. imethds G J) ` set (isuperIfs i)) entails \new old. G\resTy new\resTy old" apply (unfold wf_idecl_def o_def) @@ -751,7 +751,7 @@ show "\is_static im \ accmodi im = Public" proof - let ?inherited = "Un_tables (imethds G ` set (isuperIfs i))" - let ?new = "(o2s \ table_of (map (\(s, mh). (s, I, mh)) (imethods i)))" + let ?new = "(Option.set \ table_of (map (\(s, mh). (s, I, mh)) (imethods i)))" from if_I wf im have imethds:"im \ (?inherited \\ ?new) sig" by (simp add: imethds_rec) from wf if_I have @@ -1783,7 +1783,7 @@ by (blast dest: subint1D) let ?newMethods - = "(o2s \ table_of (map (\(sig, mh). (sig, I, mh)) (imethods i)))" + = "(Option.set \ table_of (map (\(sig, mh). (sig, I, mh)) (imethods i)))" show "?Concl I" proof (cases "?newMethods sig = {}") case True @@ -1864,7 +1864,7 @@ apply (drule (1) wf_prog_idecl) apply (frule (3) imethds_wf_mhead [OF _ _ wf_idecl_supD [THEN conjunct1 [THEN is_acc_ifaceD [THEN conjunct1]]]]) -apply (case_tac "(o2s \ table_of (map (\(s, mh). (s, y, mh)) (imethods i))) +apply (case_tac "(Option.set \ table_of (map (\(s, mh). (s, y, mh)) (imethods i))) sig ={}") apply force diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Bali/WellType.thy --- a/src/HOL/Bali/WellType.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Bali/WellType.thy Wed Mar 04 11:05:29 2009 +0100 @@ -87,11 +87,11 @@ defs cmheads_def: "cmheads G S C - \ \sig. (\(Cls,mthd). (ClassT Cls,(mhead mthd))) ` o2s (accmethd G S C sig)" + \ \sig. (\(Cls,mthd). (ClassT Cls,(mhead mthd))) ` Option.set (accmethd G S C sig)" Objectmheads_def: "Objectmheads G S \ \sig. (\(Cls,mthd). (ClassT Cls,(mhead mthd))) - ` o2s (filter_tab (\sig m. accmodi m \ Private) (accmethd G S Object) sig)" + ` Option.set (filter_tab (\sig m. accmodi m \ Private) (accmethd G S Object) sig)" accObjectmheads_def: "accObjectmheads G S T \ if G\RefT T accessible_in (pid S) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Datatype.thy --- a/src/HOL/Datatype.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Datatype.thy Wed Mar 04 11:05:29 2009 +0100 @@ -576,122 +576,4 @@ hide (open) const Suml Sumr Projl Projr - -subsection {* The option datatype *} - -datatype 'a option = None | Some 'a - -lemma not_None_eq [iff]: "(x ~= None) = (EX y. x = Some y)" - by (induct x) auto - -lemma not_Some_eq [iff]: "(ALL y. x ~= Some y) = (x = None)" - by (induct x) auto - -text{*Although it may appear that both of these equalities are helpful -only when applied to assumptions, in practice it seems better to give -them the uniform iff attribute. *} - -lemma option_caseE: - assumes c: "(case x of None => P | Some y => Q y)" - obtains - (None) "x = None" and P - | (Some) y where "x = Some y" and "Q y" - using c by (cases x) simp_all - -lemma insert_None_conv_UNIV: "insert None (range Some) = UNIV" - by (rule set_ext, case_tac x) auto - -lemma inj_Some [simp]: "inj_on Some A" - by (rule inj_onI) simp - - -subsubsection {* Operations *} - -consts - the :: "'a option => 'a" -primrec - "the (Some x) = x" - -consts - o2s :: "'a option => 'a set" -primrec - "o2s None = {}" - "o2s (Some x) = {x}" - -lemma ospec [dest]: "(ALL x:o2s A. P x) ==> A = Some x ==> P x" - by simp - -declaration {* fn _ => - Classical.map_cs (fn cs => cs addSD2 ("ospec", thm "ospec")) -*} - -lemma elem_o2s [iff]: "(x : o2s xo) = (xo = Some x)" - by (cases xo) auto - -lemma o2s_empty_eq [simp]: "(o2s xo = {}) = (xo = None)" - by (cases xo) auto - -definition - option_map :: "('a \ 'b) \ 'a option \ 'b option" -where - [code del]: "option_map = (%f y. case y of None => None | Some x => Some (f x))" - -lemma option_map_None [simp, code]: "option_map f None = None" - by (simp add: option_map_def) - -lemma option_map_Some [simp, code]: "option_map f (Some x) = Some (f x)" - by (simp add: option_map_def) - -lemma option_map_is_None [iff]: - "(option_map f opt = None) = (opt = None)" - by (simp add: option_map_def split add: option.split) - -lemma option_map_eq_Some [iff]: - "(option_map f xo = Some y) = (EX z. xo = Some z & f z = y)" - by (simp add: option_map_def split add: option.split) - -lemma option_map_comp: - "option_map f (option_map g opt) = option_map (f o g) opt" - by (simp add: option_map_def split add: option.split) - -lemma option_map_o_sum_case [simp]: - "option_map f o sum_case g h = sum_case (option_map f o g) (option_map f o h)" - by (rule ext) (simp split: sum.split) - - -subsubsection {* Code generator setup *} - -definition - is_none :: "'a option \ bool" where - is_none_none [code post, symmetric, code inline]: "is_none x \ x = None" - -lemma is_none_code [code]: - shows "is_none None \ True" - and "is_none (Some x) \ False" - unfolding is_none_none [symmetric] by simp_all - -hide (open) const is_none - -code_type option - (SML "_ option") - (OCaml "_ option") - (Haskell "Maybe _") - -code_const None and Some - (SML "NONE" and "SOME") - (OCaml "None" and "Some _") - (Haskell "Nothing" and "Just") - -code_instance option :: eq - (Haskell -) - -code_const "eq_class.eq \ 'a\eq option \ 'a option \ bool" - (Haskell infixl 4 "==") - -code_reserved SML - option NONE SOME - -code_reserved OCaml - option None Some - end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Decision_Procs/MIR.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Decision_Procs/cooper_tac.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Decision_Procs/ferrack_tac.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Decision_Procs/mir_tac.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Deriv.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Divides.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Extraction.thy --- a/src/HOL/Extraction.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Extraction.thy Wed Mar 04 11:05:29 2009 +0100 @@ -6,7 +6,7 @@ header {* Program extraction for HOL *} theory Extraction -imports Datatype +imports Option uses "Tools/rewrite_hol_proof.ML" begin diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Fact.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Finite_Set.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/FrechetDeriv.thy --- a/src/HOL/FrechetDeriv.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,503 +0,0 @@ -(* Title : FrechetDeriv.thy - ID : $Id$ - Author : Brian Huffman -*) - -header {* Frechet Derivative *} - -theory FrechetDeriv -imports Lim -begin - -definition - fderiv :: - "['a::real_normed_vector \ 'b::real_normed_vector, 'a, 'a \ 'b] \ bool" - -- {* Frechet derivative: D is derivative of function f at x *} - ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where - "FDERIV f x :> D = (bounded_linear D \ - (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0)" - -lemma FDERIV_I: - "\bounded_linear D; (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0\ - \ FDERIV f x :> D" -by (simp add: fderiv_def) - -lemma FDERIV_D: - "FDERIV f x :> D \ (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0" -by (simp add: fderiv_def) - -lemma FDERIV_bounded_linear: "FDERIV f x :> D \ bounded_linear D" -by (simp add: fderiv_def) - -lemma bounded_linear_zero: - "bounded_linear (\x::'a::real_normed_vector. 0::'b::real_normed_vector)" -proof - show "(0::'b) = 0 + 0" by simp - fix r show "(0::'b) = scaleR r 0" by simp - have "\x::'a. norm (0::'b) \ norm x * 0" by simp - thus "\K. \x::'a. norm (0::'b) \ norm x * K" .. -qed - -lemma FDERIV_const: "FDERIV (\x. k) x :> (\h. 0)" -by (simp add: fderiv_def bounded_linear_zero) - -lemma bounded_linear_ident: - "bounded_linear (\x::'a::real_normed_vector. x)" -proof - fix x y :: 'a show "x + y = x + y" by simp - fix r and x :: 'a show "scaleR r x = scaleR r x" by simp - have "\x::'a. norm x \ norm x * 1" by simp - thus "\K. \x::'a. norm x \ norm x * K" .. -qed - -lemma FDERIV_ident: "FDERIV (\x. x) x :> (\h. h)" -by (simp add: fderiv_def bounded_linear_ident) - -subsection {* Addition *} - -lemma add_diff_add: - fixes a b c d :: "'a::ab_group_add" - shows "(a + c) - (b + d) = (a - b) + (c - d)" -by simp - -lemma bounded_linear_add: - assumes "bounded_linear f" - assumes "bounded_linear g" - shows "bounded_linear (\x. f x + g x)" -proof - - interpret f: bounded_linear f by fact - interpret g: bounded_linear g by fact - show ?thesis apply (unfold_locales) - apply (simp only: f.add g.add add_ac) - apply (simp only: f.scaleR g.scaleR scaleR_right_distrib) - apply (rule f.pos_bounded [THEN exE], rename_tac Kf) - apply (rule g.pos_bounded [THEN exE], rename_tac Kg) - apply (rule_tac x="Kf + Kg" in exI, safe) - apply (subst right_distrib) - apply (rule order_trans [OF norm_triangle_ineq]) - apply (rule add_mono, erule spec, erule spec) - done -qed - -lemma norm_ratio_ineq: - fixes x y :: "'a::real_normed_vector" - fixes h :: "'b::real_normed_vector" - shows "norm (x + y) / norm h \ norm x / norm h + norm y / norm h" -apply (rule ord_le_eq_trans) -apply (rule divide_right_mono) -apply (rule norm_triangle_ineq) -apply (rule norm_ge_zero) -apply (rule add_divide_distrib) -done - -lemma FDERIV_add: - assumes f: "FDERIV f x :> F" - assumes g: "FDERIV g x :> G" - shows "FDERIV (\x. f x + g x) x :> (\h. F h + G h)" -proof (rule FDERIV_I) - show "bounded_linear (\h. F h + G h)" - apply (rule bounded_linear_add) - apply (rule FDERIV_bounded_linear [OF f]) - apply (rule FDERIV_bounded_linear [OF g]) - done -next - have f': "(\h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0" - using f by (rule FDERIV_D) - have g': "(\h. norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0" - using g by (rule FDERIV_D) - from f' g' - have "(\h. norm (f (x + h) - f x - F h) / norm h - + norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0" - by (rule LIM_add_zero) - thus "(\h. norm (f (x + h) + g (x + h) - (f x + g x) - (F h + G h)) - / norm h) -- 0 --> 0" - apply (rule real_LIM_sandwich_zero) - apply (simp add: divide_nonneg_pos) - apply (simp only: add_diff_add) - apply (rule norm_ratio_ineq) - done -qed - -subsection {* Subtraction *} - -lemma bounded_linear_minus: - assumes "bounded_linear f" - shows "bounded_linear (\x. - f x)" -proof - - interpret f: bounded_linear f by fact - show ?thesis apply (unfold_locales) - apply (simp add: f.add) - apply (simp add: f.scaleR) - apply (simp add: f.bounded) - done -qed - -lemma FDERIV_minus: - "FDERIV f x :> F \ FDERIV (\x. - f x) x :> (\h. - F h)" -apply (rule FDERIV_I) -apply (rule bounded_linear_minus) -apply (erule FDERIV_bounded_linear) -apply (simp only: fderiv_def minus_diff_minus norm_minus_cancel) -done - -lemma FDERIV_diff: - "\FDERIV f x :> F; FDERIV g x :> G\ - \ FDERIV (\x. f x - g x) x :> (\h. F h - G h)" -by (simp only: diff_minus FDERIV_add FDERIV_minus) - -subsection {* Continuity *} - -lemma FDERIV_isCont: - assumes f: "FDERIV f x :> F" - shows "isCont f x" -proof - - from f interpret F: bounded_linear "F" by (rule FDERIV_bounded_linear) - have "(\h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0" - by (rule FDERIV_D [OF f]) - hence "(\h. norm (f (x + h) - f x - F h) / norm h * norm h) -- 0 --> 0" - by (intro LIM_mult_zero LIM_norm_zero LIM_ident) - hence "(\h. norm (f (x + h) - f x - F h)) -- 0 --> 0" - by (simp cong: LIM_cong) - hence "(\h. f (x + h) - f x - F h) -- 0 --> 0" - by (rule LIM_norm_zero_cancel) - hence "(\h. f (x + h) - f x - F h + F h) -- 0 --> 0" - by (intro LIM_add_zero F.LIM_zero LIM_ident) - hence "(\h. f (x + h) - f x) -- 0 --> 0" - by simp - thus "isCont f x" - unfolding isCont_iff by (rule LIM_zero_cancel) -qed - -subsection {* Composition *} - -lemma real_divide_cancel_lemma: - fixes a b c :: real - shows "(b = 0 \ a = 0) \ (a / b) * (b / c) = a / c" -by simp - -lemma bounded_linear_compose: - assumes "bounded_linear f" - assumes "bounded_linear g" - shows "bounded_linear (\x. f (g x))" -proof - - interpret f: bounded_linear f by fact - interpret g: bounded_linear g by fact - show ?thesis proof (unfold_locales) - fix x y show "f (g (x + y)) = f (g x) + f (g y)" - by (simp only: f.add g.add) - next - fix r x show "f (g (scaleR r x)) = scaleR r (f (g x))" - by (simp only: f.scaleR g.scaleR) - next - from f.pos_bounded - obtain Kf where f: "\x. norm (f x) \ norm x * Kf" and Kf: "0 < Kf" by fast - from g.pos_bounded - obtain Kg where g: "\x. norm (g x) \ norm x * Kg" by fast - show "\K. \x. norm (f (g x)) \ norm x * K" - proof (intro exI allI) - fix x - have "norm (f (g x)) \ norm (g x) * Kf" - using f . - also have "\ \ (norm x * Kg) * Kf" - using g Kf [THEN order_less_imp_le] by (rule mult_right_mono) - also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)" - by (rule mult_assoc) - finally show "norm (f (g x)) \ norm x * (Kg * Kf)" . - qed - qed -qed - -lemma FDERIV_compose: - fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" - fixes g :: "'b::real_normed_vector \ 'c::real_normed_vector" - assumes f: "FDERIV f x :> F" - assumes g: "FDERIV g (f x) :> G" - shows "FDERIV (\x. g (f x)) x :> (\h. G (F h))" -proof (rule FDERIV_I) - from FDERIV_bounded_linear [OF g] FDERIV_bounded_linear [OF f] - show "bounded_linear (\h. G (F h))" - by (rule bounded_linear_compose) -next - let ?Rf = "\h. f (x + h) - f x - F h" - let ?Rg = "\k. g (f x + k) - g (f x) - G k" - let ?k = "\h. f (x + h) - f x" - let ?Nf = "\h. norm (?Rf h) / norm h" - let ?Ng = "\h. norm (?Rg (?k h)) / norm (?k h)" - from f interpret F!: bounded_linear "F" by (rule FDERIV_bounded_linear) - from g interpret G!: bounded_linear "G" by (rule FDERIV_bounded_linear) - from F.bounded obtain kF where kF: "\x. norm (F x) \ norm x * kF" by fast - from G.bounded obtain kG where kG: "\x. norm (G x) \ norm x * kG" by fast - - let ?fun2 = "\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)" - - show "(\h. norm (g (f (x + h)) - g (f x) - G (F h)) / norm h) -- 0 --> 0" - proof (rule real_LIM_sandwich_zero) - have Nf: "?Nf -- 0 --> 0" - using FDERIV_D [OF f] . - - have Ng1: "isCont (\k. norm (?Rg k) / norm k) 0" - by (simp add: isCont_def FDERIV_D [OF g]) - have Ng2: "?k -- 0 --> 0" - apply (rule LIM_zero) - apply (fold isCont_iff) - apply (rule FDERIV_isCont [OF f]) - done - have Ng: "?Ng -- 0 --> 0" - using isCont_LIM_compose [OF Ng1 Ng2] by simp - - have "(\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) - -- 0 --> 0 * kG + 0 * (0 + kF)" - by (intro LIM_add LIM_mult LIM_const Nf Ng) - thus "(\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) -- 0 --> 0" - by simp - next - fix h::'a assume h: "h \ 0" - thus "0 \ norm (g (f (x + h)) - g (f x) - G (F h)) / norm h" - by (simp add: divide_nonneg_pos) - next - fix h::'a assume h: "h \ 0" - have "g (f (x + h)) - g (f x) - G (F h) = G (?Rf h) + ?Rg (?k h)" - by (simp add: G.diff) - hence "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h - = norm (G (?Rf h) + ?Rg (?k h)) / norm h" - by (rule arg_cong) - also have "\ \ norm (G (?Rf h)) / norm h + norm (?Rg (?k h)) / norm h" - by (rule norm_ratio_ineq) - also have "\ \ ?Nf h * kG + ?Ng h * (?Nf h + kF)" - proof (rule add_mono) - show "norm (G (?Rf h)) / norm h \ ?Nf h * kG" - apply (rule ord_le_eq_trans) - apply (rule divide_right_mono [OF kG norm_ge_zero]) - apply simp - done - next - have "norm (?Rg (?k h)) / norm h = ?Ng h * (norm (?k h) / norm h)" - apply (rule real_divide_cancel_lemma [symmetric]) - apply (simp add: G.zero) - done - also have "\ \ ?Ng h * (?Nf h + kF)" - proof (rule mult_left_mono) - have "norm (?k h) / norm h = norm (?Rf h + F h) / norm h" - by simp - also have "\ \ ?Nf h + norm (F h) / norm h" - by (rule norm_ratio_ineq) - also have "\ \ ?Nf h + kF" - apply (rule add_left_mono) - apply (subst pos_divide_le_eq, simp add: h) - apply (subst mult_commute) - apply (rule kF) - done - finally show "norm (?k h) / norm h \ ?Nf h + kF" . - next - show "0 \ ?Ng h" - apply (case_tac "f (x + h) - f x = 0", simp) - apply (rule divide_nonneg_pos [OF norm_ge_zero]) - apply simp - done - qed - finally show "norm (?Rg (?k h)) / norm h \ ?Ng h * (?Nf h + kF)" . - qed - finally show "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h - \ ?Nf h * kG + ?Ng h * (?Nf h + kF)" . - qed -qed - -subsection {* Product Rule *} - -lemma (in bounded_bilinear) FDERIV_lemma: - "a' ** b' - a ** b - (a ** B + A ** b) - = a ** (b' - b - B) + (a' - a - A) ** b' + A ** (b' - b)" -by (simp add: diff_left diff_right) - -lemma (in bounded_bilinear) FDERIV: - fixes x :: "'d::real_normed_vector" - assumes f: "FDERIV f x :> F" - assumes g: "FDERIV g x :> G" - shows "FDERIV (\x. f x ** g x) x :> (\h. f x ** G h + F h ** g x)" -proof (rule FDERIV_I) - show "bounded_linear (\h. f x ** G h + F h ** g x)" - apply (rule bounded_linear_add) - apply (rule bounded_linear_compose [OF bounded_linear_right]) - apply (rule FDERIV_bounded_linear [OF g]) - apply (rule bounded_linear_compose [OF bounded_linear_left]) - apply (rule FDERIV_bounded_linear [OF f]) - done -next - from bounded_linear.bounded [OF FDERIV_bounded_linear [OF f]] - obtain KF where norm_F: "\x. norm (F x) \ norm x * KF" by fast - - from pos_bounded obtain K where K: "0 < K" and norm_prod: - "\a b. norm (a ** b) \ norm a * norm b * K" by fast - - let ?Rf = "\h. f (x + h) - f x - F h" - let ?Rg = "\h. g (x + h) - g x - G h" - - let ?fun1 = "\h. - norm (f x ** ?Rg h + ?Rf h ** g (x + h) + F h ** (g (x + h) - g x)) / - norm h" - - let ?fun2 = "\h. - norm (f x) * (norm (?Rg h) / norm h) * K + - norm (?Rf h) / norm h * norm (g (x + h)) * K + - KF * norm (g (x + h) - g x) * K" - - have "?fun1 -- 0 --> 0" - proof (rule real_LIM_sandwich_zero) - from f g isCont_iff [THEN iffD1, OF FDERIV_isCont [OF g]] - have "?fun2 -- 0 --> - norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K" - by (intro LIM_add LIM_mult LIM_const LIM_norm LIM_zero FDERIV_D) - thus "?fun2 -- 0 --> 0" - by simp - next - fix h::'d assume "h \ 0" - thus "0 \ ?fun1 h" - by (simp add: divide_nonneg_pos) - next - fix h::'d assume "h \ 0" - have "?fun1 h \ (norm (f x) * norm (?Rg h) * K + - norm (?Rf h) * norm (g (x + h)) * K + - norm h * KF * norm (g (x + h) - g x) * K) / norm h" - by (intro - divide_right_mono mult_mono' - order_trans [OF norm_triangle_ineq add_mono] - order_trans [OF norm_prod mult_right_mono] - mult_nonneg_nonneg order_refl norm_ge_zero norm_F - K [THEN order_less_imp_le] - ) - also have "\ = ?fun2 h" - by (simp add: add_divide_distrib) - finally show "?fun1 h \ ?fun2 h" . - qed - thus "(\h. - norm (f (x + h) ** g (x + h) - f x ** g x - (f x ** G h + F h ** g x)) - / norm h) -- 0 --> 0" - by (simp only: FDERIV_lemma) -qed - -lemmas FDERIV_mult = mult.FDERIV - -lemmas FDERIV_scaleR = scaleR.FDERIV - - -subsection {* Powers *} - -lemma FDERIV_power_Suc: - fixes x :: "'a::{real_normed_algebra,recpower,comm_ring_1}" - shows "FDERIV (\x. x ^ Suc n) x :> (\h. (1 + of_nat n) * x ^ n * h)" - apply (induct n) - apply (simp add: power_Suc FDERIV_ident) - apply (drule FDERIV_mult [OF FDERIV_ident]) - apply (simp only: of_nat_Suc left_distrib mult_1_left) - apply (simp only: power_Suc right_distrib add_ac mult_ac) -done - -lemma FDERIV_power: - fixes x :: "'a::{real_normed_algebra,recpower,comm_ring_1}" - shows "FDERIV (\x. x ^ n) x :> (\h. of_nat n * x ^ (n - 1) * h)" - apply (cases n) - apply (simp add: FDERIV_const) - apply (simp add: FDERIV_power_Suc) - done - - -subsection {* Inverse *} - -lemma inverse_diff_inverse: - "\(a::'a::division_ring) \ 0; b \ 0\ - \ inverse a - inverse b = - (inverse a * (a - b) * inverse b)" -by (simp add: right_diff_distrib left_diff_distrib mult_assoc) - -lemmas bounded_linear_mult_const = - mult.bounded_linear_left [THEN bounded_linear_compose] - -lemmas bounded_linear_const_mult = - mult.bounded_linear_right [THEN bounded_linear_compose] - -lemma FDERIV_inverse: - fixes x :: "'a::real_normed_div_algebra" - assumes x: "x \ 0" - shows "FDERIV inverse x :> (\h. - (inverse x * h * inverse x))" - (is "FDERIV ?inv _ :> _") -proof (rule FDERIV_I) - show "bounded_linear (\h. - (?inv x * h * ?inv x))" - apply (rule bounded_linear_minus) - apply (rule bounded_linear_mult_const) - apply (rule bounded_linear_const_mult) - apply (rule bounded_linear_ident) - done -next - show "(\h. norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h) - -- 0 --> 0" - proof (rule LIM_equal2) - show "0 < norm x" using x by simp - next - fix h::'a - assume 1: "h \ 0" - assume "norm (h - 0) < norm x" - hence "h \ -x" by clarsimp - hence 2: "x + h \ 0" - apply (rule contrapos_nn) - apply (rule sym) - apply (erule equals_zero_I) - done - show "norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h - = norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h" - apply (subst inverse_diff_inverse [OF 2 x]) - apply (subst minus_diff_minus) - apply (subst norm_minus_cancel) - apply (simp add: left_diff_distrib) - done - next - show "(\h. norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h) - -- 0 --> 0" - proof (rule real_LIM_sandwich_zero) - show "(\h. norm (?inv (x + h) - ?inv x) * norm (?inv x)) - -- 0 --> 0" - apply (rule LIM_mult_left_zero) - apply (rule LIM_norm_zero) - apply (rule LIM_zero) - apply (rule LIM_offset_zero) - apply (rule LIM_inverse) - apply (rule LIM_ident) - apply (rule x) - done - next - fix h::'a assume h: "h \ 0" - show "0 \ norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h" - apply (rule divide_nonneg_pos) - apply (rule norm_ge_zero) - apply (simp add: h) - done - next - fix h::'a assume h: "h \ 0" - have "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h - \ norm (?inv (x + h) - ?inv x) * norm h * norm (?inv x) / norm h" - apply (rule divide_right_mono [OF _ norm_ge_zero]) - apply (rule order_trans [OF norm_mult_ineq]) - apply (rule mult_right_mono [OF _ norm_ge_zero]) - apply (rule norm_mult_ineq) - done - also have "\ = norm (?inv (x + h) - ?inv x) * norm (?inv x)" - by simp - finally show "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h - \ norm (?inv (x + h) - ?inv x) * norm (?inv x)" . - qed - qed -qed - -subsection {* Alternate definition *} - -lemma field_fderiv_def: - fixes x :: "'a::real_normed_field" shows - "FDERIV f x :> (\h. h * D) = (\h. (f (x + h) - f x) / h) -- 0 --> D" - apply (unfold fderiv_def) - apply (simp add: mult.bounded_linear_left) - apply (simp cong: LIM_cong add: nonzero_norm_divide [symmetric]) - apply (subst diff_divide_distrib) - apply (subst times_divide_eq_left [symmetric]) - apply (simp cong: LIM_cong) - apply (simp add: LIM_norm_zero_iff LIM_zero_iff) -done - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/GCD.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Groebner_Basis.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/HOL.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Import/Generate-HOL/GenHOL4Base.thy --- a/src/HOL/Import/Generate-HOL/GenHOL4Base.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Import/Generate-HOL/GenHOL4Base.thy Wed Mar 04 11:05:29 2009 +0100 @@ -88,14 +88,14 @@ import_theory option; type_maps - option > Datatype.option; + option > Option.option; const_maps - NONE > Datatype.option.None - SOME > Datatype.option.Some - option_case > Datatype.option.option_case - OPTION_MAP > Datatype.option_map - THE > Datatype.the + NONE > Option.option.None + SOME > Option.option.Some + option_case > Option.option.option_case + OPTION_MAP > Option.map + THE > Option.the IS_SOME > HOL4Compat.IS_SOME IS_NONE > HOL4Compat.IS_NONE OPTION_JOIN > HOL4Compat.OPTION_JOIN; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Import/HOL4Compat.thy --- a/src/HOL/Import/HOL4Compat.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Import/HOL4Compat.thy Wed Mar 04 11:05:29 2009 +0100 @@ -73,7 +73,7 @@ lemma option_case_def: "(!u f. option_case u f None = u) & (!u f x. option_case u f (Some x) = f x)" by simp -lemma OPTION_MAP_DEF: "(!f x. option_map f (Some x) = Some (f x)) & (!f. option_map f None = None)" +lemma OPTION_MAP_DEF: "(!f x. Option.map f (Some x) = Some (f x)) & (!f. Option.map f None = None)" by simp consts diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Int.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/IntDiv.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/IsaMakefile Wed Mar 04 11:05:29 2009 +0100 @@ -127,6 +127,7 @@ Nat.thy \ OrderedGroup.thy \ Orderings.thy \ + Option.thy \ Plain.thy \ Power.thy \ Predicate.thy \ diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/AssocList.thy --- a/src/HOL/Library/AssocList.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Library/AssocList.thy Wed Mar 04 11:05:29 2009 +0100 @@ -429,7 +429,7 @@ subsection {* @{const map_ran} *} -lemma map_ran_conv: "map_of (map_ran f al) k = option_map (f k) (map_of al k)" +lemma map_ran_conv: "map_of (map_ran f al) k = Option.map (f k) (map_of al k)" by (induct al) auto lemma dom_map_ran: "fst ` set (map_ran f al) = fst ` set al" diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Bit.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Bit.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,122 @@ +(* Title: Bit.thy + Author: Brian Huffman +*) + +header {* The Field of Integers mod 2 *} + +theory Bit +imports Main +begin + +subsection {* Bits as a datatype *} + +typedef (open) bit = "UNIV :: bool set" .. + +instantiation bit :: "{zero, one}" +begin + +definition zero_bit_def: + "0 = Abs_bit False" + +definition one_bit_def: + "1 = Abs_bit True" + +instance .. + +end + +rep_datatype (bit) "0::bit" "1::bit" +proof - + fix P and x :: bit + assume "P (0::bit)" and "P (1::bit)" + then have "\b. P (Abs_bit b)" + unfolding zero_bit_def one_bit_def + by (simp add: all_bool_eq) + then show "P x" + by (induct x) simp +next + show "(0::bit) \ (1::bit)" + unfolding zero_bit_def one_bit_def + by (simp add: Abs_bit_inject) +qed + +lemma bit_not_0_iff [iff]: "(x::bit) \ 0 \ x = 1" + by (induct x) simp_all + +lemma bit_not_1_iff [iff]: "(x::bit) \ 1 \ x = 0" + by (induct x) simp_all + + +subsection {* Type @{typ bit} forms a field *} + +instantiation bit :: "{field, division_by_zero}" +begin + +definition plus_bit_def: + "x + y = (case x of 0 \ y | 1 \ (case y of 0 \ 1 | 1 \ 0))" + +definition times_bit_def: + "x * y = (case x of 0 \ 0 | 1 \ y)" + +definition uminus_bit_def [simp]: + "- x = (x :: bit)" + +definition minus_bit_def [simp]: + "x - y = (x + y :: bit)" + +definition inverse_bit_def [simp]: + "inverse x = (x :: bit)" + +definition divide_bit_def [simp]: + "x / y = (x * y :: bit)" + +lemmas field_bit_defs = + plus_bit_def times_bit_def minus_bit_def uminus_bit_def + divide_bit_def inverse_bit_def + +instance proof +qed (unfold field_bit_defs, auto split: bit.split) + +end + +lemma bit_add_self: "x + x = (0 :: bit)" + unfolding plus_bit_def by (simp split: bit.split) + +lemma bit_mult_eq_1_iff [simp]: "x * y = (1 :: bit) \ x = 1 \ y = 1" + unfolding times_bit_def by (simp split: bit.split) + +text {* Not sure whether the next two should be simp rules. *} + +lemma bit_add_eq_0_iff: "x + y = (0 :: bit) \ x = y" + unfolding plus_bit_def by (simp split: bit.split) + +lemma bit_add_eq_1_iff: "x + y = (1 :: bit) \ x \ y" + unfolding plus_bit_def by (simp split: bit.split) + + +subsection {* Numerals at type @{typ bit} *} + +instantiation bit :: number_ring +begin + +definition number_of_bit_def: + "(number_of w :: bit) = of_int w" + +instance proof +qed (rule number_of_bit_def) + +end + +text {* All numerals reduce to either 0 or 1. *} + +lemma bit_minus1 [simp]: "-1 = (1 :: bit)" + by (simp only: number_of_Min uminus_bit_def) + +lemma bit_number_of_even [simp]: "number_of (Int.Bit0 w) = (0 :: bit)" + by (simp only: number_of_Bit0 add_0_left bit_add_self) + +lemma bit_number_of_odd [simp]: "number_of (Int.Bit1 w) = (1 :: bit)" + by (simp only: number_of_Bit1 add_assoc bit_add_self + monoid_add_class.add_0_right) + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Euclidean_Space.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Float.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/FrechetDeriv.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/FrechetDeriv.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,503 @@ +(* Title : FrechetDeriv.thy + ID : $Id$ + Author : Brian Huffman +*) + +header {* Frechet Derivative *} + +theory FrechetDeriv +imports Lim +begin + +definition + fderiv :: + "['a::real_normed_vector \ 'b::real_normed_vector, 'a, 'a \ 'b] \ bool" + -- {* Frechet derivative: D is derivative of function f at x *} + ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where + "FDERIV f x :> D = (bounded_linear D \ + (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0)" + +lemma FDERIV_I: + "\bounded_linear D; (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0\ + \ FDERIV f x :> D" +by (simp add: fderiv_def) + +lemma FDERIV_D: + "FDERIV f x :> D \ (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0" +by (simp add: fderiv_def) + +lemma FDERIV_bounded_linear: "FDERIV f x :> D \ bounded_linear D" +by (simp add: fderiv_def) + +lemma bounded_linear_zero: + "bounded_linear (\x::'a::real_normed_vector. 0::'b::real_normed_vector)" +proof + show "(0::'b) = 0 + 0" by simp + fix r show "(0::'b) = scaleR r 0" by simp + have "\x::'a. norm (0::'b) \ norm x * 0" by simp + thus "\K. \x::'a. norm (0::'b) \ norm x * K" .. +qed + +lemma FDERIV_const: "FDERIV (\x. k) x :> (\h. 0)" +by (simp add: fderiv_def bounded_linear_zero) + +lemma bounded_linear_ident: + "bounded_linear (\x::'a::real_normed_vector. x)" +proof + fix x y :: 'a show "x + y = x + y" by simp + fix r and x :: 'a show "scaleR r x = scaleR r x" by simp + have "\x::'a. norm x \ norm x * 1" by simp + thus "\K. \x::'a. norm x \ norm x * K" .. +qed + +lemma FDERIV_ident: "FDERIV (\x. x) x :> (\h. h)" +by (simp add: fderiv_def bounded_linear_ident) + +subsection {* Addition *} + +lemma add_diff_add: + fixes a b c d :: "'a::ab_group_add" + shows "(a + c) - (b + d) = (a - b) + (c - d)" +by simp + +lemma bounded_linear_add: + assumes "bounded_linear f" + assumes "bounded_linear g" + shows "bounded_linear (\x. f x + g x)" +proof - + interpret f: bounded_linear f by fact + interpret g: bounded_linear g by fact + show ?thesis apply (unfold_locales) + apply (simp only: f.add g.add add_ac) + apply (simp only: f.scaleR g.scaleR scaleR_right_distrib) + apply (rule f.pos_bounded [THEN exE], rename_tac Kf) + apply (rule g.pos_bounded [THEN exE], rename_tac Kg) + apply (rule_tac x="Kf + Kg" in exI, safe) + apply (subst right_distrib) + apply (rule order_trans [OF norm_triangle_ineq]) + apply (rule add_mono, erule spec, erule spec) + done +qed + +lemma norm_ratio_ineq: + fixes x y :: "'a::real_normed_vector" + fixes h :: "'b::real_normed_vector" + shows "norm (x + y) / norm h \ norm x / norm h + norm y / norm h" +apply (rule ord_le_eq_trans) +apply (rule divide_right_mono) +apply (rule norm_triangle_ineq) +apply (rule norm_ge_zero) +apply (rule add_divide_distrib) +done + +lemma FDERIV_add: + assumes f: "FDERIV f x :> F" + assumes g: "FDERIV g x :> G" + shows "FDERIV (\x. f x + g x) x :> (\h. F h + G h)" +proof (rule FDERIV_I) + show "bounded_linear (\h. F h + G h)" + apply (rule bounded_linear_add) + apply (rule FDERIV_bounded_linear [OF f]) + apply (rule FDERIV_bounded_linear [OF g]) + done +next + have f': "(\h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0" + using f by (rule FDERIV_D) + have g': "(\h. norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0" + using g by (rule FDERIV_D) + from f' g' + have "(\h. norm (f (x + h) - f x - F h) / norm h + + norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0" + by (rule LIM_add_zero) + thus "(\h. norm (f (x + h) + g (x + h) - (f x + g x) - (F h + G h)) + / norm h) -- 0 --> 0" + apply (rule real_LIM_sandwich_zero) + apply (simp add: divide_nonneg_pos) + apply (simp only: add_diff_add) + apply (rule norm_ratio_ineq) + done +qed + +subsection {* Subtraction *} + +lemma bounded_linear_minus: + assumes "bounded_linear f" + shows "bounded_linear (\x. - f x)" +proof - + interpret f: bounded_linear f by fact + show ?thesis apply (unfold_locales) + apply (simp add: f.add) + apply (simp add: f.scaleR) + apply (simp add: f.bounded) + done +qed + +lemma FDERIV_minus: + "FDERIV f x :> F \ FDERIV (\x. - f x) x :> (\h. - F h)" +apply (rule FDERIV_I) +apply (rule bounded_linear_minus) +apply (erule FDERIV_bounded_linear) +apply (simp only: fderiv_def minus_diff_minus norm_minus_cancel) +done + +lemma FDERIV_diff: + "\FDERIV f x :> F; FDERIV g x :> G\ + \ FDERIV (\x. f x - g x) x :> (\h. F h - G h)" +by (simp only: diff_minus FDERIV_add FDERIV_minus) + +subsection {* Continuity *} + +lemma FDERIV_isCont: + assumes f: "FDERIV f x :> F" + shows "isCont f x" +proof - + from f interpret F: bounded_linear "F" by (rule FDERIV_bounded_linear) + have "(\h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0" + by (rule FDERIV_D [OF f]) + hence "(\h. norm (f (x + h) - f x - F h) / norm h * norm h) -- 0 --> 0" + by (intro LIM_mult_zero LIM_norm_zero LIM_ident) + hence "(\h. norm (f (x + h) - f x - F h)) -- 0 --> 0" + by (simp cong: LIM_cong) + hence "(\h. f (x + h) - f x - F h) -- 0 --> 0" + by (rule LIM_norm_zero_cancel) + hence "(\h. f (x + h) - f x - F h + F h) -- 0 --> 0" + by (intro LIM_add_zero F.LIM_zero LIM_ident) + hence "(\h. f (x + h) - f x) -- 0 --> 0" + by simp + thus "isCont f x" + unfolding isCont_iff by (rule LIM_zero_cancel) +qed + +subsection {* Composition *} + +lemma real_divide_cancel_lemma: + fixes a b c :: real + shows "(b = 0 \ a = 0) \ (a / b) * (b / c) = a / c" +by simp + +lemma bounded_linear_compose: + assumes "bounded_linear f" + assumes "bounded_linear g" + shows "bounded_linear (\x. f (g x))" +proof - + interpret f: bounded_linear f by fact + interpret g: bounded_linear g by fact + show ?thesis proof (unfold_locales) + fix x y show "f (g (x + y)) = f (g x) + f (g y)" + by (simp only: f.add g.add) + next + fix r x show "f (g (scaleR r x)) = scaleR r (f (g x))" + by (simp only: f.scaleR g.scaleR) + next + from f.pos_bounded + obtain Kf where f: "\x. norm (f x) \ norm x * Kf" and Kf: "0 < Kf" by fast + from g.pos_bounded + obtain Kg where g: "\x. norm (g x) \ norm x * Kg" by fast + show "\K. \x. norm (f (g x)) \ norm x * K" + proof (intro exI allI) + fix x + have "norm (f (g x)) \ norm (g x) * Kf" + using f . + also have "\ \ (norm x * Kg) * Kf" + using g Kf [THEN order_less_imp_le] by (rule mult_right_mono) + also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)" + by (rule mult_assoc) + finally show "norm (f (g x)) \ norm x * (Kg * Kf)" . + qed + qed +qed + +lemma FDERIV_compose: + fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" + fixes g :: "'b::real_normed_vector \ 'c::real_normed_vector" + assumes f: "FDERIV f x :> F" + assumes g: "FDERIV g (f x) :> G" + shows "FDERIV (\x. g (f x)) x :> (\h. G (F h))" +proof (rule FDERIV_I) + from FDERIV_bounded_linear [OF g] FDERIV_bounded_linear [OF f] + show "bounded_linear (\h. G (F h))" + by (rule bounded_linear_compose) +next + let ?Rf = "\h. f (x + h) - f x - F h" + let ?Rg = "\k. g (f x + k) - g (f x) - G k" + let ?k = "\h. f (x + h) - f x" + let ?Nf = "\h. norm (?Rf h) / norm h" + let ?Ng = "\h. norm (?Rg (?k h)) / norm (?k h)" + from f interpret F!: bounded_linear "F" by (rule FDERIV_bounded_linear) + from g interpret G!: bounded_linear "G" by (rule FDERIV_bounded_linear) + from F.bounded obtain kF where kF: "\x. norm (F x) \ norm x * kF" by fast + from G.bounded obtain kG where kG: "\x. norm (G x) \ norm x * kG" by fast + + let ?fun2 = "\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)" + + show "(\h. norm (g (f (x + h)) - g (f x) - G (F h)) / norm h) -- 0 --> 0" + proof (rule real_LIM_sandwich_zero) + have Nf: "?Nf -- 0 --> 0" + using FDERIV_D [OF f] . + + have Ng1: "isCont (\k. norm (?Rg k) / norm k) 0" + by (simp add: isCont_def FDERIV_D [OF g]) + have Ng2: "?k -- 0 --> 0" + apply (rule LIM_zero) + apply (fold isCont_iff) + apply (rule FDERIV_isCont [OF f]) + done + have Ng: "?Ng -- 0 --> 0" + using isCont_LIM_compose [OF Ng1 Ng2] by simp + + have "(\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) + -- 0 --> 0 * kG + 0 * (0 + kF)" + by (intro LIM_add LIM_mult LIM_const Nf Ng) + thus "(\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) -- 0 --> 0" + by simp + next + fix h::'a assume h: "h \ 0" + thus "0 \ norm (g (f (x + h)) - g (f x) - G (F h)) / norm h" + by (simp add: divide_nonneg_pos) + next + fix h::'a assume h: "h \ 0" + have "g (f (x + h)) - g (f x) - G (F h) = G (?Rf h) + ?Rg (?k h)" + by (simp add: G.diff) + hence "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h + = norm (G (?Rf h) + ?Rg (?k h)) / norm h" + by (rule arg_cong) + also have "\ \ norm (G (?Rf h)) / norm h + norm (?Rg (?k h)) / norm h" + by (rule norm_ratio_ineq) + also have "\ \ ?Nf h * kG + ?Ng h * (?Nf h + kF)" + proof (rule add_mono) + show "norm (G (?Rf h)) / norm h \ ?Nf h * kG" + apply (rule ord_le_eq_trans) + apply (rule divide_right_mono [OF kG norm_ge_zero]) + apply simp + done + next + have "norm (?Rg (?k h)) / norm h = ?Ng h * (norm (?k h) / norm h)" + apply (rule real_divide_cancel_lemma [symmetric]) + apply (simp add: G.zero) + done + also have "\ \ ?Ng h * (?Nf h + kF)" + proof (rule mult_left_mono) + have "norm (?k h) / norm h = norm (?Rf h + F h) / norm h" + by simp + also have "\ \ ?Nf h + norm (F h) / norm h" + by (rule norm_ratio_ineq) + also have "\ \ ?Nf h + kF" + apply (rule add_left_mono) + apply (subst pos_divide_le_eq, simp add: h) + apply (subst mult_commute) + apply (rule kF) + done + finally show "norm (?k h) / norm h \ ?Nf h + kF" . + next + show "0 \ ?Ng h" + apply (case_tac "f (x + h) - f x = 0", simp) + apply (rule divide_nonneg_pos [OF norm_ge_zero]) + apply simp + done + qed + finally show "norm (?Rg (?k h)) / norm h \ ?Ng h * (?Nf h + kF)" . + qed + finally show "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h + \ ?Nf h * kG + ?Ng h * (?Nf h + kF)" . + qed +qed + +subsection {* Product Rule *} + +lemma (in bounded_bilinear) FDERIV_lemma: + "a' ** b' - a ** b - (a ** B + A ** b) + = a ** (b' - b - B) + (a' - a - A) ** b' + A ** (b' - b)" +by (simp add: diff_left diff_right) + +lemma (in bounded_bilinear) FDERIV: + fixes x :: "'d::real_normed_vector" + assumes f: "FDERIV f x :> F" + assumes g: "FDERIV g x :> G" + shows "FDERIV (\x. f x ** g x) x :> (\h. f x ** G h + F h ** g x)" +proof (rule FDERIV_I) + show "bounded_linear (\h. f x ** G h + F h ** g x)" + apply (rule bounded_linear_add) + apply (rule bounded_linear_compose [OF bounded_linear_right]) + apply (rule FDERIV_bounded_linear [OF g]) + apply (rule bounded_linear_compose [OF bounded_linear_left]) + apply (rule FDERIV_bounded_linear [OF f]) + done +next + from bounded_linear.bounded [OF FDERIV_bounded_linear [OF f]] + obtain KF where norm_F: "\x. norm (F x) \ norm x * KF" by fast + + from pos_bounded obtain K where K: "0 < K" and norm_prod: + "\a b. norm (a ** b) \ norm a * norm b * K" by fast + + let ?Rf = "\h. f (x + h) - f x - F h" + let ?Rg = "\h. g (x + h) - g x - G h" + + let ?fun1 = "\h. + norm (f x ** ?Rg h + ?Rf h ** g (x + h) + F h ** (g (x + h) - g x)) / + norm h" + + let ?fun2 = "\h. + norm (f x) * (norm (?Rg h) / norm h) * K + + norm (?Rf h) / norm h * norm (g (x + h)) * K + + KF * norm (g (x + h) - g x) * K" + + have "?fun1 -- 0 --> 0" + proof (rule real_LIM_sandwich_zero) + from f g isCont_iff [THEN iffD1, OF FDERIV_isCont [OF g]] + have "?fun2 -- 0 --> + norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K" + by (intro LIM_add LIM_mult LIM_const LIM_norm LIM_zero FDERIV_D) + thus "?fun2 -- 0 --> 0" + by simp + next + fix h::'d assume "h \ 0" + thus "0 \ ?fun1 h" + by (simp add: divide_nonneg_pos) + next + fix h::'d assume "h \ 0" + have "?fun1 h \ (norm (f x) * norm (?Rg h) * K + + norm (?Rf h) * norm (g (x + h)) * K + + norm h * KF * norm (g (x + h) - g x) * K) / norm h" + by (intro + divide_right_mono mult_mono' + order_trans [OF norm_triangle_ineq add_mono] + order_trans [OF norm_prod mult_right_mono] + mult_nonneg_nonneg order_refl norm_ge_zero norm_F + K [THEN order_less_imp_le] + ) + also have "\ = ?fun2 h" + by (simp add: add_divide_distrib) + finally show "?fun1 h \ ?fun2 h" . + qed + thus "(\h. + norm (f (x + h) ** g (x + h) - f x ** g x - (f x ** G h + F h ** g x)) + / norm h) -- 0 --> 0" + by (simp only: FDERIV_lemma) +qed + +lemmas FDERIV_mult = mult.FDERIV + +lemmas FDERIV_scaleR = scaleR.FDERIV + + +subsection {* Powers *} + +lemma FDERIV_power_Suc: + fixes x :: "'a::{real_normed_algebra,recpower,comm_ring_1}" + shows "FDERIV (\x. x ^ Suc n) x :> (\h. (1 + of_nat n) * x ^ n * h)" + apply (induct n) + apply (simp add: power_Suc FDERIV_ident) + apply (drule FDERIV_mult [OF FDERIV_ident]) + apply (simp only: of_nat_Suc left_distrib mult_1_left) + apply (simp only: power_Suc right_distrib add_ac mult_ac) +done + +lemma FDERIV_power: + fixes x :: "'a::{real_normed_algebra,recpower,comm_ring_1}" + shows "FDERIV (\x. x ^ n) x :> (\h. of_nat n * x ^ (n - 1) * h)" + apply (cases n) + apply (simp add: FDERIV_const) + apply (simp add: FDERIV_power_Suc) + done + + +subsection {* Inverse *} + +lemma inverse_diff_inverse: + "\(a::'a::division_ring) \ 0; b \ 0\ + \ inverse a - inverse b = - (inverse a * (a - b) * inverse b)" +by (simp add: right_diff_distrib left_diff_distrib mult_assoc) + +lemmas bounded_linear_mult_const = + mult.bounded_linear_left [THEN bounded_linear_compose] + +lemmas bounded_linear_const_mult = + mult.bounded_linear_right [THEN bounded_linear_compose] + +lemma FDERIV_inverse: + fixes x :: "'a::real_normed_div_algebra" + assumes x: "x \ 0" + shows "FDERIV inverse x :> (\h. - (inverse x * h * inverse x))" + (is "FDERIV ?inv _ :> _") +proof (rule FDERIV_I) + show "bounded_linear (\h. - (?inv x * h * ?inv x))" + apply (rule bounded_linear_minus) + apply (rule bounded_linear_mult_const) + apply (rule bounded_linear_const_mult) + apply (rule bounded_linear_ident) + done +next + show "(\h. norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h) + -- 0 --> 0" + proof (rule LIM_equal2) + show "0 < norm x" using x by simp + next + fix h::'a + assume 1: "h \ 0" + assume "norm (h - 0) < norm x" + hence "h \ -x" by clarsimp + hence 2: "x + h \ 0" + apply (rule contrapos_nn) + apply (rule sym) + apply (erule equals_zero_I) + done + show "norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h + = norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h" + apply (subst inverse_diff_inverse [OF 2 x]) + apply (subst minus_diff_minus) + apply (subst norm_minus_cancel) + apply (simp add: left_diff_distrib) + done + next + show "(\h. norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h) + -- 0 --> 0" + proof (rule real_LIM_sandwich_zero) + show "(\h. norm (?inv (x + h) - ?inv x) * norm (?inv x)) + -- 0 --> 0" + apply (rule LIM_mult_left_zero) + apply (rule LIM_norm_zero) + apply (rule LIM_zero) + apply (rule LIM_offset_zero) + apply (rule LIM_inverse) + apply (rule LIM_ident) + apply (rule x) + done + next + fix h::'a assume h: "h \ 0" + show "0 \ norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h" + apply (rule divide_nonneg_pos) + apply (rule norm_ge_zero) + apply (simp add: h) + done + next + fix h::'a assume h: "h \ 0" + have "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h + \ norm (?inv (x + h) - ?inv x) * norm h * norm (?inv x) / norm h" + apply (rule divide_right_mono [OF _ norm_ge_zero]) + apply (rule order_trans [OF norm_mult_ineq]) + apply (rule mult_right_mono [OF _ norm_ge_zero]) + apply (rule norm_mult_ineq) + done + also have "\ = norm (?inv (x + h) - ?inv x) * norm (?inv x)" + by simp + finally show "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h + \ norm (?inv (x + h) - ?inv x) * norm (?inv x)" . + qed + qed +qed + +subsection {* Alternate definition *} + +lemma field_fderiv_def: + fixes x :: "'a::real_normed_field" shows + "FDERIV f x :> (\h. h * D) = (\h. (f (x + h) - f x) / h) -- 0 --> D" + apply (unfold fderiv_def) + apply (simp add: mult.bounded_linear_left) + apply (simp cong: LIM_cong add: nonzero_norm_divide [symmetric]) + apply (subst diff_divide_distrib) + apply (subst times_divide_eq_left [symmetric]) + apply (simp cong: LIM_cong) + apply (simp add: LIM_norm_zero_iff LIM_zero_iff) +done + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Fundamental_Theorem_Algebra.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Inner_Product.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Inner_Product.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,296 @@ +(* Title: Inner_Product.thy + Author: Brian Huffman +*) + +header {* Inner Product Spaces and the Gradient Derivative *} + +theory Inner_Product +imports Complex FrechetDeriv +begin + +subsection {* Real inner product spaces *} + +class real_inner = real_vector + sgn_div_norm + + fixes inner :: "'a \ 'a \ real" + assumes inner_commute: "inner x y = inner y x" + and inner_left_distrib: "inner (x + y) z = inner x z + inner y z" + and inner_scaleR_left: "inner (scaleR r x) y = r * (inner x y)" + and inner_ge_zero [simp]: "0 \ inner x x" + and inner_eq_zero_iff [simp]: "inner x x = 0 \ x = 0" + and norm_eq_sqrt_inner: "norm x = sqrt (inner x x)" +begin + +lemma inner_zero_left [simp]: "inner 0 x = 0" + using inner_left_distrib [of 0 0 x] by simp + +lemma inner_minus_left [simp]: "inner (- x) y = - inner x y" + using inner_left_distrib [of x "- x" y] by simp + +lemma inner_diff_left: "inner (x - y) z = inner x z - inner y z" + by (simp add: diff_minus inner_left_distrib) + +text {* Transfer distributivity rules to right argument. *} + +lemma inner_right_distrib: "inner x (y + z) = inner x y + inner x z" + using inner_left_distrib [of y z x] by (simp only: inner_commute) + +lemma inner_scaleR_right: "inner x (scaleR r y) = r * (inner x y)" + using inner_scaleR_left [of r y x] by (simp only: inner_commute) + +lemma inner_zero_right [simp]: "inner x 0 = 0" + using inner_zero_left [of x] by (simp only: inner_commute) + +lemma inner_minus_right [simp]: "inner x (- y) = - inner x y" + using inner_minus_left [of y x] by (simp only: inner_commute) + +lemma inner_diff_right: "inner x (y - z) = inner x y - inner x z" + using inner_diff_left [of y z x] by (simp only: inner_commute) + +lemmas inner_distrib = inner_left_distrib inner_right_distrib +lemmas inner_diff = inner_diff_left inner_diff_right +lemmas inner_scaleR = inner_scaleR_left inner_scaleR_right + +lemma inner_gt_zero_iff [simp]: "0 < inner x x \ x \ 0" + by (simp add: order_less_le) + +lemma power2_norm_eq_inner: "(norm x)\ = inner x x" + by (simp add: norm_eq_sqrt_inner) + +lemma Cauchy_Schwarz_ineq: + "(inner x y)\ \ inner x x * inner y y" +proof (cases) + assume "y = 0" + thus ?thesis by simp +next + assume y: "y \ 0" + let ?r = "inner x y / inner y y" + have "0 \ inner (x - scaleR ?r y) (x - scaleR ?r y)" + by (rule inner_ge_zero) + also have "\ = inner x x - inner y x * ?r" + by (simp add: inner_diff inner_scaleR) + also have "\ = inner x x - (inner x y)\ / inner y y" + by (simp add: power2_eq_square inner_commute) + finally have "0 \ inner x x - (inner x y)\ / inner y y" . + hence "(inner x y)\ / inner y y \ inner x x" + by (simp add: le_diff_eq) + thus "(inner x y)\ \ inner x x * inner y y" + by (simp add: pos_divide_le_eq y) +qed + +lemma Cauchy_Schwarz_ineq2: + "\inner x y\ \ norm x * norm y" +proof (rule power2_le_imp_le) + have "(inner x y)\ \ inner x x * inner y y" + using Cauchy_Schwarz_ineq . + thus "\inner x y\\ \ (norm x * norm y)\" + by (simp add: power_mult_distrib power2_norm_eq_inner) + show "0 \ norm x * norm y" + unfolding norm_eq_sqrt_inner + by (intro mult_nonneg_nonneg real_sqrt_ge_zero inner_ge_zero) +qed + +subclass real_normed_vector +proof + fix a :: real and x y :: 'a + show "0 \ norm x" + unfolding norm_eq_sqrt_inner by simp + show "norm x = 0 \ x = 0" + unfolding norm_eq_sqrt_inner by simp + show "norm (x + y) \ norm x + norm y" + proof (rule power2_le_imp_le) + have "inner x y \ norm x * norm y" + by (rule order_trans [OF abs_ge_self Cauchy_Schwarz_ineq2]) + thus "(norm (x + y))\ \ (norm x + norm y)\" + unfolding power2_sum power2_norm_eq_inner + by (simp add: inner_distrib inner_commute) + show "0 \ norm x + norm y" + unfolding norm_eq_sqrt_inner + by (simp add: add_nonneg_nonneg) + qed + have "sqrt (a\ * inner x x) = \a\ * sqrt (inner x x)" + by (simp add: real_sqrt_mult_distrib) + then show "norm (a *\<^sub>R x) = \a\ * norm x" + unfolding norm_eq_sqrt_inner + by (simp add: inner_scaleR power2_eq_square mult_assoc) +qed + +end + +interpretation inner!: + bounded_bilinear "inner::'a::real_inner \ 'a \ real" +proof + fix x y z :: 'a and r :: real + show "inner (x + y) z = inner x z + inner y z" + by (rule inner_left_distrib) + show "inner x (y + z) = inner x y + inner x z" + by (rule inner_right_distrib) + show "inner (scaleR r x) y = scaleR r (inner x y)" + unfolding real_scaleR_def by (rule inner_scaleR_left) + show "inner x (scaleR r y) = scaleR r (inner x y)" + unfolding real_scaleR_def by (rule inner_scaleR_right) + show "\K. \x y::'a. norm (inner x y) \ norm x * norm y * K" + proof + show "\x y::'a. norm (inner x y) \ norm x * norm y * 1" + by (simp add: Cauchy_Schwarz_ineq2) + qed +qed + +interpretation inner_left!: + bounded_linear "\x::'a::real_inner. inner x y" + by (rule inner.bounded_linear_left) + +interpretation inner_right!: + bounded_linear "\y::'a::real_inner. inner x y" + by (rule inner.bounded_linear_right) + + +subsection {* Class instances *} + +instantiation real :: real_inner +begin + +definition inner_real_def [simp]: "inner = op *" + +instance proof + fix x y z r :: real + show "inner x y = inner y x" + unfolding inner_real_def by (rule mult_commute) + show "inner (x + y) z = inner x z + inner y z" + unfolding inner_real_def by (rule left_distrib) + show "inner (scaleR r x) y = r * inner x y" + unfolding inner_real_def real_scaleR_def by (rule mult_assoc) + show "0 \ inner x x" + unfolding inner_real_def by simp + show "inner x x = 0 \ x = 0" + unfolding inner_real_def by simp + show "norm x = sqrt (inner x x)" + unfolding inner_real_def by simp +qed + +end + +instantiation complex :: real_inner +begin + +definition inner_complex_def: + "inner x y = Re x * Re y + Im x * Im y" + +instance proof + fix x y z :: complex and r :: real + show "inner x y = inner y x" + unfolding inner_complex_def by (simp add: mult_commute) + show "inner (x + y) z = inner x z + inner y z" + unfolding inner_complex_def by (simp add: left_distrib) + show "inner (scaleR r x) y = r * inner x y" + unfolding inner_complex_def by (simp add: right_distrib) + show "0 \ inner x x" + unfolding inner_complex_def by (simp add: add_nonneg_nonneg) + show "inner x x = 0 \ x = 0" + unfolding inner_complex_def + by (simp add: add_nonneg_eq_0_iff complex_Re_Im_cancel_iff) + show "norm x = sqrt (inner x x)" + unfolding inner_complex_def complex_norm_def + by (simp add: power2_eq_square) +qed + +end + + +subsection {* Gradient derivative *} + +definition + gderiv :: + "['a::real_inner \ real, 'a, 'a] \ bool" + ("(GDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) +where + "GDERIV f x :> D \ FDERIV f x :> (\h. inner h D)" + +lemma deriv_fderiv: "DERIV f x :> D \ FDERIV f x :> (\h. h * D)" + by (simp only: deriv_def field_fderiv_def) + +lemma gderiv_deriv [simp]: "GDERIV f x :> D \ DERIV f x :> D" + by (simp only: gderiv_def deriv_fderiv inner_real_def) + +lemma GDERIV_DERIV_compose: + "\GDERIV f x :> df; DERIV g (f x) :> dg\ + \ GDERIV (\x. g (f x)) x :> scaleR dg df" + unfolding gderiv_def deriv_fderiv + apply (drule (1) FDERIV_compose) + apply (simp add: inner_scaleR_right mult_ac) + done + +lemma FDERIV_subst: "\FDERIV f x :> df; df = d\ \ FDERIV f x :> d" + by simp + +lemma GDERIV_subst: "\GDERIV f x :> df; df = d\ \ GDERIV f x :> d" + by simp + +lemma GDERIV_const: "GDERIV (\x. k) x :> 0" + unfolding gderiv_def inner_right.zero by (rule FDERIV_const) + +lemma GDERIV_add: + "\GDERIV f x :> df; GDERIV g x :> dg\ + \ GDERIV (\x. f x + g x) x :> df + dg" + unfolding gderiv_def inner_right.add by (rule FDERIV_add) + +lemma GDERIV_minus: + "GDERIV f x :> df \ GDERIV (\x. - f x) x :> - df" + unfolding gderiv_def inner_right.minus by (rule FDERIV_minus) + +lemma GDERIV_diff: + "\GDERIV f x :> df; GDERIV g x :> dg\ + \ GDERIV (\x. f x - g x) x :> df - dg" + unfolding gderiv_def inner_right.diff by (rule FDERIV_diff) + +lemma GDERIV_scaleR: + "\DERIV f x :> df; GDERIV g x :> dg\ + \ GDERIV (\x. scaleR (f x) (g x)) x + :> (scaleR (f x) dg + scaleR df (g x))" + unfolding gderiv_def deriv_fderiv inner_right.add inner_right.scaleR + apply (rule FDERIV_subst) + apply (erule (1) scaleR.FDERIV) + apply (simp add: mult_ac) + done + +lemma GDERIV_mult: + "\GDERIV f x :> df; GDERIV g x :> dg\ + \ GDERIV (\x. f x * g x) x :> scaleR (f x) dg + scaleR (g x) df" + unfolding gderiv_def + apply (rule FDERIV_subst) + apply (erule (1) FDERIV_mult) + apply (simp add: inner_distrib inner_scaleR mult_ac) + done + +lemma GDERIV_inverse: + "\GDERIV f x :> df; f x \ 0\ + \ GDERIV (\x. inverse (f x)) x :> - (inverse (f x))\ *\<^sub>R df" + apply (erule GDERIV_DERIV_compose) + apply (erule DERIV_inverse [folded numeral_2_eq_2]) + done + +lemma GDERIV_norm: + assumes "x \ 0" shows "GDERIV (\x. norm x) x :> sgn x" +proof - + have 1: "FDERIV (\x. inner x x) x :> (\h. inner x h + inner h x)" + by (intro inner.FDERIV FDERIV_ident) + have 2: "(\h. inner x h + inner h x) = (\h. inner h (scaleR 2 x))" + by (simp add: expand_fun_eq inner_scaleR inner_commute) + have "0 < inner x x" using `x \ 0` by simp + then have 3: "DERIV sqrt (inner x x) :> (inverse (sqrt (inner x x)) / 2)" + by (rule DERIV_real_sqrt) + have 4: "(inverse (sqrt (inner x x)) / 2) *\<^sub>R 2 *\<^sub>R x = sgn x" + by (simp add: sgn_div_norm norm_eq_sqrt_inner) + show ?thesis + unfolding norm_eq_sqrt_inner + apply (rule GDERIV_subst [OF _ 4]) + apply (rule GDERIV_DERIV_compose [where g=sqrt and df="scaleR 2 x"]) + apply (subst gderiv_def) + apply (rule FDERIV_subst [OF _ 2]) + apply (rule 1) + apply (rule 3) + done +qed + +lemmas FDERIV_norm = GDERIV_norm [unfolded gderiv_def] + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Library.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Numeral_Type.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Permutations.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Pocklington.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Poly_Deriv.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Poly_Deriv.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,316 @@ +(* Title: Poly_Deriv.thy + Author: Amine Chaieb + Ported to new Polynomial library by Brian Huffman +*) + +header{* Polynomials and Differentiation *} + +theory Poly_Deriv +imports Deriv Polynomial +begin + +subsection {* Derivatives of univariate polynomials *} + +definition + pderiv :: "'a::real_normed_field poly \ 'a poly" where + "pderiv = poly_rec 0 (\a p p'. p + pCons 0 p')" + +lemma pderiv_0 [simp]: "pderiv 0 = 0" + unfolding pderiv_def by (simp add: poly_rec_0) + +lemma pderiv_pCons: "pderiv (pCons a p) = p + pCons 0 (pderiv p)" + unfolding pderiv_def by (simp add: poly_rec_pCons) + +lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)" + apply (induct p arbitrary: n, simp) + apply (simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split) + done + +lemma pderiv_eq_0_iff: "pderiv p = 0 \ degree p = 0" + apply (rule iffI) + apply (cases p, simp) + apply (simp add: expand_poly_eq coeff_pderiv del: of_nat_Suc) + apply (simp add: expand_poly_eq coeff_pderiv coeff_eq_0) + done + +lemma degree_pderiv: "degree (pderiv p) = degree p - 1" + apply (rule order_antisym [OF degree_le]) + apply (simp add: coeff_pderiv coeff_eq_0) + apply (cases "degree p", simp) + apply (rule le_degree) + apply (simp add: coeff_pderiv del: of_nat_Suc) + apply (rule subst, assumption) + apply (rule leading_coeff_neq_0, clarsimp) + done + +lemma pderiv_singleton [simp]: "pderiv [:a:] = 0" +by (simp add: pderiv_pCons) + +lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q" +by (rule poly_ext, simp add: coeff_pderiv algebra_simps) + +lemma pderiv_minus: "pderiv (- p) = - pderiv p" +by (rule poly_ext, simp add: coeff_pderiv) + +lemma pderiv_diff: "pderiv (p - q) = pderiv p - pderiv q" +by (rule poly_ext, simp add: coeff_pderiv algebra_simps) + +lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)" +by (rule poly_ext, simp add: coeff_pderiv algebra_simps) + +lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p" +apply (induct p) +apply simp +apply (simp add: pderiv_add pderiv_smult pderiv_pCons algebra_simps) +done + +lemma pderiv_power_Suc: + "pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p" +apply (induct n) +apply simp +apply (subst power_Suc) +apply (subst pderiv_mult) +apply (erule ssubst) +apply (simp add: smult_add_left algebra_simps) +done + +lemma DERIV_cmult2: "DERIV f x :> D ==> DERIV (%x. (f x) * c :: real) x :> D * c" +by (simp add: DERIV_cmult mult_commute [of _ c]) + +lemma DERIV_pow2: "DERIV (%x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" +by (rule lemma_DERIV_subst, rule DERIV_pow, simp) +declare DERIV_pow2 [simp] DERIV_pow [simp] + +lemma DERIV_add_const: "DERIV f x :> D ==> DERIV (%x. a + f x :: 'a::real_normed_field) x :> D" +by (rule lemma_DERIV_subst, rule DERIV_add, auto) + +lemma poly_DERIV[simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x" +apply (induct p) +apply simp +apply (simp add: pderiv_pCons) +apply (rule lemma_DERIV_subst) +apply (rule DERIV_add DERIV_mult DERIV_const DERIV_ident | assumption)+ +apply simp +done + +text{* Consequences of the derivative theorem above*} + +lemma poly_differentiable[simp]: "(%x. poly p x) differentiable (x::real)" +apply (simp add: differentiable_def) +apply (blast intro: poly_DERIV) +done + +lemma poly_isCont[simp]: "isCont (%x. poly p x) (x::real)" +by (rule poly_DERIV [THEN DERIV_isCont]) + +lemma poly_IVT_pos: "[| a < b; poly p (a::real) < 0; 0 < poly p b |] + ==> \x. a < x & x < b & (poly p x = 0)" +apply (cut_tac f = "%x. poly p x" and a = a and b = b and y = 0 in IVT_objl) +apply (auto simp add: order_le_less) +done + +lemma poly_IVT_neg: "[| (a::real) < b; 0 < poly p a; poly p b < 0 |] + ==> \x. a < x & x < b & (poly p x = 0)" +by (insert poly_IVT_pos [where p = "- p" ]) simp + +lemma poly_MVT: "(a::real) < b ==> + \x. a < x & x < b & (poly p b - poly p a = (b - a) * poly (pderiv p) x)" +apply (drule_tac f = "poly p" in MVT, auto) +apply (rule_tac x = z in exI) +apply (auto simp add: real_mult_left_cancel poly_DERIV [THEN DERIV_unique]) +done + +text{*Lemmas for Derivatives*} + +lemma order_unique_lemma: + fixes p :: "'a::idom poly" + assumes "[:-a, 1:] ^ n dvd p \ \ [:-a, 1:] ^ Suc n dvd p" + shows "n = order a p" +unfolding Polynomial.order_def +apply (rule Least_equality [symmetric]) +apply (rule assms [THEN conjunct2]) +apply (erule contrapos_np) +apply (rule power_le_dvd) +apply (rule assms [THEN conjunct1]) +apply simp +done + +lemma lemma_order_pderiv1: + "pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q + + smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)" +apply (simp only: pderiv_mult pderiv_power_Suc) +apply (simp del: power_poly_Suc of_nat_Suc add: pderiv_pCons) +done + +lemma dvd_add_cancel1: + fixes a b c :: "'a::comm_ring_1" + shows "a dvd b + c \ a dvd b \ a dvd c" + by (drule (1) Ring_and_Field.dvd_diff, simp) + +lemma lemma_order_pderiv [rule_format]: + "\p q a. 0 < n & + pderiv p \ 0 & + p = [:- a, 1:] ^ n * q & ~ [:- a, 1:] dvd q + --> n = Suc (order a (pderiv p))" + apply (cases "n", safe, rename_tac n p q a) + apply (rule order_unique_lemma) + apply (rule conjI) + apply (subst lemma_order_pderiv1) + apply (rule dvd_add) + apply (rule dvd_mult2) + apply (rule le_imp_power_dvd, simp) + apply (rule dvd_smult) + apply (rule dvd_mult) + apply (rule dvd_refl) + apply (subst lemma_order_pderiv1) + apply (erule contrapos_nn) back + apply (subgoal_tac "[:- a, 1:] ^ Suc n dvd q * [:- a, 1:] ^ n") + apply (simp del: mult_pCons_left) + apply (drule dvd_add_cancel1) + apply (simp del: mult_pCons_left) + apply (drule dvd_smult_cancel, simp del: of_nat_Suc) + apply assumption +done + +lemma order_decomp: + "p \ 0 + ==> \q. p = [:-a, 1:] ^ (order a p) * q & + ~([:-a, 1:] dvd q)" +apply (drule order [where a=a]) +apply (erule conjE) +apply (erule dvdE) +apply (rule exI) +apply (rule conjI, assumption) +apply (erule contrapos_nn) +apply (erule ssubst) back +apply (subst power_Suc2) +apply (erule mult_dvd_mono [OF dvd_refl]) +done + +lemma order_pderiv: "[| pderiv p \ 0; order a p \ 0 |] + ==> (order a p = Suc (order a (pderiv p)))" +apply (case_tac "p = 0", simp) +apply (drule_tac a = a and p = p in order_decomp) +using neq0_conv +apply (blast intro: lemma_order_pderiv) +done + +lemma order_mult: "p * q \ 0 \ order a (p * q) = order a p + order a q" +proof - + def i \ "order a p" + def j \ "order a q" + def t \ "[:-a, 1:]" + have t_dvd_iff: "\u. t dvd u \ poly u a = 0" + unfolding t_def by (simp add: dvd_iff_poly_eq_0) + assume "p * q \ 0" + then show "order a (p * q) = i + j" + apply clarsimp + apply (drule order [where a=a and p=p, folded i_def t_def]) + apply (drule order [where a=a and p=q, folded j_def t_def]) + apply clarify + apply (rule order_unique_lemma [symmetric], fold t_def) + apply (erule dvdE)+ + apply (simp add: power_add t_dvd_iff) + done +qed + +text{*Now justify the standard squarefree decomposition, i.e. f / gcd(f,f'). *} + +lemma order_divides: "[:-a, 1:] ^ n dvd p \ p = 0 \ n \ order a p" +apply (cases "p = 0", auto) +apply (drule order_2 [where a=a and p=p]) +apply (erule contrapos_np) +apply (erule power_le_dvd) +apply simp +apply (erule power_le_dvd [OF order_1]) +done + +lemma poly_squarefree_decomp_order: + assumes "pderiv p \ 0" + and p: "p = q * d" + and p': "pderiv p = e * d" + and d: "d = r * p + s * pderiv p" + shows "order a q = (if order a p = 0 then 0 else 1)" +proof (rule classical) + assume 1: "order a q \ (if order a p = 0 then 0 else 1)" + from `pderiv p \ 0` have "p \ 0" by auto + with p have "order a p = order a q + order a d" + by (simp add: order_mult) + with 1 have "order a p \ 0" by (auto split: if_splits) + have "order a (pderiv p) = order a e + order a d" + using `pderiv p \ 0` `pderiv p = e * d` by (simp add: order_mult) + have "order a p = Suc (order a (pderiv p))" + using `pderiv p \ 0` `order a p \ 0` by (rule order_pderiv) + have "d \ 0" using `p \ 0` `p = q * d` by simp + have "([:-a, 1:] ^ (order a (pderiv p))) dvd d" + apply (simp add: d) + apply (rule dvd_add) + apply (rule dvd_mult) + apply (simp add: order_divides `p \ 0` + `order a p = Suc (order a (pderiv p))`) + apply (rule dvd_mult) + apply (simp add: order_divides) + done + then have "order a (pderiv p) \ order a d" + using `d \ 0` by (simp add: order_divides) + show ?thesis + using `order a p = order a q + order a d` + using `order a (pderiv p) = order a e + order a d` + using `order a p = Suc (order a (pderiv p))` + using `order a (pderiv p) \ order a d` + by auto +qed + +lemma poly_squarefree_decomp_order2: "[| pderiv p \ 0; + p = q * d; + pderiv p = e * d; + d = r * p + s * pderiv p + |] ==> \a. order a q = (if order a p = 0 then 0 else 1)" +apply (blast intro: poly_squarefree_decomp_order) +done + +lemma order_pderiv2: "[| pderiv p \ 0; order a p \ 0 |] + ==> (order a (pderiv p) = n) = (order a p = Suc n)" +apply (auto dest: order_pderiv) +done + +definition + rsquarefree :: "'a::idom poly => bool" where + "rsquarefree p = (p \ 0 & (\a. (order a p = 0) | (order a p = 1)))" + +lemma pderiv_iszero: "pderiv p = 0 \ \h. p = [:h:]" +apply (simp add: pderiv_eq_0_iff) +apply (case_tac p, auto split: if_splits) +done + +lemma rsquarefree_roots: + "rsquarefree p = (\a. ~(poly p a = 0 & poly (pderiv p) a = 0))" +apply (simp add: rsquarefree_def) +apply (case_tac "p = 0", simp, simp) +apply (case_tac "pderiv p = 0") +apply simp +apply (drule pderiv_iszero, clarify) +apply simp +apply (rule allI) +apply (cut_tac p = "[:h:]" and a = a in order_root) +apply simp +apply (auto simp add: order_root order_pderiv2) +apply (erule_tac x="a" in allE, simp) +done + +lemma poly_squarefree_decomp: + assumes "pderiv p \ 0" + and "p = q * d" + and "pderiv p = e * d" + and "d = r * p + s * pderiv p" + shows "rsquarefree q & (\a. (poly q a = 0) = (poly p a = 0))" +proof - + from `pderiv p \ 0` have "p \ 0" by auto + with `p = q * d` have "q \ 0" by simp + have "\a. order a q = (if order a p = 0 then 0 else 1)" + using assms by (rule poly_squarefree_decomp_order2) + with `p \ 0` `q \ 0` show ?thesis + by (simp add: rsquarefree_def order_root) +qed + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Polynomial.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Polynomial.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,1463 @@ +(* Title: HOL/Polynomial.thy + Author: Brian Huffman + Based on an earlier development by Clemens Ballarin +*) + +header {* Univariate Polynomials *} + +theory Polynomial +imports Plain SetInterval Main +begin + +subsection {* Definition of type @{text poly} *} + +typedef (Poly) 'a poly = "{f::nat \ 'a::zero. \n. \i>n. f i = 0}" + morphisms coeff Abs_poly + by auto + +lemma expand_poly_eq: "p = q \ (\n. coeff p n = coeff q n)" +by (simp add: coeff_inject [symmetric] expand_fun_eq) + +lemma poly_ext: "(\n. coeff p n = coeff q n) \ p = q" +by (simp add: expand_poly_eq) + + +subsection {* Degree of a polynomial *} + +definition + degree :: "'a::zero poly \ nat" where + "degree p = (LEAST n. \i>n. coeff p i = 0)" + +lemma coeff_eq_0: "degree p < n \ coeff p n = 0" +proof - + have "coeff p \ Poly" + by (rule coeff) + hence "\n. \i>n. coeff p i = 0" + unfolding Poly_def by simp + hence "\i>degree p. coeff p i = 0" + unfolding degree_def by (rule LeastI_ex) + moreover assume "degree p < n" + ultimately show ?thesis by simp +qed + +lemma le_degree: "coeff p n \ 0 \ n \ degree p" + by (erule contrapos_np, rule coeff_eq_0, simp) + +lemma degree_le: "\i>n. coeff p i = 0 \ degree p \ n" + unfolding degree_def by (erule Least_le) + +lemma less_degree_imp: "n < degree p \ \i>n. coeff p i \ 0" + unfolding degree_def by (drule not_less_Least, simp) + + +subsection {* The zero polynomial *} + +instantiation poly :: (zero) zero +begin + +definition + zero_poly_def: "0 = Abs_poly (\n. 0)" + +instance .. +end + +lemma coeff_0 [simp]: "coeff 0 n = 0" + unfolding zero_poly_def + by (simp add: Abs_poly_inverse Poly_def) + +lemma degree_0 [simp]: "degree 0 = 0" + by (rule order_antisym [OF degree_le le0]) simp + +lemma leading_coeff_neq_0: + assumes "p \ 0" shows "coeff p (degree p) \ 0" +proof (cases "degree p") + case 0 + from `p \ 0` have "\n. coeff p n \ 0" + by (simp add: expand_poly_eq) + then obtain n where "coeff p n \ 0" .. + hence "n \ degree p" by (rule le_degree) + with `coeff p n \ 0` and `degree p = 0` + show "coeff p (degree p) \ 0" by simp +next + case (Suc n) + from `degree p = Suc n` have "n < degree p" by simp + hence "\i>n. coeff p i \ 0" by (rule less_degree_imp) + then obtain i where "n < i" and "coeff p i \ 0" by fast + from `degree p = Suc n` and `n < i` have "degree p \ i" by simp + also from `coeff p i \ 0` have "i \ degree p" by (rule le_degree) + finally have "degree p = i" . + with `coeff p i \ 0` show "coeff p (degree p) \ 0" by simp +qed + +lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \ p = 0" + by (cases "p = 0", simp, simp add: leading_coeff_neq_0) + + +subsection {* List-style constructor for polynomials *} + +definition + pCons :: "'a::zero \ 'a poly \ 'a poly" +where + [code del]: "pCons a p = Abs_poly (nat_case a (coeff p))" + +syntax + "_poly" :: "args \ 'a poly" ("[:(_):]") + +translations + "[:x, xs:]" == "CONST pCons x [:xs:]" + "[:x:]" == "CONST pCons x 0" + "[:x:]" <= "CONST pCons x (_constrain 0 t)" + +lemma Poly_nat_case: "f \ Poly \ nat_case a f \ Poly" + unfolding Poly_def by (auto split: nat.split) + +lemma coeff_pCons: + "coeff (pCons a p) = nat_case a (coeff p)" + unfolding pCons_def + by (simp add: Abs_poly_inverse Poly_nat_case coeff) + +lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a" + by (simp add: coeff_pCons) + +lemma coeff_pCons_Suc [simp]: "coeff (pCons a p) (Suc n) = coeff p n" + by (simp add: coeff_pCons) + +lemma degree_pCons_le: "degree (pCons a p) \ Suc (degree p)" +by (rule degree_le, simp add: coeff_eq_0 coeff_pCons split: nat.split) + +lemma degree_pCons_eq: + "p \ 0 \ degree (pCons a p) = Suc (degree p)" +apply (rule order_antisym [OF degree_pCons_le]) +apply (rule le_degree, simp) +done + +lemma degree_pCons_0: "degree (pCons a 0) = 0" +apply (rule order_antisym [OF _ le0]) +apply (rule degree_le, simp add: coeff_pCons split: nat.split) +done + +lemma degree_pCons_eq_if [simp]: + "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" +apply (cases "p = 0", simp_all) +apply (rule order_antisym [OF _ le0]) +apply (rule degree_le, simp add: coeff_pCons split: nat.split) +apply (rule order_antisym [OF degree_pCons_le]) +apply (rule le_degree, simp) +done + +lemma pCons_0_0 [simp]: "pCons 0 0 = 0" +by (rule poly_ext, simp add: coeff_pCons split: nat.split) + +lemma pCons_eq_iff [simp]: + "pCons a p = pCons b q \ a = b \ p = q" +proof (safe) + assume "pCons a p = pCons b q" + then have "coeff (pCons a p) 0 = coeff (pCons b q) 0" by simp + then show "a = b" by simp +next + assume "pCons a p = pCons b q" + then have "\n. coeff (pCons a p) (Suc n) = + coeff (pCons b q) (Suc n)" by simp + then show "p = q" by (simp add: expand_poly_eq) +qed + +lemma pCons_eq_0_iff [simp]: "pCons a p = 0 \ a = 0 \ p = 0" + using pCons_eq_iff [of a p 0 0] by simp + +lemma Poly_Suc: "f \ Poly \ (\n. f (Suc n)) \ Poly" + unfolding Poly_def + by (clarify, rule_tac x=n in exI, simp) + +lemma pCons_cases [cases type: poly]: + obtains (pCons) a q where "p = pCons a q" +proof + show "p = pCons (coeff p 0) (Abs_poly (\n. coeff p (Suc n)))" + by (rule poly_ext) + (simp add: Abs_poly_inverse Poly_Suc coeff coeff_pCons + split: nat.split) +qed + +lemma pCons_induct [case_names 0 pCons, induct type: poly]: + assumes zero: "P 0" + assumes pCons: "\a p. P p \ P (pCons a p)" + shows "P p" +proof (induct p rule: measure_induct_rule [where f=degree]) + case (less p) + obtain a q where "p = pCons a q" by (rule pCons_cases) + have "P q" + proof (cases "q = 0") + case True + then show "P q" by (simp add: zero) + next + case False + then have "degree (pCons a q) = Suc (degree q)" + by (rule degree_pCons_eq) + then have "degree q < degree p" + using `p = pCons a q` by simp + then show "P q" + by (rule less.hyps) + qed + then have "P (pCons a q)" + by (rule pCons) + then show ?case + using `p = pCons a q` by simp +qed + + +subsection {* Recursion combinator for polynomials *} + +function + poly_rec :: "'b \ ('a::zero \ 'a poly \ 'b \ 'b) \ 'a poly \ 'b" +where + poly_rec_pCons_eq_if [simp del, code del]: + "poly_rec z f (pCons a p) = f a p (if p = 0 then z else poly_rec z f p)" +by (case_tac x, rename_tac q, case_tac q, auto) + +termination poly_rec +by (relation "measure (degree \ snd \ snd)", simp) + (simp add: degree_pCons_eq) + +lemma poly_rec_0: + "f 0 0 z = z \ poly_rec z f 0 = z" + using poly_rec_pCons_eq_if [of z f 0 0] by simp + +lemma poly_rec_pCons: + "f 0 0 z = z \ poly_rec z f (pCons a p) = f a p (poly_rec z f p)" + by (simp add: poly_rec_pCons_eq_if poly_rec_0) + + +subsection {* Monomials *} + +definition + monom :: "'a \ nat \ 'a::zero poly" where + "monom a m = Abs_poly (\n. if m = n then a else 0)" + +lemma coeff_monom [simp]: "coeff (monom a m) n = (if m=n then a else 0)" + unfolding monom_def + by (subst Abs_poly_inverse, auto simp add: Poly_def) + +lemma monom_0: "monom a 0 = pCons a 0" + by (rule poly_ext, simp add: coeff_pCons split: nat.split) + +lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)" + by (rule poly_ext, simp add: coeff_pCons split: nat.split) + +lemma monom_eq_0 [simp]: "monom 0 n = 0" + by (rule poly_ext) simp + +lemma monom_eq_0_iff [simp]: "monom a n = 0 \ a = 0" + by (simp add: expand_poly_eq) + +lemma monom_eq_iff [simp]: "monom a n = monom b n \ a = b" + by (simp add: expand_poly_eq) + +lemma degree_monom_le: "degree (monom a n) \ n" + by (rule degree_le, simp) + +lemma degree_monom_eq: "a \ 0 \ degree (monom a n) = n" + apply (rule order_antisym [OF degree_monom_le]) + apply (rule le_degree, simp) + done + + +subsection {* Addition and subtraction *} + +instantiation poly :: (comm_monoid_add) comm_monoid_add +begin + +definition + plus_poly_def [code del]: + "p + q = Abs_poly (\n. coeff p n + coeff q n)" + +lemma Poly_add: + fixes f g :: "nat \ 'a" + shows "\f \ Poly; g \ Poly\ \ (\n. f n + g n) \ Poly" + unfolding Poly_def + apply (clarify, rename_tac m n) + apply (rule_tac x="max m n" in exI, simp) + done + +lemma coeff_add [simp]: + "coeff (p + q) n = coeff p n + coeff q n" + unfolding plus_poly_def + by (simp add: Abs_poly_inverse coeff Poly_add) + +instance proof + fix p q r :: "'a poly" + show "(p + q) + r = p + (q + r)" + by (simp add: expand_poly_eq add_assoc) + show "p + q = q + p" + by (simp add: expand_poly_eq add_commute) + show "0 + p = p" + by (simp add: expand_poly_eq) +qed + +end + +instance poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add +proof + fix p q r :: "'a poly" + assume "p + q = p + r" thus "q = r" + by (simp add: expand_poly_eq) +qed + +instantiation poly :: (ab_group_add) ab_group_add +begin + +definition + uminus_poly_def [code del]: + "- p = Abs_poly (\n. - coeff p n)" + +definition + minus_poly_def [code del]: + "p - q = Abs_poly (\n. coeff p n - coeff q n)" + +lemma Poly_minus: + fixes f :: "nat \ 'a" + shows "f \ Poly \ (\n. - f n) \ Poly" + unfolding Poly_def by simp + +lemma Poly_diff: + fixes f g :: "nat \ 'a" + shows "\f \ Poly; g \ Poly\ \ (\n. f n - g n) \ Poly" + unfolding diff_minus by (simp add: Poly_add Poly_minus) + +lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n" + unfolding uminus_poly_def + by (simp add: Abs_poly_inverse coeff Poly_minus) + +lemma coeff_diff [simp]: + "coeff (p - q) n = coeff p n - coeff q n" + unfolding minus_poly_def + by (simp add: Abs_poly_inverse coeff Poly_diff) + +instance proof + fix p q :: "'a poly" + show "- p + p = 0" + by (simp add: expand_poly_eq) + show "p - q = p + - q" + by (simp add: expand_poly_eq diff_minus) +qed + +end + +lemma add_pCons [simp]: + "pCons a p + pCons b q = pCons (a + b) (p + q)" + by (rule poly_ext, simp add: coeff_pCons split: nat.split) + +lemma minus_pCons [simp]: + "- pCons a p = pCons (- a) (- p)" + by (rule poly_ext, simp add: coeff_pCons split: nat.split) + +lemma diff_pCons [simp]: + "pCons a p - pCons b q = pCons (a - b) (p - q)" + by (rule poly_ext, simp add: coeff_pCons split: nat.split) + +lemma degree_add_le_max: "degree (p + q) \ max (degree p) (degree q)" + by (rule degree_le, auto simp add: coeff_eq_0) + +lemma degree_add_le: + "\degree p \ n; degree q \ n\ \ degree (p + q) \ n" + by (auto intro: order_trans degree_add_le_max) + +lemma degree_add_less: + "\degree p < n; degree q < n\ \ degree (p + q) < n" + by (auto intro: le_less_trans degree_add_le_max) + +lemma degree_add_eq_right: + "degree p < degree q \ degree (p + q) = degree q" + apply (cases "q = 0", simp) + apply (rule order_antisym) + apply (simp add: degree_add_le) + apply (rule le_degree) + apply (simp add: coeff_eq_0) + done + +lemma degree_add_eq_left: + "degree q < degree p \ degree (p + q) = degree p" + using degree_add_eq_right [of q p] + by (simp add: add_commute) + +lemma degree_minus [simp]: "degree (- p) = degree p" + unfolding degree_def by simp + +lemma degree_diff_le_max: "degree (p - q) \ max (degree p) (degree q)" + using degree_add_le [where p=p and q="-q"] + by (simp add: diff_minus) + +lemma degree_diff_le: + "\degree p \ n; degree q \ n\ \ degree (p - q) \ n" + by (simp add: diff_minus degree_add_le) + +lemma degree_diff_less: + "\degree p < n; degree q < n\ \ degree (p - q) < n" + by (simp add: diff_minus degree_add_less) + +lemma add_monom: "monom a n + monom b n = monom (a + b) n" + by (rule poly_ext) simp + +lemma diff_monom: "monom a n - monom b n = monom (a - b) n" + by (rule poly_ext) simp + +lemma minus_monom: "- monom a n = monom (-a) n" + by (rule poly_ext) simp + +lemma coeff_setsum: "coeff (\x\A. p x) i = (\x\A. coeff (p x) i)" + by (cases "finite A", induct set: finite, simp_all) + +lemma monom_setsum: "monom (\x\A. a x) n = (\x\A. monom (a x) n)" + by (rule poly_ext) (simp add: coeff_setsum) + + +subsection {* Multiplication by a constant *} + +definition + smult :: "'a::comm_semiring_0 \ 'a poly \ 'a poly" where + "smult a p = Abs_poly (\n. a * coeff p n)" + +lemma Poly_smult: + fixes f :: "nat \ 'a::comm_semiring_0" + shows "f \ Poly \ (\n. a * f n) \ Poly" + unfolding Poly_def + by (clarify, rule_tac x=n in exI, simp) + +lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n" + unfolding smult_def + by (simp add: Abs_poly_inverse Poly_smult coeff) + +lemma degree_smult_le: "degree (smult a p) \ degree p" + by (rule degree_le, simp add: coeff_eq_0) + +lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p" + by (rule poly_ext, simp add: mult_assoc) + +lemma smult_0_right [simp]: "smult a 0 = 0" + by (rule poly_ext, simp) + +lemma smult_0_left [simp]: "smult 0 p = 0" + by (rule poly_ext, simp) + +lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p" + by (rule poly_ext, simp) + +lemma smult_add_right: + "smult a (p + q) = smult a p + smult a q" + by (rule poly_ext, simp add: algebra_simps) + +lemma smult_add_left: + "smult (a + b) p = smult a p + smult b p" + by (rule poly_ext, simp add: algebra_simps) + +lemma smult_minus_right [simp]: + "smult (a::'a::comm_ring) (- p) = - smult a p" + by (rule poly_ext, simp) + +lemma smult_minus_left [simp]: + "smult (- a::'a::comm_ring) p = - smult a p" + by (rule poly_ext, simp) + +lemma smult_diff_right: + "smult (a::'a::comm_ring) (p - q) = smult a p - smult a q" + by (rule poly_ext, simp add: algebra_simps) + +lemma smult_diff_left: + "smult (a - b::'a::comm_ring) p = smult a p - smult b p" + by (rule poly_ext, simp add: algebra_simps) + +lemmas smult_distribs = + smult_add_left smult_add_right + smult_diff_left smult_diff_right + +lemma smult_pCons [simp]: + "smult a (pCons b p) = pCons (a * b) (smult a p)" + by (rule poly_ext, simp add: coeff_pCons split: nat.split) + +lemma smult_monom: "smult a (monom b n) = monom (a * b) n" + by (induct n, simp add: monom_0, simp add: monom_Suc) + +lemma degree_smult_eq [simp]: + fixes a :: "'a::idom" + shows "degree (smult a p) = (if a = 0 then 0 else degree p)" + by (cases "a = 0", simp, simp add: degree_def) + +lemma smult_eq_0_iff [simp]: + fixes a :: "'a::idom" + shows "smult a p = 0 \ a = 0 \ p = 0" + by (simp add: expand_poly_eq) + + +subsection {* Multiplication of polynomials *} + +text {* TODO: move to SetInterval.thy *} +lemma setsum_atMost_Suc_shift: + fixes f :: "nat \ 'a::comm_monoid_add" + shows "(\i\Suc n. f i) = f 0 + (\i\n. f (Suc i))" +proof (induct n) + case 0 show ?case by simp +next + case (Suc n) note IH = this + have "(\i\Suc (Suc n). f i) = (\i\Suc n. f i) + f (Suc (Suc n))" + by (rule setsum_atMost_Suc) + also have "(\i\Suc n. f i) = f 0 + (\i\n. f (Suc i))" + by (rule IH) + also have "f 0 + (\i\n. f (Suc i)) + f (Suc (Suc n)) = + f 0 + ((\i\n. f (Suc i)) + f (Suc (Suc n)))" + by (rule add_assoc) + also have "(\i\n. f (Suc i)) + f (Suc (Suc n)) = (\i\Suc n. f (Suc i))" + by (rule setsum_atMost_Suc [symmetric]) + finally show ?case . +qed + +instantiation poly :: (comm_semiring_0) comm_semiring_0 +begin + +definition + times_poly_def [code del]: + "p * q = poly_rec 0 (\a p pq. smult a q + pCons 0 pq) p" + +lemma mult_poly_0_left: "(0::'a poly) * q = 0" + unfolding times_poly_def by (simp add: poly_rec_0) + +lemma mult_pCons_left [simp]: + "pCons a p * q = smult a q + pCons 0 (p * q)" + unfolding times_poly_def by (simp add: poly_rec_pCons) + +lemma mult_poly_0_right: "p * (0::'a poly) = 0" + by (induct p, simp add: mult_poly_0_left, simp) + +lemma mult_pCons_right [simp]: + "p * pCons a q = smult a p + pCons 0 (p * q)" + by (induct p, simp add: mult_poly_0_left, simp add: algebra_simps) + +lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right + +lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)" + by (induct p, simp add: mult_poly_0, simp add: smult_add_right) + +lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)" + by (induct q, simp add: mult_poly_0, simp add: smult_add_right) + +lemma mult_poly_add_left: + fixes p q r :: "'a poly" + shows "(p + q) * r = p * r + q * r" + by (induct r, simp add: mult_poly_0, + simp add: smult_distribs algebra_simps) + +instance proof + fix p q r :: "'a poly" + show 0: "0 * p = 0" + by (rule mult_poly_0_left) + show "p * 0 = 0" + by (rule mult_poly_0_right) + show "(p + q) * r = p * r + q * r" + by (rule mult_poly_add_left) + show "(p * q) * r = p * (q * r)" + by (induct p, simp add: mult_poly_0, simp add: mult_poly_add_left) + show "p * q = q * p" + by (induct p, simp add: mult_poly_0, simp) +qed + +end + +instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. + +lemma coeff_mult: + "coeff (p * q) n = (\i\n. coeff p i * coeff q (n-i))" +proof (induct p arbitrary: n) + case 0 show ?case by simp +next + case (pCons a p n) thus ?case + by (cases n, simp, simp add: setsum_atMost_Suc_shift + del: setsum_atMost_Suc) +qed + +lemma degree_mult_le: "degree (p * q) \ degree p + degree q" +apply (rule degree_le) +apply (induct p) +apply simp +apply (simp add: coeff_eq_0 coeff_pCons split: nat.split) +done + +lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)" + by (induct m, simp add: monom_0 smult_monom, simp add: monom_Suc) + + +subsection {* The unit polynomial and exponentiation *} + +instantiation poly :: (comm_semiring_1) comm_semiring_1 +begin + +definition + one_poly_def: + "1 = pCons 1 0" + +instance proof + fix p :: "'a poly" show "1 * p = p" + unfolding one_poly_def + by simp +next + show "0 \ (1::'a poly)" + unfolding one_poly_def by simp +qed + +end + +instance poly :: (comm_semiring_1_cancel) comm_semiring_1_cancel .. + +lemma coeff_1 [simp]: "coeff 1 n = (if n = 0 then 1 else 0)" + unfolding one_poly_def + by (simp add: coeff_pCons split: nat.split) + +lemma degree_1 [simp]: "degree 1 = 0" + unfolding one_poly_def + by (rule degree_pCons_0) + +text {* Lemmas about divisibility *} + +lemma dvd_smult: "p dvd q \ p dvd smult a q" +proof - + assume "p dvd q" + then obtain k where "q = p * k" .. + then have "smult a q = p * smult a k" by simp + then show "p dvd smult a q" .. +qed + +lemma dvd_smult_cancel: + fixes a :: "'a::field" + shows "p dvd smult a q \ a \ 0 \ p dvd q" + by (drule dvd_smult [where a="inverse a"]) simp + +lemma dvd_smult_iff: + fixes a :: "'a::field" + shows "a \ 0 \ p dvd smult a q \ p dvd q" + by (safe elim!: dvd_smult dvd_smult_cancel) + +instantiation poly :: (comm_semiring_1) recpower +begin + +primrec power_poly where + power_poly_0: "(p::'a poly) ^ 0 = 1" +| power_poly_Suc: "(p::'a poly) ^ (Suc n) = p * p ^ n" + +instance + by default simp_all + +end + +lemma degree_power_le: "degree (p ^ n) \ degree p * n" +by (induct n, simp, auto intro: order_trans degree_mult_le) + +instance poly :: (comm_ring) comm_ring .. + +instance poly :: (comm_ring_1) comm_ring_1 .. + +instantiation poly :: (comm_ring_1) number_ring +begin + +definition + "number_of k = (of_int k :: 'a poly)" + +instance + by default (rule number_of_poly_def) + +end + + +subsection {* Polynomials form an integral domain *} + +lemma coeff_mult_degree_sum: + "coeff (p * q) (degree p + degree q) = + coeff p (degree p) * coeff q (degree q)" + by (induct p, simp, simp add: coeff_eq_0) + +instance poly :: (idom) idom +proof + fix p q :: "'a poly" + assume "p \ 0" and "q \ 0" + have "coeff (p * q) (degree p + degree q) = + coeff p (degree p) * coeff q (degree q)" + by (rule coeff_mult_degree_sum) + also have "coeff p (degree p) * coeff q (degree q) \ 0" + using `p \ 0` and `q \ 0` by simp + finally have "\n. coeff (p * q) n \ 0" .. + thus "p * q \ 0" by (simp add: expand_poly_eq) +qed + +lemma degree_mult_eq: + fixes p q :: "'a::idom poly" + shows "\p \ 0; q \ 0\ \ degree (p * q) = degree p + degree q" +apply (rule order_antisym [OF degree_mult_le le_degree]) +apply (simp add: coeff_mult_degree_sum) +done + +lemma dvd_imp_degree_le: + fixes p q :: "'a::idom poly" + shows "\p dvd q; q \ 0\ \ degree p \ degree q" + by (erule dvdE, simp add: degree_mult_eq) + + +subsection {* Polynomials form an ordered integral domain *} + +definition + pos_poly :: "'a::ordered_idom poly \ bool" +where + "pos_poly p \ 0 < coeff p (degree p)" + +lemma pos_poly_pCons: + "pos_poly (pCons a p) \ pos_poly p \ (p = 0 \ 0 < a)" + unfolding pos_poly_def by simp + +lemma not_pos_poly_0 [simp]: "\ pos_poly 0" + unfolding pos_poly_def by simp + +lemma pos_poly_add: "\pos_poly p; pos_poly q\ \ pos_poly (p + q)" + apply (induct p arbitrary: q, simp) + apply (case_tac q, force simp add: pos_poly_pCons add_pos_pos) + done + +lemma pos_poly_mult: "\pos_poly p; pos_poly q\ \ pos_poly (p * q)" + unfolding pos_poly_def + apply (subgoal_tac "p \ 0 \ q \ 0") + apply (simp add: degree_mult_eq coeff_mult_degree_sum mult_pos_pos) + apply auto + done + +lemma pos_poly_total: "p = 0 \ pos_poly p \ pos_poly (- p)" +by (induct p) (auto simp add: pos_poly_pCons) + +instantiation poly :: (ordered_idom) ordered_idom +begin + +definition + [code del]: + "x < y \ pos_poly (y - x)" + +definition + [code del]: + "x \ y \ x = y \ pos_poly (y - x)" + +definition + [code del]: + "abs (x::'a poly) = (if x < 0 then - x else x)" + +definition + [code del]: + "sgn (x::'a poly) = (if x = 0 then 0 else if 0 < x then 1 else - 1)" + +instance proof + fix x y :: "'a poly" + show "x < y \ x \ y \ \ y \ x" + unfolding less_eq_poly_def less_poly_def + apply safe + apply simp + apply (drule (1) pos_poly_add) + apply simp + done +next + fix x :: "'a poly" show "x \ x" + unfolding less_eq_poly_def by simp +next + fix x y z :: "'a poly" + assume "x \ y" and "y \ z" thus "x \ z" + unfolding less_eq_poly_def + apply safe + apply (drule (1) pos_poly_add) + apply (simp add: algebra_simps) + done +next + fix x y :: "'a poly" + assume "x \ y" and "y \ x" thus "x = y" + unfolding less_eq_poly_def + apply safe + apply (drule (1) pos_poly_add) + apply simp + done +next + fix x y z :: "'a poly" + assume "x \ y" thus "z + x \ z + y" + unfolding less_eq_poly_def + apply safe + apply (simp add: algebra_simps) + done +next + fix x y :: "'a poly" + show "x \ y \ y \ x" + unfolding less_eq_poly_def + using pos_poly_total [of "x - y"] + by auto +next + fix x y z :: "'a poly" + assume "x < y" and "0 < z" + thus "z * x < z * y" + unfolding less_poly_def + by (simp add: right_diff_distrib [symmetric] pos_poly_mult) +next + fix x :: "'a poly" + show "\x\ = (if x < 0 then - x else x)" + by (rule abs_poly_def) +next + fix x :: "'a poly" + show "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" + by (rule sgn_poly_def) +qed + +end + +text {* TODO: Simplification rules for comparisons *} + + +subsection {* Long division of polynomials *} + +definition + pdivmod_rel :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly \ bool" +where + [code del]: + "pdivmod_rel x y q r \ + x = q * y + r \ (if y = 0 then q = 0 else r = 0 \ degree r < degree y)" + +lemma pdivmod_rel_0: + "pdivmod_rel 0 y 0 0" + unfolding pdivmod_rel_def by simp + +lemma pdivmod_rel_by_0: + "pdivmod_rel x 0 0 x" + unfolding pdivmod_rel_def by simp + +lemma eq_zero_or_degree_less: + assumes "degree p \ n" and "coeff p n = 0" + shows "p = 0 \ degree p < n" +proof (cases n) + case 0 + with `degree p \ n` and `coeff p n = 0` + have "coeff p (degree p) = 0" by simp + then have "p = 0" by simp + then show ?thesis .. +next + case (Suc m) + have "\i>n. coeff p i = 0" + using `degree p \ n` by (simp add: coeff_eq_0) + then have "\i\n. coeff p i = 0" + using `coeff p n = 0` by (simp add: le_less) + then have "\i>m. coeff p i = 0" + using `n = Suc m` by (simp add: less_eq_Suc_le) + then have "degree p \ m" + by (rule degree_le) + then have "degree p < n" + using `n = Suc m` by (simp add: less_Suc_eq_le) + then show ?thesis .. +qed + +lemma pdivmod_rel_pCons: + assumes rel: "pdivmod_rel x y q r" + assumes y: "y \ 0" + assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" + shows "pdivmod_rel (pCons a x) y (pCons b q) (pCons a r - smult b y)" + (is "pdivmod_rel ?x y ?q ?r") +proof - + have x: "x = q * y + r" and r: "r = 0 \ degree r < degree y" + using assms unfolding pdivmod_rel_def by simp_all + + have 1: "?x = ?q * y + ?r" + using b x by simp + + have 2: "?r = 0 \ degree ?r < degree y" + proof (rule eq_zero_or_degree_less) + show "degree ?r \ degree y" + proof (rule degree_diff_le) + show "degree (pCons a r) \ degree y" + using r by auto + show "degree (smult b y) \ degree y" + by (rule degree_smult_le) + qed + next + show "coeff ?r (degree y) = 0" + using `y \ 0` unfolding b by simp + qed + + from 1 2 show ?thesis + unfolding pdivmod_rel_def + using `y \ 0` by simp +qed + +lemma pdivmod_rel_exists: "\q r. pdivmod_rel x y q r" +apply (cases "y = 0") +apply (fast intro!: pdivmod_rel_by_0) +apply (induct x) +apply (fast intro!: pdivmod_rel_0) +apply (fast intro!: pdivmod_rel_pCons) +done + +lemma pdivmod_rel_unique: + assumes 1: "pdivmod_rel x y q1 r1" + assumes 2: "pdivmod_rel x y q2 r2" + shows "q1 = q2 \ r1 = r2" +proof (cases "y = 0") + assume "y = 0" with assms show ?thesis + by (simp add: pdivmod_rel_def) +next + assume [simp]: "y \ 0" + from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \ degree r1 < degree y" + unfolding pdivmod_rel_def by simp_all + from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \ degree r2 < degree y" + unfolding pdivmod_rel_def by simp_all + from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" + by (simp add: algebra_simps) + from r1 r2 have r3: "(r2 - r1) = 0 \ degree (r2 - r1) < degree y" + by (auto intro: degree_diff_less) + + show "q1 = q2 \ r1 = r2" + proof (rule ccontr) + assume "\ (q1 = q2 \ r1 = r2)" + with q3 have "q1 \ q2" and "r1 \ r2" by auto + with r3 have "degree (r2 - r1) < degree y" by simp + also have "degree y \ degree (q1 - q2) + degree y" by simp + also have "\ = degree ((q1 - q2) * y)" + using `q1 \ q2` by (simp add: degree_mult_eq) + also have "\ = degree (r2 - r1)" + using q3 by simp + finally have "degree (r2 - r1) < degree (r2 - r1)" . + then show "False" by simp + qed +qed + +lemma pdivmod_rel_0_iff: "pdivmod_rel 0 y q r \ q = 0 \ r = 0" +by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_0) + +lemma pdivmod_rel_by_0_iff: "pdivmod_rel x 0 q r \ q = 0 \ r = x" +by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_by_0) + +lemmas pdivmod_rel_unique_div = + pdivmod_rel_unique [THEN conjunct1, standard] + +lemmas pdivmod_rel_unique_mod = + pdivmod_rel_unique [THEN conjunct2, standard] + +instantiation poly :: (field) ring_div +begin + +definition div_poly where + [code del]: "x div y = (THE q. \r. pdivmod_rel x y q r)" + +definition mod_poly where + [code del]: "x mod y = (THE r. \q. pdivmod_rel x y q r)" + +lemma div_poly_eq: + "pdivmod_rel x y q r \ x div y = q" +unfolding div_poly_def +by (fast elim: pdivmod_rel_unique_div) + +lemma mod_poly_eq: + "pdivmod_rel x y q r \ x mod y = r" +unfolding mod_poly_def +by (fast elim: pdivmod_rel_unique_mod) + +lemma pdivmod_rel: + "pdivmod_rel x y (x div y) (x mod y)" +proof - + from pdivmod_rel_exists + obtain q r where "pdivmod_rel x y q r" by fast + thus ?thesis + by (simp add: div_poly_eq mod_poly_eq) +qed + +instance proof + fix x y :: "'a poly" + show "x div y * y + x mod y = x" + using pdivmod_rel [of x y] + by (simp add: pdivmod_rel_def) +next + fix x :: "'a poly" + have "pdivmod_rel x 0 0 x" + by (rule pdivmod_rel_by_0) + thus "x div 0 = 0" + by (rule div_poly_eq) +next + fix y :: "'a poly" + have "pdivmod_rel 0 y 0 0" + by (rule pdivmod_rel_0) + thus "0 div y = 0" + by (rule div_poly_eq) +next + fix x y z :: "'a poly" + assume "y \ 0" + hence "pdivmod_rel (x + z * y) y (z + x div y) (x mod y)" + using pdivmod_rel [of x y] + by (simp add: pdivmod_rel_def left_distrib) + thus "(x + z * y) div y = z + x div y" + by (rule div_poly_eq) +qed + +end + +lemma degree_mod_less: + "y \ 0 \ x mod y = 0 \ degree (x mod y) < degree y" + using pdivmod_rel [of x y] + unfolding pdivmod_rel_def by simp + +lemma div_poly_less: "degree x < degree y \ x div y = 0" +proof - + assume "degree x < degree y" + hence "pdivmod_rel x y 0 x" + by (simp add: pdivmod_rel_def) + thus "x div y = 0" by (rule div_poly_eq) +qed + +lemma mod_poly_less: "degree x < degree y \ x mod y = x" +proof - + assume "degree x < degree y" + hence "pdivmod_rel x y 0 x" + by (simp add: pdivmod_rel_def) + thus "x mod y = x" by (rule mod_poly_eq) +qed + +lemma pdivmod_rel_smult_left: + "pdivmod_rel x y q r + \ pdivmod_rel (smult a x) y (smult a q) (smult a r)" + unfolding pdivmod_rel_def by (simp add: smult_add_right) + +lemma div_smult_left: "(smult a x) div y = smult a (x div y)" + by (rule div_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel) + +lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)" + by (rule mod_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel) + +lemma poly_div_minus_left [simp]: + fixes x y :: "'a::field poly" + shows "(- x) div y = - (x div y)" + using div_smult_left [of "- 1::'a"] by simp + +lemma poly_mod_minus_left [simp]: + fixes x y :: "'a::field poly" + shows "(- x) mod y = - (x mod y)" + using mod_smult_left [of "- 1::'a"] by simp + +lemma pdivmod_rel_smult_right: + "\a \ 0; pdivmod_rel x y q r\ + \ pdivmod_rel x (smult a y) (smult (inverse a) q) r" + unfolding pdivmod_rel_def by simp + +lemma div_smult_right: + "a \ 0 \ x div (smult a y) = smult (inverse a) (x div y)" + by (rule div_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel) + +lemma mod_smult_right: "a \ 0 \ x mod (smult a y) = x mod y" + by (rule mod_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel) + +lemma poly_div_minus_right [simp]: + fixes x y :: "'a::field poly" + shows "x div (- y) = - (x div y)" + using div_smult_right [of "- 1::'a"] + by (simp add: nonzero_inverse_minus_eq) + +lemma poly_mod_minus_right [simp]: + fixes x y :: "'a::field poly" + shows "x mod (- y) = x mod y" + using mod_smult_right [of "- 1::'a"] by simp + +lemma pdivmod_rel_mult: + "\pdivmod_rel x y q r; pdivmod_rel q z q' r'\ + \ pdivmod_rel x (y * z) q' (y * r' + r)" +apply (cases "z = 0", simp add: pdivmod_rel_def) +apply (cases "y = 0", simp add: pdivmod_rel_by_0_iff pdivmod_rel_0_iff) +apply (cases "r = 0") +apply (cases "r' = 0") +apply (simp add: pdivmod_rel_def) +apply (simp add: pdivmod_rel_def ring_simps degree_mult_eq) +apply (cases "r' = 0") +apply (simp add: pdivmod_rel_def degree_mult_eq) +apply (simp add: pdivmod_rel_def ring_simps) +apply (simp add: degree_mult_eq degree_add_less) +done + +lemma poly_div_mult_right: + fixes x y z :: "'a::field poly" + shows "x div (y * z) = (x div y) div z" + by (rule div_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+) + +lemma poly_mod_mult_right: + fixes x y z :: "'a::field poly" + shows "x mod (y * z) = y * (x div y mod z) + x mod y" + by (rule mod_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+) + +lemma mod_pCons: + fixes a and x + assumes y: "y \ 0" + defines b: "b \ coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" + shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)" +unfolding b +apply (rule mod_poly_eq) +apply (rule pdivmod_rel_pCons [OF pdivmod_rel y refl]) +done + + +subsection {* Evaluation of polynomials *} + +definition + poly :: "'a::comm_semiring_0 poly \ 'a \ 'a" where + "poly = poly_rec (\x. 0) (\a p f x. a + x * f x)" + +lemma poly_0 [simp]: "poly 0 x = 0" + unfolding poly_def by (simp add: poly_rec_0) + +lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x" + unfolding poly_def by (simp add: poly_rec_pCons) + +lemma poly_1 [simp]: "poly 1 x = 1" + unfolding one_poly_def by simp + +lemma poly_monom: + fixes a x :: "'a::{comm_semiring_1,recpower}" + shows "poly (monom a n) x = a * x ^ n" + by (induct n, simp add: monom_0, simp add: monom_Suc power_Suc mult_ac) + +lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" + apply (induct p arbitrary: q, simp) + apply (case_tac q, simp, simp add: algebra_simps) + done + +lemma poly_minus [simp]: + fixes x :: "'a::comm_ring" + shows "poly (- p) x = - poly p x" + by (induct p, simp_all) + +lemma poly_diff [simp]: + fixes x :: "'a::comm_ring" + shows "poly (p - q) x = poly p x - poly q x" + by (simp add: diff_minus) + +lemma poly_setsum: "poly (\k\A. p k) x = (\k\A. poly (p k) x)" + by (cases "finite A", induct set: finite, simp_all) + +lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x" + by (induct p, simp, simp add: algebra_simps) + +lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x" + by (induct p, simp_all, simp add: algebra_simps) + +lemma poly_power [simp]: + fixes p :: "'a::{comm_semiring_1,recpower} poly" + shows "poly (p ^ n) x = poly p x ^ n" + by (induct n, simp, simp add: power_Suc) + + +subsection {* Synthetic division *} + +text {* + Synthetic division is simply division by the + linear polynomial @{term "x - c"}. +*} + +definition + synthetic_divmod :: "'a::comm_semiring_0 poly \ 'a \ 'a poly \ 'a" +where [code del]: + "synthetic_divmod p c = + poly_rec (0, 0) (\a p (q, r). (pCons r q, a + c * r)) p" + +definition + synthetic_div :: "'a::comm_semiring_0 poly \ 'a \ 'a poly" +where + "synthetic_div p c = fst (synthetic_divmod p c)" + +lemma synthetic_divmod_0 [simp]: + "synthetic_divmod 0 c = (0, 0)" + unfolding synthetic_divmod_def + by (simp add: poly_rec_0) + +lemma synthetic_divmod_pCons [simp]: + "synthetic_divmod (pCons a p) c = + (\(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" + unfolding synthetic_divmod_def + by (simp add: poly_rec_pCons) + +lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c" + by (induct p, simp, simp add: split_def) + +lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0" + unfolding synthetic_div_def by simp + +lemma synthetic_div_pCons [simp]: + "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" + unfolding synthetic_div_def + by (simp add: split_def snd_synthetic_divmod) + +lemma synthetic_div_eq_0_iff: + "synthetic_div p c = 0 \ degree p = 0" + by (induct p, simp, case_tac p, simp) + +lemma degree_synthetic_div: + "degree (synthetic_div p c) = degree p - 1" + by (induct p, simp, simp add: synthetic_div_eq_0_iff) + +lemma synthetic_div_correct: + "p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)" + by (induct p) simp_all + +lemma synthetic_div_unique_lemma: "smult c p = pCons a p \ p = 0" +by (induct p arbitrary: a) simp_all + +lemma synthetic_div_unique: + "p + smult c q = pCons r q \ r = poly p c \ q = synthetic_div p c" +apply (induct p arbitrary: q r) +apply (simp, frule synthetic_div_unique_lemma, simp) +apply (case_tac q, force) +done + +lemma synthetic_div_correct': + fixes c :: "'a::comm_ring_1" + shows "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" + using synthetic_div_correct [of p c] + by (simp add: algebra_simps) + +lemma poly_eq_0_iff_dvd: + fixes c :: "'a::idom" + shows "poly p c = 0 \ [:-c, 1:] dvd p" +proof + assume "poly p c = 0" + with synthetic_div_correct' [of c p] + have "p = [:-c, 1:] * synthetic_div p c" by simp + then show "[:-c, 1:] dvd p" .. +next + assume "[:-c, 1:] dvd p" + then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) + then show "poly p c = 0" by simp +qed + +lemma dvd_iff_poly_eq_0: + fixes c :: "'a::idom" + shows "[:c, 1:] dvd p \ poly p (-c) = 0" + by (simp add: poly_eq_0_iff_dvd) + +lemma poly_roots_finite: + fixes p :: "'a::idom poly" + shows "p \ 0 \ finite {x. poly p x = 0}" +proof (induct n \ "degree p" arbitrary: p) + case (0 p) + then obtain a where "a \ 0" and "p = [:a:]" + by (cases p, simp split: if_splits) + then show "finite {x. poly p x = 0}" by simp +next + case (Suc n p) + show "finite {x. poly p x = 0}" + proof (cases "\x. poly p x = 0") + case False + then show "finite {x. poly p x = 0}" by simp + next + case True + then obtain a where "poly p a = 0" .. + then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd) + then obtain k where k: "p = [:-a, 1:] * k" .. + with `p \ 0` have "k \ 0" by auto + with k have "degree p = Suc (degree k)" + by (simp add: degree_mult_eq del: mult_pCons_left) + with `Suc n = degree p` have "n = degree k" by simp + with `k \ 0` have "finite {x. poly k x = 0}" by (rule Suc.hyps) + then have "finite (insert a {x. poly k x = 0})" by simp + then show "finite {x. poly p x = 0}" + by (simp add: k uminus_add_conv_diff Collect_disj_eq + del: mult_pCons_left) + qed +qed + +lemma poly_zero: + fixes p :: "'a::{idom,ring_char_0} poly" + shows "poly p = poly 0 \ p = 0" +apply (cases "p = 0", simp_all) +apply (drule poly_roots_finite) +apply (auto simp add: infinite_UNIV_char_0) +done + +lemma poly_eq_iff: + fixes p q :: "'a::{idom,ring_char_0} poly" + shows "poly p = poly q \ p = q" + using poly_zero [of "p - q"] + by (simp add: expand_fun_eq) + + +subsection {* Composition of polynomials *} + +definition + pcompose :: "'a::comm_semiring_0 poly \ 'a poly \ 'a poly" +where + "pcompose p q = poly_rec 0 (\a _ c. [:a:] + q * c) p" + +lemma pcompose_0 [simp]: "pcompose 0 q = 0" + unfolding pcompose_def by (simp add: poly_rec_0) + +lemma pcompose_pCons: + "pcompose (pCons a p) q = [:a:] + q * pcompose p q" + unfolding pcompose_def by (simp add: poly_rec_pCons) + +lemma poly_pcompose: "poly (pcompose p q) x = poly p (poly q x)" + by (induct p) (simp_all add: pcompose_pCons) + +lemma degree_pcompose_le: + "degree (pcompose p q) \ degree p * degree q" +apply (induct p, simp) +apply (simp add: pcompose_pCons, clarify) +apply (rule degree_add_le, simp) +apply (rule order_trans [OF degree_mult_le], simp) +done + + +subsection {* Order of polynomial roots *} + +definition + order :: "'a::idom \ 'a poly \ nat" +where + [code del]: + "order a p = (LEAST n. \ [:-a, 1:] ^ Suc n dvd p)" + +lemma coeff_linear_power: + fixes a :: "'a::comm_semiring_1" + shows "coeff ([:a, 1:] ^ n) n = 1" +apply (induct n, simp_all) +apply (subst coeff_eq_0) +apply (auto intro: le_less_trans degree_power_le) +done + +lemma degree_linear_power: + fixes a :: "'a::comm_semiring_1" + shows "degree ([:a, 1:] ^ n) = n" +apply (rule order_antisym) +apply (rule ord_le_eq_trans [OF degree_power_le], simp) +apply (rule le_degree, simp add: coeff_linear_power) +done + +lemma order_1: "[:-a, 1:] ^ order a p dvd p" +apply (cases "p = 0", simp) +apply (cases "order a p", simp) +apply (subgoal_tac "nat < (LEAST n. \ [:-a, 1:] ^ Suc n dvd p)") +apply (drule not_less_Least, simp) +apply (fold order_def, simp) +done + +lemma order_2: "p \ 0 \ \ [:-a, 1:] ^ Suc (order a p) dvd p" +unfolding order_def +apply (rule LeastI_ex) +apply (rule_tac x="degree p" in exI) +apply (rule notI) +apply (drule (1) dvd_imp_degree_le) +apply (simp only: degree_linear_power) +done + +lemma order: + "p \ 0 \ [:-a, 1:] ^ order a p dvd p \ \ [:-a, 1:] ^ Suc (order a p) dvd p" +by (rule conjI [OF order_1 order_2]) + +lemma order_degree: + assumes p: "p \ 0" + shows "order a p \ degree p" +proof - + have "order a p = degree ([:-a, 1:] ^ order a p)" + by (simp only: degree_linear_power) + also have "\ \ degree p" + using order_1 p by (rule dvd_imp_degree_le) + finally show ?thesis . +qed + +lemma order_root: "poly p a = 0 \ p = 0 \ order a p \ 0" +apply (cases "p = 0", simp_all) +apply (rule iffI) +apply (rule ccontr, simp) +apply (frule order_2 [where a=a], simp) +apply (simp add: poly_eq_0_iff_dvd) +apply (simp add: poly_eq_0_iff_dvd) +apply (simp only: order_def) +apply (drule not_less_Least, simp) +done + + +subsection {* Configuration of the code generator *} + +code_datatype "0::'a::zero poly" pCons + +declare pCons_0_0 [code post] + +instantiation poly :: ("{zero,eq}") eq +begin + +definition [code del]: + "eq_class.eq (p::'a poly) q \ p = q" + +instance + by default (rule eq_poly_def) + +end + +lemma eq_poly_code [code]: + "eq_class.eq (0::_ poly) (0::_ poly) \ True" + "eq_class.eq (0::_ poly) (pCons b q) \ eq_class.eq 0 b \ eq_class.eq 0 q" + "eq_class.eq (pCons a p) (0::_ poly) \ eq_class.eq a 0 \ eq_class.eq p 0" + "eq_class.eq (pCons a p) (pCons b q) \ eq_class.eq a b \ eq_class.eq p q" +unfolding eq by simp_all + +lemmas coeff_code [code] = + coeff_0 coeff_pCons_0 coeff_pCons_Suc + +lemmas degree_code [code] = + degree_0 degree_pCons_eq_if + +lemmas monom_poly_code [code] = + monom_0 monom_Suc + +lemma add_poly_code [code]: + "0 + q = (q :: _ poly)" + "p + 0 = (p :: _ poly)" + "pCons a p + pCons b q = pCons (a + b) (p + q)" +by simp_all + +lemma minus_poly_code [code]: + "- 0 = (0 :: _ poly)" + "- pCons a p = pCons (- a) (- p)" +by simp_all + +lemma diff_poly_code [code]: + "0 - q = (- q :: _ poly)" + "p - 0 = (p :: _ poly)" + "pCons a p - pCons b q = pCons (a - b) (p - q)" +by simp_all + +lemmas smult_poly_code [code] = + smult_0_right smult_pCons + +lemma mult_poly_code [code]: + "0 * q = (0 :: _ poly)" + "pCons a p * q = smult a q + pCons 0 (p * q)" +by simp_all + +lemmas poly_code [code] = + poly_0 poly_pCons + +lemmas synthetic_divmod_code [code] = + synthetic_divmod_0 synthetic_divmod_pCons + +text {* code generator setup for div and mod *} + +definition + pdivmod :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly" +where + [code del]: "pdivmod x y = (x div y, x mod y)" + +lemma div_poly_code [code]: "x div y = fst (pdivmod x y)" + unfolding pdivmod_def by simp + +lemma mod_poly_code [code]: "x mod y = snd (pdivmod x y)" + unfolding pdivmod_def by simp + +lemma pdivmod_0 [code]: "pdivmod 0 y = (0, 0)" + unfolding pdivmod_def by simp + +lemma pdivmod_pCons [code]: + "pdivmod (pCons a x) y = + (if y = 0 then (0, pCons a x) else + (let (q, r) = pdivmod x y; + b = coeff (pCons a r) (degree y) / coeff y (degree y) + in (pCons b q, pCons a r - smult b y)))" +apply (simp add: pdivmod_def Let_def, safe) +apply (rule div_poly_eq) +apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) +apply (rule mod_poly_eq) +apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) +done + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Primes.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Product_Vector.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Product_Vector.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,273 @@ +(* Title: HOL/Library/Product_Vector.thy + Author: Brian Huffman +*) + +header {* Cartesian Products as Vector Spaces *} + +theory Product_Vector +imports Inner_Product Product_plus +begin + +subsection {* Product is a real vector space *} + +instantiation "*" :: (real_vector, real_vector) real_vector +begin + +definition scaleR_prod_def: + "scaleR r A = (scaleR r (fst A), scaleR r (snd A))" + +lemma fst_scaleR [simp]: "fst (scaleR r A) = scaleR r (fst A)" + unfolding scaleR_prod_def by simp + +lemma snd_scaleR [simp]: "snd (scaleR r A) = scaleR r (snd A)" + unfolding scaleR_prod_def by simp + +lemma scaleR_Pair [simp]: "scaleR r (a, b) = (scaleR r a, scaleR r b)" + unfolding scaleR_prod_def by simp + +instance proof + fix a b :: real and x y :: "'a \ 'b" + show "scaleR a (x + y) = scaleR a x + scaleR a y" + by (simp add: expand_prod_eq scaleR_right_distrib) + show "scaleR (a + b) x = scaleR a x + scaleR b x" + by (simp add: expand_prod_eq scaleR_left_distrib) + show "scaleR a (scaleR b x) = scaleR (a * b) x" + by (simp add: expand_prod_eq) + show "scaleR 1 x = x" + by (simp add: expand_prod_eq) +qed + +end + +subsection {* Product is a normed vector space *} + +instantiation + "*" :: (real_normed_vector, real_normed_vector) real_normed_vector +begin + +definition norm_prod_def: + "norm x = sqrt ((norm (fst x))\ + (norm (snd x))\)" + +definition sgn_prod_def: + "sgn (x::'a \ 'b) = scaleR (inverse (norm x)) x" + +lemma norm_Pair: "norm (a, b) = sqrt ((norm a)\ + (norm b)\)" + unfolding norm_prod_def by simp + +instance proof + fix r :: real and x y :: "'a \ 'b" + show "0 \ norm x" + unfolding norm_prod_def by simp + show "norm x = 0 \ x = 0" + unfolding norm_prod_def + by (simp add: expand_prod_eq) + show "norm (x + y) \ norm x + norm y" + unfolding norm_prod_def + apply (rule order_trans [OF _ real_sqrt_sum_squares_triangle_ineq]) + apply (simp add: add_mono power_mono norm_triangle_ineq) + done + show "norm (scaleR r x) = \r\ * norm x" + unfolding norm_prod_def + apply (simp add: norm_scaleR power_mult_distrib) + apply (simp add: right_distrib [symmetric]) + apply (simp add: real_sqrt_mult_distrib) + done + show "sgn x = scaleR (inverse (norm x)) x" + by (rule sgn_prod_def) +qed + +end + +subsection {* Product is an inner product space *} + +instantiation "*" :: (real_inner, real_inner) real_inner +begin + +definition inner_prod_def: + "inner x y = inner (fst x) (fst y) + inner (snd x) (snd y)" + +lemma inner_Pair [simp]: "inner (a, b) (c, d) = inner a c + inner b d" + unfolding inner_prod_def by simp + +instance proof + fix r :: real + fix x y z :: "'a::real_inner * 'b::real_inner" + show "inner x y = inner y x" + unfolding inner_prod_def + by (simp add: inner_commute) + show "inner (x + y) z = inner x z + inner y z" + unfolding inner_prod_def + by (simp add: inner_left_distrib) + show "inner (scaleR r x) y = r * inner x y" + unfolding inner_prod_def + by (simp add: inner_scaleR_left right_distrib) + show "0 \ inner x x" + unfolding inner_prod_def + by (intro add_nonneg_nonneg inner_ge_zero) + show "inner x x = 0 \ x = 0" + unfolding inner_prod_def expand_prod_eq + by (simp add: add_nonneg_eq_0_iff) + show "norm x = sqrt (inner x x)" + unfolding norm_prod_def inner_prod_def + by (simp add: power2_norm_eq_inner) +qed + +end + +subsection {* Pair operations are linear and continuous *} + +interpretation fst!: bounded_linear fst + apply (unfold_locales) + apply (rule fst_add) + apply (rule fst_scaleR) + apply (rule_tac x="1" in exI, simp add: norm_Pair) + done + +interpretation snd!: bounded_linear snd + apply (unfold_locales) + apply (rule snd_add) + apply (rule snd_scaleR) + apply (rule_tac x="1" in exI, simp add: norm_Pair) + done + +text {* TODO: move to NthRoot *} +lemma sqrt_add_le_add_sqrt: + assumes x: "0 \ x" and y: "0 \ y" + shows "sqrt (x + y) \ sqrt x + sqrt y" +apply (rule power2_le_imp_le) +apply (simp add: real_sum_squared_expand add_nonneg_nonneg x y) +apply (simp add: mult_nonneg_nonneg x y) +apply (simp add: add_nonneg_nonneg x y) +done + +lemma bounded_linear_Pair: + assumes f: "bounded_linear f" + assumes g: "bounded_linear g" + shows "bounded_linear (\x. (f x, g x))" +proof + interpret f: bounded_linear f by fact + interpret g: bounded_linear g by fact + fix x y and r :: real + show "(f (x + y), g (x + y)) = (f x, g x) + (f y, g y)" + by (simp add: f.add g.add) + show "(f (r *\<^sub>R x), g (r *\<^sub>R x)) = r *\<^sub>R (f x, g x)" + by (simp add: f.scaleR g.scaleR) + obtain Kf where "0 < Kf" and norm_f: "\x. norm (f x) \ norm x * Kf" + using f.pos_bounded by fast + obtain Kg where "0 < Kg" and norm_g: "\x. norm (g x) \ norm x * Kg" + using g.pos_bounded by fast + have "\x. norm (f x, g x) \ norm x * (Kf + Kg)" + apply (rule allI) + apply (simp add: norm_Pair) + apply (rule order_trans [OF sqrt_add_le_add_sqrt], simp, simp) + apply (simp add: right_distrib) + apply (rule add_mono [OF norm_f norm_g]) + done + then show "\K. \x. norm (f x, g x) \ norm x * K" .. +qed + +text {* + TODO: The next three proofs are nearly identical to each other. + Is there a good way to factor out the common parts? +*} + +lemma LIMSEQ_Pair: + assumes "X ----> a" and "Y ----> b" + shows "(\n. (X n, Y n)) ----> (a, b)" +proof (rule LIMSEQ_I) + fix r :: real assume "0 < r" + then have "0 < r / sqrt 2" (is "0 < ?s") + by (simp add: divide_pos_pos) + obtain M where M: "\n\M. norm (X n - a) < ?s" + using LIMSEQ_D [OF `X ----> a` `0 < ?s`] .. + obtain N where N: "\n\N. norm (Y n - b) < ?s" + using LIMSEQ_D [OF `Y ----> b` `0 < ?s`] .. + have "\n\max M N. norm ((X n, Y n) - (a, b)) < r" + using M N by (simp add: real_sqrt_sum_squares_less norm_Pair) + then show "\n0. \n\n0. norm ((X n, Y n) - (a, b)) < r" .. +qed + +lemma Cauchy_Pair: + assumes "Cauchy X" and "Cauchy Y" + shows "Cauchy (\n. (X n, Y n))" +proof (rule CauchyI) + fix r :: real assume "0 < r" + then have "0 < r / sqrt 2" (is "0 < ?s") + by (simp add: divide_pos_pos) + obtain M where M: "\m\M. \n\M. norm (X m - X n) < ?s" + using CauchyD [OF `Cauchy X` `0 < ?s`] .. + obtain N where N: "\m\N. \n\N. norm (Y m - Y n) < ?s" + using CauchyD [OF `Cauchy Y` `0 < ?s`] .. + have "\m\max M N. \n\max M N. norm ((X m, Y m) - (X n, Y n)) < r" + using M N by (simp add: real_sqrt_sum_squares_less norm_Pair) + then show "\n0. \m\n0. \n\n0. norm ((X m, Y m) - (X n, Y n)) < r" .. +qed + +lemma LIM_Pair: + assumes "f -- x --> a" and "g -- x --> b" + shows "(\x. (f x, g x)) -- x --> (a, b)" +proof (rule LIM_I) + fix r :: real assume "0 < r" + then have "0 < r / sqrt 2" (is "0 < ?e") + by (simp add: divide_pos_pos) + obtain s where s: "0 < s" + "\z. z \ x \ norm (z - x) < s \ norm (f z - a) < ?e" + using LIM_D [OF `f -- x --> a` `0 < ?e`] by fast + obtain t where t: "0 < t" + "\z. z \ x \ norm (z - x) < t \ norm (g z - b) < ?e" + using LIM_D [OF `g -- x --> b` `0 < ?e`] by fast + have "0 < min s t \ + (\z. z \ x \ norm (z - x) < min s t \ norm ((f z, g z) - (a, b)) < r)" + using s t by (simp add: real_sqrt_sum_squares_less norm_Pair) + then show + "\s>0. \z. z \ x \ norm (z - x) < s \ norm ((f z, g z) - (a, b)) < r" .. +qed + +lemma isCont_Pair [simp]: + "\isCont f x; isCont g x\ \ isCont (\x. (f x, g x)) x" + unfolding isCont_def by (rule LIM_Pair) + + +subsection {* Product is a complete vector space *} + +instance "*" :: (banach, banach) banach +proof + fix X :: "nat \ 'a \ 'b" assume "Cauchy X" + have 1: "(\n. fst (X n)) ----> lim (\n. fst (X n))" + using fst.Cauchy [OF `Cauchy X`] + by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) + have 2: "(\n. snd (X n)) ----> lim (\n. snd (X n))" + using snd.Cauchy [OF `Cauchy X`] + by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) + have "X ----> (lim (\n. fst (X n)), lim (\n. snd (X n)))" + using LIMSEQ_Pair [OF 1 2] by simp + then show "convergent X" + by (rule convergentI) +qed + +subsection {* Frechet derivatives involving pairs *} + +lemma FDERIV_Pair: + assumes f: "FDERIV f x :> f'" and g: "FDERIV g x :> g'" + shows "FDERIV (\x. (f x, g x)) x :> (\h. (f' h, g' h))" +apply (rule FDERIV_I) +apply (rule bounded_linear_Pair) +apply (rule FDERIV_bounded_linear [OF f]) +apply (rule FDERIV_bounded_linear [OF g]) +apply (simp add: norm_Pair) +apply (rule real_LIM_sandwich_zero) +apply (rule LIM_add_zero) +apply (rule FDERIV_D [OF f]) +apply (rule FDERIV_D [OF g]) +apply (rename_tac h) +apply (simp add: divide_nonneg_pos) +apply (rename_tac h) +apply (subst add_divide_distrib [symmetric]) +apply (rule divide_right_mono [OF _ norm_ge_zero]) +apply (rule order_trans [OF sqrt_add_le_add_sqrt]) +apply simp +apply simp +apply simp +done + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/Product_plus.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Product_plus.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,115 @@ +(* Title: HOL/Library/Product_plus.thy + Author: Brian Huffman +*) + +header {* Additive group operations on product types *} + +theory Product_plus +imports Main +begin + +subsection {* Operations *} + +instantiation "*" :: (zero, zero) zero +begin + +definition zero_prod_def: "0 = (0, 0)" + +instance .. +end + +instantiation "*" :: (plus, plus) plus +begin + +definition plus_prod_def: + "x + y = (fst x + fst y, snd x + snd y)" + +instance .. +end + +instantiation "*" :: (minus, minus) minus +begin + +definition minus_prod_def: + "x - y = (fst x - fst y, snd x - snd y)" + +instance .. +end + +instantiation "*" :: (uminus, uminus) uminus +begin + +definition uminus_prod_def: + "- x = (- fst x, - snd x)" + +instance .. +end + +lemma fst_zero [simp]: "fst 0 = 0" + unfolding zero_prod_def by simp + +lemma snd_zero [simp]: "snd 0 = 0" + unfolding zero_prod_def by simp + +lemma fst_add [simp]: "fst (x + y) = fst x + fst y" + unfolding plus_prod_def by simp + +lemma snd_add [simp]: "snd (x + y) = snd x + snd y" + unfolding plus_prod_def by simp + +lemma fst_diff [simp]: "fst (x - y) = fst x - fst y" + unfolding minus_prod_def by simp + +lemma snd_diff [simp]: "snd (x - y) = snd x - snd y" + unfolding minus_prod_def by simp + +lemma fst_uminus [simp]: "fst (- x) = - fst x" + unfolding uminus_prod_def by simp + +lemma snd_uminus [simp]: "snd (- x) = - snd x" + unfolding uminus_prod_def by simp + +lemma add_Pair [simp]: "(a, b) + (c, d) = (a + c, b + d)" + unfolding plus_prod_def by simp + +lemma diff_Pair [simp]: "(a, b) - (c, d) = (a - c, b - d)" + unfolding minus_prod_def by simp + +lemma uminus_Pair [simp, code]: "- (a, b) = (- a, - b)" + unfolding uminus_prod_def by simp + +lemmas expand_prod_eq = Pair_fst_snd_eq + + +subsection {* Class instances *} + +instance "*" :: (semigroup_add, semigroup_add) semigroup_add + by default (simp add: expand_prod_eq add_assoc) + +instance "*" :: (ab_semigroup_add, ab_semigroup_add) ab_semigroup_add + by default (simp add: expand_prod_eq add_commute) + +instance "*" :: (monoid_add, monoid_add) monoid_add + by default (simp_all add: expand_prod_eq) + +instance "*" :: (comm_monoid_add, comm_monoid_add) comm_monoid_add + by default (simp add: expand_prod_eq) + +instance "*" :: + (cancel_semigroup_add, cancel_semigroup_add) cancel_semigroup_add + by default (simp_all add: expand_prod_eq) + +instance "*" :: + (cancel_ab_semigroup_add, cancel_ab_semigroup_add) cancel_ab_semigroup_add + by default (simp add: expand_prod_eq) + +instance "*" :: + (cancel_comm_monoid_add, cancel_comm_monoid_add) cancel_comm_monoid_add .. + +instance "*" :: (group_add, group_add) group_add + by default (simp_all add: expand_prod_eq diff_minus) + +instance "*" :: (ab_group_add, ab_group_add) ab_group_add + by default (simp_all add: expand_prod_eq) + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Library/RBT.thy --- a/src/HOL/Library/RBT.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Library/RBT.thy Wed Mar 04 11:05:29 2009 +0100 @@ -891,7 +891,7 @@ theorem mapwk_isrbt[simp]: "isrbt (mapwithkey f t) = isrbt t" unfolding isrbt_def by (simp add: mapwk_inv1 mapwk_inv2 mapwk_st mapwk_treec) -theorem map_of_mapwk[simp]: "map_of (mapwithkey f t) x = option_map (f x) (map_of t x)" +theorem map_of_mapwk[simp]: "map_of (mapwithkey f t) x = Option.map (f x) (map_of t x)" by (induct t) auto definition map @@ -899,7 +899,7 @@ theorem map_keys[simp]: "keys (map f t) = keys t" unfolding map_def by simp theorem map_isrbt[simp]: "isrbt (map f t) = isrbt t" unfolding map_def by simp -theorem map_of_map[simp]: "map_of (map f t) = option_map f o map_of t" +theorem map_of_map[simp]: "map_of (map f t) = Option.map f o map_of t" by (rule ext) (simp add:map_def) subsection {* Fold *} diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/List.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Map.thy --- a/src/HOL/Map.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Map.thy Wed Mar 04 11:05:29 2009 +0100 @@ -242,17 +242,17 @@ "map_of xs k = Some z \ P k z \ map_of (filter (split P) xs) k = Some z" by (induct xs) auto -lemma map_of_map: "map_of (map (%(a,b). (a,f b)) xs) x = option_map f (map_of xs x)" +lemma map_of_map: "map_of (map (%(a,b). (a,f b)) xs) x = Option.map f (map_of xs x)" by (induct xs) auto -subsection {* @{term [source] option_map} related *} +subsection {* @{const Option.map} related *} -lemma option_map_o_empty [simp]: "option_map f o empty = empty" +lemma option_map_o_empty [simp]: "Option.map f o empty = empty" by (rule ext) simp lemma option_map_o_map_upd [simp]: - "option_map f o m(a|->b) = (option_map f o m)(a|->f b)" + "Option.map f o m(a|->b) = (Option.map f o m)(a|->f b)" by (rule ext) simp diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/MicroJava/BV/Effect.thy --- a/src/HOL/MicroJava/BV/Effect.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/MicroJava/BV/Effect.thy Wed Mar 04 11:05:29 2009 +0100 @@ -105,7 +105,7 @@ (xcpt_names (i,G,pc,et))" norm_eff :: "instr \ jvm_prog \ state_type option \ state_type option" - "norm_eff i G == option_map (\s. eff' (i,G,s))" + "norm_eff i G == Option.map (\s. eff' (i,G,s))" eff :: "instr \ jvm_prog \ p_count \ exception_table \ state_type option \ succ_type" "eff i G pc et s == (map (\pc'. (pc',norm_eff i G s)) (succs i pc)) @ (xcpt_eff i G pc s et)" diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/MicroJava/BV/Opt.thy --- a/src/HOL/MicroJava/BV/Opt.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/MicroJava/BV/Opt.thy Wed Mar 04 11:05:29 2009 +0100 @@ -286,8 +286,8 @@ lemma option_map_in_optionI: "\ ox : opt S; !x:S. ox = Some x \ f x : S \ - \ option_map f ox : opt S"; -apply (unfold option_map_def) + \ Option.map f ox : opt S"; +apply (unfold Option.map_def) apply (simp split: option.split) apply blast done diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/MicroJava/Comp/AuxLemmas.thy --- a/src/HOL/MicroJava/Comp/AuxLemmas.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/MicroJava/Comp/AuxLemmas.thy Wed Mar 04 11:05:29 2009 +0100 @@ -126,11 +126,11 @@ by (induct xs,auto) lemma map_of_map2: "\ x \ set xs. (fst (f x)) = (fst x) \ - map_of (map f xs) a = option_map (\ b. (snd (f (a, b)))) (map_of xs a)" + map_of (map f xs) a = Option.map (\ b. (snd (f (a, b)))) (map_of xs a)" by (induct xs, auto) -lemma option_map_of [simp]: "(option_map f (map_of xs k) = None) = ((map_of xs k) = None)" -by (simp add: option_map_def split: option.split) +lemma option_map_of [simp]: "(Option.map f (map_of xs k) = None) = ((map_of xs k) = None)" +by (simp add: Option.map_def split: option.split) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/MicroJava/Comp/CorrCompTp.thy --- a/src/HOL/MicroJava/Comp/CorrCompTp.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy Wed Mar 04 11:05:29 2009 +0100 @@ -553,7 +553,7 @@ lemma match_xctable_offset: " (match_exception_table G cn (pc + n) (offset_xctable n et)) = - (option_map (\ pc'. pc' + n) (match_exception_table G cn pc et))" + (Option.map (\ pc'. pc' + n) (match_exception_table G cn pc et))" apply (induct et) apply (simp add: offset_xctable_def)+ apply (case_tac "match_exception_entry G cn pc a") @@ -675,7 +675,7 @@ in app_jumps_lem) apply (simp add: nth_append)+ (* subgoal \ st. mt ! pc'' = Some st *) - apply (simp add: norm_eff_def option_map_def nth_append) + apply (simp add: norm_eff_def Option.map_def nth_append) apply (case_tac "mt ! pc''") apply simp+ done diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/MicroJava/Comp/LemmasComp.thy --- a/src/HOL/MicroJava/Comp/LemmasComp.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/MicroJava/Comp/LemmasComp.thy Wed Mar 04 11:05:29 2009 +0100 @@ -271,7 +271,7 @@ lemma map_of_map_fst: "\ inj f; \x\set xs. fst (f x) = fst x; \x\set xs. fst (g x) = fst x \ \ map_of (map g xs) k - = option_map (\ e. (snd (g ((inv f) (k, e))))) (map_of (map f xs) k)" + = Option.map (\ e. (snd (g ((inv f) (k, e))))) (map_of (map f xs) k)" apply (induct xs) apply simp apply (simp del: split_paired_All) @@ -288,13 +288,13 @@ lemma comp_method [rule_format (no_asm)]: "\ ws_prog G; is_class G C\ \ ((method (comp G, C) S) = - option_map (\ (D,rT,b). (D, rT, mtd_mb (compMethod G D (S, rT, b)))) + Option.map (\ (D,rT,b). (D, rT, mtd_mb (compMethod G D (S, rT, b)))) (method (G, C) S))" apply (simp add: method_def) apply (frule wf_subcls1) apply (simp add: comp_class_rec) apply (simp add: map_compose [THEN sym] split_iter split_compose del: map_compose) -apply (rule_tac R="%x y. ((x S) = (option_map (\(D, rT, b). +apply (rule_tac R="%x y. ((x S) = (Option.map (\(D, rT, b). (D, rT, snd (snd (compMethod G D (S, rT, b))))) (y S)))" in class_rec_relation) apply assumption diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/MicroJava/J/Conform.thy --- a/src/HOL/MicroJava/J/Conform.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/MicroJava/J/Conform.thy Wed Mar 04 11:05:29 2009 +0100 @@ -17,7 +17,7 @@ conf :: "'c prog => aheap => val => ty => bool" ("_,_ |- _ ::<= _" [51,51,51,51] 50) - "G,h|-v::<=T == \T'. typeof (option_map obj_ty o h) v = Some T' \ G\T'\T" + "G,h|-v::<=T == \T'. typeof (Option.map obj_ty o h) v = Some T' \ G\T'\T" lconf :: "'c prog => aheap => ('a \ val) => ('a \ ty) => bool" ("_,_ |- _ [::<=] _" [51,51,51,51] 50) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/MicroJava/J/Exceptions.thy --- a/src/HOL/MicroJava/J/Exceptions.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/MicroJava/J/Exceptions.thy Wed Mar 04 11:05:29 2009 +0100 @@ -21,7 +21,7 @@ cname_of :: "aheap \ val \ cname" translations - "cname_of hp v" == "fst (the (hp (the_Addr v)))" + "cname_of hp v" == "fst (CONST the (hp (the_Addr v)))" constdefs diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/MicroJava/J/State.thy --- a/src/HOL/MicroJava/J/State.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/MicroJava/J/State.thy Wed Mar 04 11:05:29 2009 +0100 @@ -41,7 +41,7 @@ "Norm s" == "(None,s)" "abrupt" => "fst" "store" => "snd" - "lookup_obj s a'" == "the (heap s (the_Addr a'))" + "lookup_obj s a'" == "CONST the (heap s (the_Addr a'))" constdefs diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/NanoJava/State.thy --- a/src/HOL/NanoJava/State.thy Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/NanoJava/State.thy Wed Mar 04 11:05:29 2009 +0100 @@ -33,7 +33,7 @@ constdefs init_vars:: "('a \ 'b) => ('a \ val)" - "init_vars m == option_map (\T. Null) o m" + "init_vars m == Option.map (\T. Null) o m" text {* private: *} types heap = "loc \ obj" diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Nat.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/NatBin.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Nominal/Nominal.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Nominal/nominal_atoms.ML --- a/src/HOL/Nominal/nominal_atoms.ML Wed Mar 04 11:05:02 2009 +0100 +++ b/src/HOL/Nominal/nominal_atoms.ML Wed Mar 04 11:05:29 2009 +0100 @@ -539,7 +539,7 @@ thy |> AxClass.prove_arity ("fun",[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_fun) |> AxClass.prove_arity ("Nominal.noption",[[cls_name]],[cls_name]) (pt_proof pt_thm_noptn) - |> AxClass.prove_arity ("Datatype.option",[[cls_name]],[cls_name]) (pt_proof pt_thm_optn) + |> AxClass.prove_arity ("Option.option",[[cls_name]],[cls_name]) (pt_proof pt_thm_optn) |> AxClass.prove_arity ("List.list",[[cls_name]],[cls_name]) (pt_proof pt_thm_list) |> AxClass.prove_arity ("*",[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_prod) |> AxClass.prove_arity ("Nominal.nprod",[[cls_name],[cls_name]],[cls_name]) @@ -606,7 +606,7 @@ |> AxClass.prove_arity ("Nominal.nprod",[[cls_name],[cls_name]],[cls_name]) (fs_proof fs_thm_nprod) |> AxClass.prove_arity ("List.list",[[cls_name]],[cls_name]) (fs_proof fs_thm_list) - |> AxClass.prove_arity ("Datatype.option",[[cls_name]],[cls_name]) (fs_proof fs_thm_optn) + |> AxClass.prove_arity ("Option.option",[[cls_name]],[cls_name]) (fs_proof fs_thm_optn) end) ak_names thy20; (******** cp__ class instances ********) @@ -687,7 +687,7 @@ |> AxClass.prove_arity ("*",[[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_prod) |> AxClass.prove_arity ("List.list",[[cls_name]],[cls_name]) (cp_proof cp_thm_list) |> AxClass.prove_arity ("fun",[[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_fun) - |> AxClass.prove_arity ("Datatype.option",[[cls_name]],[cls_name]) (cp_proof cp_thm_optn) + |> AxClass.prove_arity ("Option.option",[[cls_name]],[cls_name]) (cp_proof cp_thm_optn) |> AxClass.prove_arity ("Nominal.noption",[[cls_name]],[cls_name]) (cp_proof cp_thm_noptn) end) ak_names thy) ak_names thy25; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Nominal/nominal_inductive.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Nominal/nominal_inductive2.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/NumberTheory/Chinese.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/NumberTheory/IntPrimes.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Polynomial.thy --- a/src/HOL/Polynomial.thy Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1305 +0,0 @@ -(* Title: HOL/Polynomial.thy - Author: Brian Huffman - Based on an earlier development by Clemens Ballarin -*) - -header {* Univariate Polynomials *} - -theory Polynomial -imports Plain SetInterval Main -begin - -subsection {* Definition of type @{text poly} *} - -typedef (Poly) 'a poly = "{f::nat \ 'a::zero. \n. \i>n. f i = 0}" - morphisms coeff Abs_poly - by auto - -lemma expand_poly_eq: "p = q \ (\n. coeff p n = coeff q n)" -by (simp add: coeff_inject [symmetric] expand_fun_eq) - -lemma poly_ext: "(\n. coeff p n = coeff q n) \ p = q" -by (simp add: expand_poly_eq) - - -subsection {* Degree of a polynomial *} - -definition - degree :: "'a::zero poly \ nat" where - "degree p = (LEAST n. \i>n. coeff p i = 0)" - -lemma coeff_eq_0: "degree p < n \ coeff p n = 0" -proof - - have "coeff p \ Poly" - by (rule coeff) - hence "\n. \i>n. coeff p i = 0" - unfolding Poly_def by simp - hence "\i>degree p. coeff p i = 0" - unfolding degree_def by (rule LeastI_ex) - moreover assume "degree p < n" - ultimately show ?thesis by simp -qed - -lemma le_degree: "coeff p n \ 0 \ n \ degree p" - by (erule contrapos_np, rule coeff_eq_0, simp) - -lemma degree_le: "\i>n. coeff p i = 0 \ degree p \ n" - unfolding degree_def by (erule Least_le) - -lemma less_degree_imp: "n < degree p \ \i>n. coeff p i \ 0" - unfolding degree_def by (drule not_less_Least, simp) - - -subsection {* The zero polynomial *} - -instantiation poly :: (zero) zero -begin - -definition - zero_poly_def: "0 = Abs_poly (\n. 0)" - -instance .. -end - -lemma coeff_0 [simp]: "coeff 0 n = 0" - unfolding zero_poly_def - by (simp add: Abs_poly_inverse Poly_def) - -lemma degree_0 [simp]: "degree 0 = 0" - by (rule order_antisym [OF degree_le le0]) simp - -lemma leading_coeff_neq_0: - assumes "p \ 0" shows "coeff p (degree p) \ 0" -proof (cases "degree p") - case 0 - from `p \ 0` have "\n. coeff p n \ 0" - by (simp add: expand_poly_eq) - then obtain n where "coeff p n \ 0" .. - hence "n \ degree p" by (rule le_degree) - with `coeff p n \ 0` and `degree p = 0` - show "coeff p (degree p) \ 0" by simp -next - case (Suc n) - from `degree p = Suc n` have "n < degree p" by simp - hence "\i>n. coeff p i \ 0" by (rule less_degree_imp) - then obtain i where "n < i" and "coeff p i \ 0" by fast - from `degree p = Suc n` and `n < i` have "degree p \ i" by simp - also from `coeff p i \ 0` have "i \ degree p" by (rule le_degree) - finally have "degree p = i" . - with `coeff p i \ 0` show "coeff p (degree p) \ 0" by simp -qed - -lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \ p = 0" - by (cases "p = 0", simp, simp add: leading_coeff_neq_0) - - -subsection {* List-style constructor for polynomials *} - -definition - pCons :: "'a::zero \ 'a poly \ 'a poly" -where - [code del]: "pCons a p = Abs_poly (nat_case a (coeff p))" - -syntax - "_poly" :: "args \ 'a poly" ("[:(_):]") - -translations - "[:x, xs:]" == "CONST pCons x [:xs:]" - "[:x:]" == "CONST pCons x 0" - -lemma Poly_nat_case: "f \ Poly \ nat_case a f \ Poly" - unfolding Poly_def by (auto split: nat.split) - -lemma coeff_pCons: - "coeff (pCons a p) = nat_case a (coeff p)" - unfolding pCons_def - by (simp add: Abs_poly_inverse Poly_nat_case coeff) - -lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a" - by (simp add: coeff_pCons) - -lemma coeff_pCons_Suc [simp]: "coeff (pCons a p) (Suc n) = coeff p n" - by (simp add: coeff_pCons) - -lemma degree_pCons_le: "degree (pCons a p) \ Suc (degree p)" -by (rule degree_le, simp add: coeff_eq_0 coeff_pCons split: nat.split) - -lemma degree_pCons_eq: - "p \ 0 \ degree (pCons a p) = Suc (degree p)" -apply (rule order_antisym [OF degree_pCons_le]) -apply (rule le_degree, simp) -done - -lemma degree_pCons_0: "degree (pCons a 0) = 0" -apply (rule order_antisym [OF _ le0]) -apply (rule degree_le, simp add: coeff_pCons split: nat.split) -done - -lemma degree_pCons_eq_if [simp]: - "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" -apply (cases "p = 0", simp_all) -apply (rule order_antisym [OF _ le0]) -apply (rule degree_le, simp add: coeff_pCons split: nat.split) -apply (rule order_antisym [OF degree_pCons_le]) -apply (rule le_degree, simp) -done - -lemma pCons_0_0 [simp]: "pCons 0 0 = 0" -by (rule poly_ext, simp add: coeff_pCons split: nat.split) - -lemma pCons_eq_iff [simp]: - "pCons a p = pCons b q \ a = b \ p = q" -proof (safe) - assume "pCons a p = pCons b q" - then have "coeff (pCons a p) 0 = coeff (pCons b q) 0" by simp - then show "a = b" by simp -next - assume "pCons a p = pCons b q" - then have "\n. coeff (pCons a p) (Suc n) = - coeff (pCons b q) (Suc n)" by simp - then show "p = q" by (simp add: expand_poly_eq) -qed - -lemma pCons_eq_0_iff [simp]: "pCons a p = 0 \ a = 0 \ p = 0" - using pCons_eq_iff [of a p 0 0] by simp - -lemma Poly_Suc: "f \ Poly \ (\n. f (Suc n)) \ Poly" - unfolding Poly_def - by (clarify, rule_tac x=n in exI, simp) - -lemma pCons_cases [cases type: poly]: - obtains (pCons) a q where "p = pCons a q" -proof - show "p = pCons (coeff p 0) (Abs_poly (\n. coeff p (Suc n)))" - by (rule poly_ext) - (simp add: Abs_poly_inverse Poly_Suc coeff coeff_pCons - split: nat.split) -qed - -lemma pCons_induct [case_names 0 pCons, induct type: poly]: - assumes zero: "P 0" - assumes pCons: "\a p. P p \ P (pCons a p)" - shows "P p" -proof (induct p rule: measure_induct_rule [where f=degree]) - case (less p) - obtain a q where "p = pCons a q" by (rule pCons_cases) - have "P q" - proof (cases "q = 0") - case True - then show "P q" by (simp add: zero) - next - case False - then have "degree (pCons a q) = Suc (degree q)" - by (rule degree_pCons_eq) - then have "degree q < degree p" - using `p = pCons a q` by simp - then show "P q" - by (rule less.hyps) - qed - then have "P (pCons a q)" - by (rule pCons) - then show ?case - using `p = pCons a q` by simp -qed - - -subsection {* Recursion combinator for polynomials *} - -function - poly_rec :: "'b \ ('a::zero \ 'a poly \ 'b \ 'b) \ 'a poly \ 'b" -where - poly_rec_pCons_eq_if [simp del, code del]: - "poly_rec z f (pCons a p) = f a p (if p = 0 then z else poly_rec z f p)" -by (case_tac x, rename_tac q, case_tac q, auto) - -termination poly_rec -by (relation "measure (degree \ snd \ snd)", simp) - (simp add: degree_pCons_eq) - -lemma poly_rec_0: - "f 0 0 z = z \ poly_rec z f 0 = z" - using poly_rec_pCons_eq_if [of z f 0 0] by simp - -lemma poly_rec_pCons: - "f 0 0 z = z \ poly_rec z f (pCons a p) = f a p (poly_rec z f p)" - by (simp add: poly_rec_pCons_eq_if poly_rec_0) - - -subsection {* Monomials *} - -definition - monom :: "'a \ nat \ 'a::zero poly" where - "monom a m = Abs_poly (\n. if m = n then a else 0)" - -lemma coeff_monom [simp]: "coeff (monom a m) n = (if m=n then a else 0)" - unfolding monom_def - by (subst Abs_poly_inverse, auto simp add: Poly_def) - -lemma monom_0: "monom a 0 = pCons a 0" - by (rule poly_ext, simp add: coeff_pCons split: nat.split) - -lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)" - by (rule poly_ext, simp add: coeff_pCons split: nat.split) - -lemma monom_eq_0 [simp]: "monom 0 n = 0" - by (rule poly_ext) simp - -lemma monom_eq_0_iff [simp]: "monom a n = 0 \ a = 0" - by (simp add: expand_poly_eq) - -lemma monom_eq_iff [simp]: "monom a n = monom b n \ a = b" - by (simp add: expand_poly_eq) - -lemma degree_monom_le: "degree (monom a n) \ n" - by (rule degree_le, simp) - -lemma degree_monom_eq: "a \ 0 \ degree (monom a n) = n" - apply (rule order_antisym [OF degree_monom_le]) - apply (rule le_degree, simp) - done - - -subsection {* Addition and subtraction *} - -instantiation poly :: (comm_monoid_add) comm_monoid_add -begin - -definition - plus_poly_def [code del]: - "p + q = Abs_poly (\n. coeff p n + coeff q n)" - -lemma Poly_add: - fixes f g :: "nat \ 'a" - shows "\f \ Poly; g \ Poly\ \ (\n. f n + g n) \ Poly" - unfolding Poly_def - apply (clarify, rename_tac m n) - apply (rule_tac x="max m n" in exI, simp) - done - -lemma coeff_add [simp]: - "coeff (p + q) n = coeff p n + coeff q n" - unfolding plus_poly_def - by (simp add: Abs_poly_inverse coeff Poly_add) - -instance proof - fix p q r :: "'a poly" - show "(p + q) + r = p + (q + r)" - by (simp add: expand_poly_eq add_assoc) - show "p + q = q + p" - by (simp add: expand_poly_eq add_commute) - show "0 + p = p" - by (simp add: expand_poly_eq) -qed - -end - -instance poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add -proof - fix p q r :: "'a poly" - assume "p + q = p + r" thus "q = r" - by (simp add: expand_poly_eq) -qed - -instantiation poly :: (ab_group_add) ab_group_add -begin - -definition - uminus_poly_def [code del]: - "- p = Abs_poly (\n. - coeff p n)" - -definition - minus_poly_def [code del]: - "p - q = Abs_poly (\n. coeff p n - coeff q n)" - -lemma Poly_minus: - fixes f :: "nat \ 'a" - shows "f \ Poly \ (\n. - f n) \ Poly" - unfolding Poly_def by simp - -lemma Poly_diff: - fixes f g :: "nat \ 'a" - shows "\f \ Poly; g \ Poly\ \ (\n. f n - g n) \ Poly" - unfolding diff_minus by (simp add: Poly_add Poly_minus) - -lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n" - unfolding uminus_poly_def - by (simp add: Abs_poly_inverse coeff Poly_minus) - -lemma coeff_diff [simp]: - "coeff (p - q) n = coeff p n - coeff q n" - unfolding minus_poly_def - by (simp add: Abs_poly_inverse coeff Poly_diff) - -instance proof - fix p q :: "'a poly" - show "- p + p = 0" - by (simp add: expand_poly_eq) - show "p - q = p + - q" - by (simp add: expand_poly_eq diff_minus) -qed - -end - -lemma add_pCons [simp]: - "pCons a p + pCons b q = pCons (a + b) (p + q)" - by (rule poly_ext, simp add: coeff_pCons split: nat.split) - -lemma minus_pCons [simp]: - "- pCons a p = pCons (- a) (- p)" - by (rule poly_ext, simp add: coeff_pCons split: nat.split) - -lemma diff_pCons [simp]: - "pCons a p - pCons b q = pCons (a - b) (p - q)" - by (rule poly_ext, simp add: coeff_pCons split: nat.split) - -lemma degree_add_le_max: "degree (p + q) \ max (degree p) (degree q)" - by (rule degree_le, auto simp add: coeff_eq_0) - -lemma degree_add_le: - "\degree p \ n; degree q \ n\ \ degree (p + q) \ n" - by (auto intro: order_trans degree_add_le_max) - -lemma degree_add_less: - "\degree p < n; degree q < n\ \ degree (p + q) < n" - by (auto intro: le_less_trans degree_add_le_max) - -lemma degree_add_eq_right: - "degree p < degree q \ degree (p + q) = degree q" - apply (cases "q = 0", simp) - apply (rule order_antisym) - apply (simp add: degree_add_le) - apply (rule le_degree) - apply (simp add: coeff_eq_0) - done - -lemma degree_add_eq_left: - "degree q < degree p \ degree (p + q) = degree p" - using degree_add_eq_right [of q p] - by (simp add: add_commute) - -lemma degree_minus [simp]: "degree (- p) = degree p" - unfolding degree_def by simp - -lemma degree_diff_le_max: "degree (p - q) \ max (degree p) (degree q)" - using degree_add_le [where p=p and q="-q"] - by (simp add: diff_minus) - -lemma degree_diff_le: - "\degree p \ n; degree q \ n\ \ degree (p - q) \ n" - by (simp add: diff_minus degree_add_le) - -lemma degree_diff_less: - "\degree p < n; degree q < n\ \ degree (p - q) < n" - by (simp add: diff_minus degree_add_less) - -lemma add_monom: "monom a n + monom b n = monom (a + b) n" - by (rule poly_ext) simp - -lemma diff_monom: "monom a n - monom b n = monom (a - b) n" - by (rule poly_ext) simp - -lemma minus_monom: "- monom a n = monom (-a) n" - by (rule poly_ext) simp - -lemma coeff_setsum: "coeff (\x\A. p x) i = (\x\A. coeff (p x) i)" - by (cases "finite A", induct set: finite, simp_all) - -lemma monom_setsum: "monom (\x\A. a x) n = (\x\A. monom (a x) n)" - by (rule poly_ext) (simp add: coeff_setsum) - - -subsection {* Multiplication by a constant *} - -definition - smult :: "'a::comm_semiring_0 \ 'a poly \ 'a poly" where - "smult a p = Abs_poly (\n. a * coeff p n)" - -lemma Poly_smult: - fixes f :: "nat \ 'a::comm_semiring_0" - shows "f \ Poly \ (\n. a * f n) \ Poly" - unfolding Poly_def - by (clarify, rule_tac x=n in exI, simp) - -lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n" - unfolding smult_def - by (simp add: Abs_poly_inverse Poly_smult coeff) - -lemma degree_smult_le: "degree (smult a p) \ degree p" - by (rule degree_le, simp add: coeff_eq_0) - -lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p" - by (rule poly_ext, simp add: mult_assoc) - -lemma smult_0_right [simp]: "smult a 0 = 0" - by (rule poly_ext, simp) - -lemma smult_0_left [simp]: "smult 0 p = 0" - by (rule poly_ext, simp) - -lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p" - by (rule poly_ext, simp) - -lemma smult_add_right: - "smult a (p + q) = smult a p + smult a q" - by (rule poly_ext, simp add: algebra_simps) - -lemma smult_add_left: - "smult (a + b) p = smult a p + smult b p" - by (rule poly_ext, simp add: algebra_simps) - -lemma smult_minus_right [simp]: - "smult (a::'a::comm_ring) (- p) = - smult a p" - by (rule poly_ext, simp) - -lemma smult_minus_left [simp]: - "smult (- a::'a::comm_ring) p = - smult a p" - by (rule poly_ext, simp) - -lemma smult_diff_right: - "smult (a::'a::comm_ring) (p - q) = smult a p - smult a q" - by (rule poly_ext, simp add: algebra_simps) - -lemma smult_diff_left: - "smult (a - b::'a::comm_ring) p = smult a p - smult b p" - by (rule poly_ext, simp add: algebra_simps) - -lemmas smult_distribs = - smult_add_left smult_add_right - smult_diff_left smult_diff_right - -lemma smult_pCons [simp]: - "smult a (pCons b p) = pCons (a * b) (smult a p)" - by (rule poly_ext, simp add: coeff_pCons split: nat.split) - -lemma smult_monom: "smult a (monom b n) = monom (a * b) n" - by (induct n, simp add: monom_0, simp add: monom_Suc) - -lemma degree_smult_eq [simp]: - fixes a :: "'a::idom" - shows "degree (smult a p) = (if a = 0 then 0 else degree p)" - by (cases "a = 0", simp, simp add: degree_def) - -lemma smult_eq_0_iff [simp]: - fixes a :: "'a::idom" - shows "smult a p = 0 \ a = 0 \ p = 0" - by (simp add: expand_poly_eq) - - -subsection {* Multiplication of polynomials *} - -text {* TODO: move to SetInterval.thy *} -lemma setsum_atMost_Suc_shift: - fixes f :: "nat \ 'a::comm_monoid_add" - shows "(\i\Suc n. f i) = f 0 + (\i\n. f (Suc i))" -proof (induct n) - case 0 show ?case by simp -next - case (Suc n) note IH = this - have "(\i\Suc (Suc n). f i) = (\i\Suc n. f i) + f (Suc (Suc n))" - by (rule setsum_atMost_Suc) - also have "(\i\Suc n. f i) = f 0 + (\i\n. f (Suc i))" - by (rule IH) - also have "f 0 + (\i\n. f (Suc i)) + f (Suc (Suc n)) = - f 0 + ((\i\n. f (Suc i)) + f (Suc (Suc n)))" - by (rule add_assoc) - also have "(\i\n. f (Suc i)) + f (Suc (Suc n)) = (\i\Suc n. f (Suc i))" - by (rule setsum_atMost_Suc [symmetric]) - finally show ?case . -qed - -instantiation poly :: (comm_semiring_0) comm_semiring_0 -begin - -definition - times_poly_def [code del]: - "p * q = poly_rec 0 (\a p pq. smult a q + pCons 0 pq) p" - -lemma mult_poly_0_left: "(0::'a poly) * q = 0" - unfolding times_poly_def by (simp add: poly_rec_0) - -lemma mult_pCons_left [simp]: - "pCons a p * q = smult a q + pCons 0 (p * q)" - unfolding times_poly_def by (simp add: poly_rec_pCons) - -lemma mult_poly_0_right: "p * (0::'a poly) = 0" - by (induct p, simp add: mult_poly_0_left, simp) - -lemma mult_pCons_right [simp]: - "p * pCons a q = smult a p + pCons 0 (p * q)" - by (induct p, simp add: mult_poly_0_left, simp add: algebra_simps) - -lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right - -lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)" - by (induct p, simp add: mult_poly_0, simp add: smult_add_right) - -lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)" - by (induct q, simp add: mult_poly_0, simp add: smult_add_right) - -lemma mult_poly_add_left: - fixes p q r :: "'a poly" - shows "(p + q) * r = p * r + q * r" - by (induct r, simp add: mult_poly_0, - simp add: smult_distribs algebra_simps) - -instance proof - fix p q r :: "'a poly" - show 0: "0 * p = 0" - by (rule mult_poly_0_left) - show "p * 0 = 0" - by (rule mult_poly_0_right) - show "(p + q) * r = p * r + q * r" - by (rule mult_poly_add_left) - show "(p * q) * r = p * (q * r)" - by (induct p, simp add: mult_poly_0, simp add: mult_poly_add_left) - show "p * q = q * p" - by (induct p, simp add: mult_poly_0, simp) -qed - -end - -instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. - -lemma coeff_mult: - "coeff (p * q) n = (\i\n. coeff p i * coeff q (n-i))" -proof (induct p arbitrary: n) - case 0 show ?case by simp -next - case (pCons a p n) thus ?case - by (cases n, simp, simp add: setsum_atMost_Suc_shift - del: setsum_atMost_Suc) -qed - -lemma degree_mult_le: "degree (p * q) \ degree p + degree q" -apply (rule degree_le) -apply (induct p) -apply simp -apply (simp add: coeff_eq_0 coeff_pCons split: nat.split) -done - -lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)" - by (induct m, simp add: monom_0 smult_monom, simp add: monom_Suc) - - -subsection {* The unit polynomial and exponentiation *} - -instantiation poly :: (comm_semiring_1) comm_semiring_1 -begin - -definition - one_poly_def: - "1 = pCons 1 0" - -instance proof - fix p :: "'a poly" show "1 * p = p" - unfolding one_poly_def - by simp -next - show "0 \ (1::'a poly)" - unfolding one_poly_def by simp -qed - -end - -instance poly :: (comm_semiring_1_cancel) comm_semiring_1_cancel .. - -lemma coeff_1 [simp]: "coeff 1 n = (if n = 0 then 1 else 0)" - unfolding one_poly_def - by (simp add: coeff_pCons split: nat.split) - -lemma degree_1 [simp]: "degree 1 = 0" - unfolding one_poly_def - by (rule degree_pCons_0) - -instantiation poly :: (comm_semiring_1) recpower -begin - -primrec power_poly where - power_poly_0: "(p::'a poly) ^ 0 = 1" -| power_poly_Suc: "(p::'a poly) ^ (Suc n) = p * p ^ n" - -instance - by default simp_all - -end - -instance poly :: (comm_ring) comm_ring .. - -instance poly :: (comm_ring_1) comm_ring_1 .. - -instantiation poly :: (comm_ring_1) number_ring -begin - -definition - "number_of k = (of_int k :: 'a poly)" - -instance - by default (rule number_of_poly_def) - -end - - -subsection {* Polynomials form an integral domain *} - -lemma coeff_mult_degree_sum: - "coeff (p * q) (degree p + degree q) = - coeff p (degree p) * coeff q (degree q)" - by (induct p, simp, simp add: coeff_eq_0) - -instance poly :: (idom) idom -proof - fix p q :: "'a poly" - assume "p \ 0" and "q \ 0" - have "coeff (p * q) (degree p + degree q) = - coeff p (degree p) * coeff q (degree q)" - by (rule coeff_mult_degree_sum) - also have "coeff p (degree p) * coeff q (degree q) \ 0" - using `p \ 0` and `q \ 0` by simp - finally have "\n. coeff (p * q) n \ 0" .. - thus "p * q \ 0" by (simp add: expand_poly_eq) -qed - -lemma degree_mult_eq: - fixes p q :: "'a::idom poly" - shows "\p \ 0; q \ 0\ \ degree (p * q) = degree p + degree q" -apply (rule order_antisym [OF degree_mult_le le_degree]) -apply (simp add: coeff_mult_degree_sum) -done - -lemma dvd_imp_degree_le: - fixes p q :: "'a::idom poly" - shows "\p dvd q; q \ 0\ \ degree p \ degree q" - by (erule dvdE, simp add: degree_mult_eq) - - -subsection {* Polynomials form an ordered integral domain *} - -definition - pos_poly :: "'a::ordered_idom poly \ bool" -where - "pos_poly p \ 0 < coeff p (degree p)" - -lemma pos_poly_pCons: - "pos_poly (pCons a p) \ pos_poly p \ (p = 0 \ 0 < a)" - unfolding pos_poly_def by simp - -lemma not_pos_poly_0 [simp]: "\ pos_poly 0" - unfolding pos_poly_def by simp - -lemma pos_poly_add: "\pos_poly p; pos_poly q\ \ pos_poly (p + q)" - apply (induct p arbitrary: q, simp) - apply (case_tac q, force simp add: pos_poly_pCons add_pos_pos) - done - -lemma pos_poly_mult: "\pos_poly p; pos_poly q\ \ pos_poly (p * q)" - unfolding pos_poly_def - apply (subgoal_tac "p \ 0 \ q \ 0") - apply (simp add: degree_mult_eq coeff_mult_degree_sum mult_pos_pos) - apply auto - done - -lemma pos_poly_total: "p = 0 \ pos_poly p \ pos_poly (- p)" -by (induct p) (auto simp add: pos_poly_pCons) - -instantiation poly :: (ordered_idom) ordered_idom -begin - -definition - [code del]: - "x < y \ pos_poly (y - x)" - -definition - [code del]: - "x \ y \ x = y \ pos_poly (y - x)" - -definition - [code del]: - "abs (x::'a poly) = (if x < 0 then - x else x)" - -definition - [code del]: - "sgn (x::'a poly) = (if x = 0 then 0 else if 0 < x then 1 else - 1)" - -instance proof - fix x y :: "'a poly" - show "x < y \ x \ y \ \ y \ x" - unfolding less_eq_poly_def less_poly_def - apply safe - apply simp - apply (drule (1) pos_poly_add) - apply simp - done -next - fix x :: "'a poly" show "x \ x" - unfolding less_eq_poly_def by simp -next - fix x y z :: "'a poly" - assume "x \ y" and "y \ z" thus "x \ z" - unfolding less_eq_poly_def - apply safe - apply (drule (1) pos_poly_add) - apply (simp add: algebra_simps) - done -next - fix x y :: "'a poly" - assume "x \ y" and "y \ x" thus "x = y" - unfolding less_eq_poly_def - apply safe - apply (drule (1) pos_poly_add) - apply simp - done -next - fix x y z :: "'a poly" - assume "x \ y" thus "z + x \ z + y" - unfolding less_eq_poly_def - apply safe - apply (simp add: algebra_simps) - done -next - fix x y :: "'a poly" - show "x \ y \ y \ x" - unfolding less_eq_poly_def - using pos_poly_total [of "x - y"] - by auto -next - fix x y z :: "'a poly" - assume "x < y" and "0 < z" - thus "z * x < z * y" - unfolding less_poly_def - by (simp add: right_diff_distrib [symmetric] pos_poly_mult) -next - fix x :: "'a poly" - show "\x\ = (if x < 0 then - x else x)" - by (rule abs_poly_def) -next - fix x :: "'a poly" - show "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" - by (rule sgn_poly_def) -qed - -end - -text {* TODO: Simplification rules for comparisons *} - - -subsection {* Long division of polynomials *} - -definition - pdivmod_rel :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly \ bool" -where - [code del]: - "pdivmod_rel x y q r \ - x = q * y + r \ (if y = 0 then q = 0 else r = 0 \ degree r < degree y)" - -lemma pdivmod_rel_0: - "pdivmod_rel 0 y 0 0" - unfolding pdivmod_rel_def by simp - -lemma pdivmod_rel_by_0: - "pdivmod_rel x 0 0 x" - unfolding pdivmod_rel_def by simp - -lemma eq_zero_or_degree_less: - assumes "degree p \ n" and "coeff p n = 0" - shows "p = 0 \ degree p < n" -proof (cases n) - case 0 - with `degree p \ n` and `coeff p n = 0` - have "coeff p (degree p) = 0" by simp - then have "p = 0" by simp - then show ?thesis .. -next - case (Suc m) - have "\i>n. coeff p i = 0" - using `degree p \ n` by (simp add: coeff_eq_0) - then have "\i\n. coeff p i = 0" - using `coeff p n = 0` by (simp add: le_less) - then have "\i>m. coeff p i = 0" - using `n = Suc m` by (simp add: less_eq_Suc_le) - then have "degree p \ m" - by (rule degree_le) - then have "degree p < n" - using `n = Suc m` by (simp add: less_Suc_eq_le) - then show ?thesis .. -qed - -lemma pdivmod_rel_pCons: - assumes rel: "pdivmod_rel x y q r" - assumes y: "y \ 0" - assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" - shows "pdivmod_rel (pCons a x) y (pCons b q) (pCons a r - smult b y)" - (is "pdivmod_rel ?x y ?q ?r") -proof - - have x: "x = q * y + r" and r: "r = 0 \ degree r < degree y" - using assms unfolding pdivmod_rel_def by simp_all - - have 1: "?x = ?q * y + ?r" - using b x by simp - - have 2: "?r = 0 \ degree ?r < degree y" - proof (rule eq_zero_or_degree_less) - show "degree ?r \ degree y" - proof (rule degree_diff_le) - show "degree (pCons a r) \ degree y" - using r by auto - show "degree (smult b y) \ degree y" - by (rule degree_smult_le) - qed - next - show "coeff ?r (degree y) = 0" - using `y \ 0` unfolding b by simp - qed - - from 1 2 show ?thesis - unfolding pdivmod_rel_def - using `y \ 0` by simp -qed - -lemma pdivmod_rel_exists: "\q r. pdivmod_rel x y q r" -apply (cases "y = 0") -apply (fast intro!: pdivmod_rel_by_0) -apply (induct x) -apply (fast intro!: pdivmod_rel_0) -apply (fast intro!: pdivmod_rel_pCons) -done - -lemma pdivmod_rel_unique: - assumes 1: "pdivmod_rel x y q1 r1" - assumes 2: "pdivmod_rel x y q2 r2" - shows "q1 = q2 \ r1 = r2" -proof (cases "y = 0") - assume "y = 0" with assms show ?thesis - by (simp add: pdivmod_rel_def) -next - assume [simp]: "y \ 0" - from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \ degree r1 < degree y" - unfolding pdivmod_rel_def by simp_all - from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \ degree r2 < degree y" - unfolding pdivmod_rel_def by simp_all - from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" - by (simp add: algebra_simps) - from r1 r2 have r3: "(r2 - r1) = 0 \ degree (r2 - r1) < degree y" - by (auto intro: degree_diff_less) - - show "q1 = q2 \ r1 = r2" - proof (rule ccontr) - assume "\ (q1 = q2 \ r1 = r2)" - with q3 have "q1 \ q2" and "r1 \ r2" by auto - with r3 have "degree (r2 - r1) < degree y" by simp - also have "degree y \ degree (q1 - q2) + degree y" by simp - also have "\ = degree ((q1 - q2) * y)" - using `q1 \ q2` by (simp add: degree_mult_eq) - also have "\ = degree (r2 - r1)" - using q3 by simp - finally have "degree (r2 - r1) < degree (r2 - r1)" . - then show "False" by simp - qed -qed - -lemma pdivmod_rel_0_iff: "pdivmod_rel 0 y q r \ q = 0 \ r = 0" -by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_0) - -lemma pdivmod_rel_by_0_iff: "pdivmod_rel x 0 q r \ q = 0 \ r = x" -by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_by_0) - -lemmas pdivmod_rel_unique_div = - pdivmod_rel_unique [THEN conjunct1, standard] - -lemmas pdivmod_rel_unique_mod = - pdivmod_rel_unique [THEN conjunct2, standard] - -instantiation poly :: (field) ring_div -begin - -definition div_poly where - [code del]: "x div y = (THE q. \r. pdivmod_rel x y q r)" - -definition mod_poly where - [code del]: "x mod y = (THE r. \q. pdivmod_rel x y q r)" - -lemma div_poly_eq: - "pdivmod_rel x y q r \ x div y = q" -unfolding div_poly_def -by (fast elim: pdivmod_rel_unique_div) - -lemma mod_poly_eq: - "pdivmod_rel x y q r \ x mod y = r" -unfolding mod_poly_def -by (fast elim: pdivmod_rel_unique_mod) - -lemma pdivmod_rel: - "pdivmod_rel x y (x div y) (x mod y)" -proof - - from pdivmod_rel_exists - obtain q r where "pdivmod_rel x y q r" by fast - thus ?thesis - by (simp add: div_poly_eq mod_poly_eq) -qed - -instance proof - fix x y :: "'a poly" - show "x div y * y + x mod y = x" - using pdivmod_rel [of x y] - by (simp add: pdivmod_rel_def) -next - fix x :: "'a poly" - have "pdivmod_rel x 0 0 x" - by (rule pdivmod_rel_by_0) - thus "x div 0 = 0" - by (rule div_poly_eq) -next - fix y :: "'a poly" - have "pdivmod_rel 0 y 0 0" - by (rule pdivmod_rel_0) - thus "0 div y = 0" - by (rule div_poly_eq) -next - fix x y z :: "'a poly" - assume "y \ 0" - hence "pdivmod_rel (x + z * y) y (z + x div y) (x mod y)" - using pdivmod_rel [of x y] - by (simp add: pdivmod_rel_def left_distrib) - thus "(x + z * y) div y = z + x div y" - by (rule div_poly_eq) -qed - -end - -lemma degree_mod_less: - "y \ 0 \ x mod y = 0 \ degree (x mod y) < degree y" - using pdivmod_rel [of x y] - unfolding pdivmod_rel_def by simp - -lemma div_poly_less: "degree x < degree y \ x div y = 0" -proof - - assume "degree x < degree y" - hence "pdivmod_rel x y 0 x" - by (simp add: pdivmod_rel_def) - thus "x div y = 0" by (rule div_poly_eq) -qed - -lemma mod_poly_less: "degree x < degree y \ x mod y = x" -proof - - assume "degree x < degree y" - hence "pdivmod_rel x y 0 x" - by (simp add: pdivmod_rel_def) - thus "x mod y = x" by (rule mod_poly_eq) -qed - -lemma pdivmod_rel_smult_left: - "pdivmod_rel x y q r - \ pdivmod_rel (smult a x) y (smult a q) (smult a r)" - unfolding pdivmod_rel_def by (simp add: smult_add_right) - -lemma div_smult_left: "(smult a x) div y = smult a (x div y)" - by (rule div_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel) - -lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)" - by (rule mod_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel) - -lemma pdivmod_rel_smult_right: - "\a \ 0; pdivmod_rel x y q r\ - \ pdivmod_rel x (smult a y) (smult (inverse a) q) r" - unfolding pdivmod_rel_def by simp - -lemma div_smult_right: - "a \ 0 \ x div (smult a y) = smult (inverse a) (x div y)" - by (rule div_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel) - -lemma mod_smult_right: "a \ 0 \ x mod (smult a y) = x mod y" - by (rule mod_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel) - -lemma pdivmod_rel_mult: - "\pdivmod_rel x y q r; pdivmod_rel q z q' r'\ - \ pdivmod_rel x (y * z) q' (y * r' + r)" -apply (cases "z = 0", simp add: pdivmod_rel_def) -apply (cases "y = 0", simp add: pdivmod_rel_by_0_iff pdivmod_rel_0_iff) -apply (cases "r = 0") -apply (cases "r' = 0") -apply (simp add: pdivmod_rel_def) -apply (simp add: pdivmod_rel_def ring_simps degree_mult_eq) -apply (cases "r' = 0") -apply (simp add: pdivmod_rel_def degree_mult_eq) -apply (simp add: pdivmod_rel_def ring_simps) -apply (simp add: degree_mult_eq degree_add_less) -done - -lemma poly_div_mult_right: - fixes x y z :: "'a::field poly" - shows "x div (y * z) = (x div y) div z" - by (rule div_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+) - -lemma poly_mod_mult_right: - fixes x y z :: "'a::field poly" - shows "x mod (y * z) = y * (x div y mod z) + x mod y" - by (rule mod_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+) - -lemma mod_pCons: - fixes a and x - assumes y: "y \ 0" - defines b: "b \ coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" - shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)" -unfolding b -apply (rule mod_poly_eq) -apply (rule pdivmod_rel_pCons [OF pdivmod_rel y refl]) -done - - -subsection {* Evaluation of polynomials *} - -definition - poly :: "'a::comm_semiring_0 poly \ 'a \ 'a" where - "poly = poly_rec (\x. 0) (\a p f x. a + x * f x)" - -lemma poly_0 [simp]: "poly 0 x = 0" - unfolding poly_def by (simp add: poly_rec_0) - -lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x" - unfolding poly_def by (simp add: poly_rec_pCons) - -lemma poly_1 [simp]: "poly 1 x = 1" - unfolding one_poly_def by simp - -lemma poly_monom: - fixes a x :: "'a::{comm_semiring_1,recpower}" - shows "poly (monom a n) x = a * x ^ n" - by (induct n, simp add: monom_0, simp add: monom_Suc power_Suc mult_ac) - -lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" - apply (induct p arbitrary: q, simp) - apply (case_tac q, simp, simp add: algebra_simps) - done - -lemma poly_minus [simp]: - fixes x :: "'a::comm_ring" - shows "poly (- p) x = - poly p x" - by (induct p, simp_all) - -lemma poly_diff [simp]: - fixes x :: "'a::comm_ring" - shows "poly (p - q) x = poly p x - poly q x" - by (simp add: diff_minus) - -lemma poly_setsum: "poly (\k\A. p k) x = (\k\A. poly (p k) x)" - by (cases "finite A", induct set: finite, simp_all) - -lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x" - by (induct p, simp, simp add: algebra_simps) - -lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x" - by (induct p, simp_all, simp add: algebra_simps) - -lemma poly_power [simp]: - fixes p :: "'a::{comm_semiring_1,recpower} poly" - shows "poly (p ^ n) x = poly p x ^ n" - by (induct n, simp, simp add: power_Suc) - - -subsection {* Synthetic division *} - -definition - synthetic_divmod :: "'a::comm_semiring_0 poly \ 'a \ 'a poly \ 'a" -where [code del]: - "synthetic_divmod p c = - poly_rec (0, 0) (\a p (q, r). (pCons r q, a + c * r)) p" - -definition - synthetic_div :: "'a::comm_semiring_0 poly \ 'a \ 'a poly" -where - "synthetic_div p c = fst (synthetic_divmod p c)" - -lemma synthetic_divmod_0 [simp]: - "synthetic_divmod 0 c = (0, 0)" - unfolding synthetic_divmod_def - by (simp add: poly_rec_0) - -lemma synthetic_divmod_pCons [simp]: - "synthetic_divmod (pCons a p) c = - (\(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" - unfolding synthetic_divmod_def - by (simp add: poly_rec_pCons) - -lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c" - by (induct p, simp, simp add: split_def) - -lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0" - unfolding synthetic_div_def by simp - -lemma synthetic_div_pCons [simp]: - "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" - unfolding synthetic_div_def - by (simp add: split_def snd_synthetic_divmod) - -lemma synthetic_div_eq_0_iff: - "synthetic_div p c = 0 \ degree p = 0" - by (induct p, simp, case_tac p, simp) - -lemma degree_synthetic_div: - "degree (synthetic_div p c) = degree p - 1" - by (induct p, simp, simp add: synthetic_div_eq_0_iff) - -lemma synthetic_div_correct: - "p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)" - by (induct p) simp_all - -lemma synthetic_div_unique_lemma: "smult c p = pCons a p \ p = 0" -by (induct p arbitrary: a) simp_all - -lemma synthetic_div_unique: - "p + smult c q = pCons r q \ r = poly p c \ q = synthetic_div p c" -apply (induct p arbitrary: q r) -apply (simp, frule synthetic_div_unique_lemma, simp) -apply (case_tac q, force) -done - -lemma synthetic_div_correct': - fixes c :: "'a::comm_ring_1" - shows "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" - using synthetic_div_correct [of p c] - by (simp add: algebra_simps) - -lemma poly_eq_0_iff_dvd: - fixes c :: "'a::idom" - shows "poly p c = 0 \ [:-c, 1:] dvd p" -proof - assume "poly p c = 0" - with synthetic_div_correct' [of c p] - have "p = [:-c, 1:] * synthetic_div p c" by simp - then show "[:-c, 1:] dvd p" .. -next - assume "[:-c, 1:] dvd p" - then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) - then show "poly p c = 0" by simp -qed - -lemma dvd_iff_poly_eq_0: - fixes c :: "'a::idom" - shows "[:c, 1:] dvd p \ poly p (-c) = 0" - by (simp add: poly_eq_0_iff_dvd) - -lemma poly_roots_finite: - fixes p :: "'a::idom poly" - shows "p \ 0 \ finite {x. poly p x = 0}" -proof (induct n \ "degree p" arbitrary: p) - case (0 p) - then obtain a where "a \ 0" and "p = [:a:]" - by (cases p, simp split: if_splits) - then show "finite {x. poly p x = 0}" by simp -next - case (Suc n p) - show "finite {x. poly p x = 0}" - proof (cases "\x. poly p x = 0") - case False - then show "finite {x. poly p x = 0}" by simp - next - case True - then obtain a where "poly p a = 0" .. - then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd) - then obtain k where k: "p = [:-a, 1:] * k" .. - with `p \ 0` have "k \ 0" by auto - with k have "degree p = Suc (degree k)" - by (simp add: degree_mult_eq del: mult_pCons_left) - with `Suc n = degree p` have "n = degree k" by simp - with `k \ 0` have "finite {x. poly k x = 0}" by (rule Suc.hyps) - then have "finite (insert a {x. poly k x = 0})" by simp - then show "finite {x. poly p x = 0}" - by (simp add: k uminus_add_conv_diff Collect_disj_eq - del: mult_pCons_left) - qed -qed - - -subsection {* Configuration of the code generator *} - -code_datatype "0::'a::zero poly" pCons - -declare pCons_0_0 [code post] - -instantiation poly :: ("{zero,eq}") eq -begin - -definition [code del]: - "eq_class.eq (p::'a poly) q \ p = q" - -instance - by default (rule eq_poly_def) - -end - -lemma eq_poly_code [code]: - "eq_class.eq (0::_ poly) (0::_ poly) \ True" - "eq_class.eq (0::_ poly) (pCons b q) \ eq_class.eq 0 b \ eq_class.eq 0 q" - "eq_class.eq (pCons a p) (0::_ poly) \ eq_class.eq a 0 \ eq_class.eq p 0" - "eq_class.eq (pCons a p) (pCons b q) \ eq_class.eq a b \ eq_class.eq p q" -unfolding eq by simp_all - -lemmas coeff_code [code] = - coeff_0 coeff_pCons_0 coeff_pCons_Suc - -lemmas degree_code [code] = - degree_0 degree_pCons_eq_if - -lemmas monom_poly_code [code] = - monom_0 monom_Suc - -lemma add_poly_code [code]: - "0 + q = (q :: _ poly)" - "p + 0 = (p :: _ poly)" - "pCons a p + pCons b q = pCons (a + b) (p + q)" -by simp_all - -lemma minus_poly_code [code]: - "- 0 = (0 :: _ poly)" - "- pCons a p = pCons (- a) (- p)" -by simp_all - -lemma diff_poly_code [code]: - "0 - q = (- q :: _ poly)" - "p - 0 = (p :: _ poly)" - "pCons a p - pCons b q = pCons (a - b) (p - q)" -by simp_all - -lemmas smult_poly_code [code] = - smult_0_right smult_pCons - -lemma mult_poly_code [code]: - "0 * q = (0 :: _ poly)" - "pCons a p * q = smult a q + pCons 0 (p * q)" -by simp_all - -lemmas poly_code [code] = - poly_0 poly_pCons - -lemmas synthetic_divmod_code [code] = - synthetic_divmod_0 synthetic_divmod_pCons - -text {* code generator setup for div and mod *} - -definition - pdivmod :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly" -where - [code del]: "pdivmod x y = (x div y, x mod y)" - -lemma div_poly_code [code]: "x div y = fst (pdivmod x y)" - unfolding pdivmod_def by simp - -lemma mod_poly_code [code]: "x mod y = snd (pdivmod x y)" - unfolding pdivmod_def by simp - -lemma pdivmod_0 [code]: "pdivmod 0 y = (0, 0)" - unfolding pdivmod_def by simp - -lemma pdivmod_pCons [code]: - "pdivmod (pCons a x) y = - (if y = 0 then (0, pCons a x) else - (let (q, r) = pdivmod x y; - b = coeff (pCons a r) (degree y) / coeff y (degree y) - in (pCons b q, pCons a r - smult b y)))" -apply (simp add: pdivmod_def Let_def, safe) -apply (rule div_poly_eq) -apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) -apply (rule mod_poly_eq) -apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) -done - -end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Power.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Presburger.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/RComplete.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Rational.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/RealDef.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/RealVector.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Ring_and_Field.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/SEQ.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/SetInterval.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Tools/Qelim/presburger.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Tools/atp_wrapper.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Tools/datatype_codegen.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Tools/datatype_package.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Tools/inductive_package.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Tools/inductive_set_package.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Tools/int_factor_simprocs.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Tools/refute.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Tools/res_atp.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Tools/res_clause.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Tools/res_hol_clause.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Tools/specification_package.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Transitive_Closure.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/Word/Num_Lemmas.thy diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOL/ex/Serbian.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Serbian.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,217 @@ +(* -*- coding: utf-8 -*- :encoding=utf-8: + Author: Filip Maric + +Example theory involving Unicode characters (UTF-8 encoding) -- +Conversion between Serbian cyrillic and latin letters (ÑрпÑка ћирилица и латиница) +*) + +header {* A Serbian theory *} + +theory Serbian +imports Main +begin + +text{* Serbian cyrillic letters *} +datatype azbuka = + azbA ("Ð") +| azbB ("Б") +| azbV ("Ð’") +| azbG ("Г") +| azbD ("Д") +| azbDj ("Ђ") +| azbE ("Е") +| azbZv ("Ж") +| azbZ ("З") +| azbI ("И") +| azbJ ("Ј") +| azbK ("К") +| azbL ("Л") +| azbLj ("Љ") +| azbM ("Ðœ") +| azbN ("Ð") +| azbNj ("Њ") +| azbO ("О") +| azbP ("П") +| azbR ("Р") +| azbS ("С") +| azbT ("Т") +| azbC' ("Ћ") +| azbU ("У") +| azbF ("Ф") +| azbH ("Ð¥") +| azbC ("Ц") +| azbCv ("Ч") +| azbDzv ("Ð") +| azbSv ("Ш") +| azbSpc + +thm azbuka.induct + +text{* Serbian latin letters *} +datatype abeceda = + abcA ("A") +| abcB ("B") +| abcC ("C") +| abcCv ("ÄŒ") +| abcC' ("Ć") +| abcD ("D") +| abcE ("E") +| abcF ("F") +| abcG ("G") +| abcH ("H") +| abcI ("I") +| abcJ ("J") +| abcK ("K") +| abcL ("L") +| abcM ("M") +| abcN ("N") +| abcO ("O") +| abcP ("P") +| abcR ("R") +| abcS ("S") +| abcSv ("Å ") +| abcT ("T") +| abcU ("U") +| abcV ("V") +| abcZ ("Z") +| abcvZ ("Ž") +| abcSpc + +thm abeceda.induct + + +text{* Conversion from cyrillic to latin - + this conversion is valid in all cases *} +primrec azb2abc_aux :: "azbuka \ abeceda list" +where + "azb2abc_aux Ð = [A]" +| "azb2abc_aux Б = [B]" +| "azb2abc_aux Ð’ = [V]" +| "azb2abc_aux Г = [G]" +| "azb2abc_aux Д = [D]" +| "azb2abc_aux Ђ = [D, J]" +| "azb2abc_aux Е = [E]" +| "azb2abc_aux Ж = [Ž]" +| "azb2abc_aux З = [Z]" +| "azb2abc_aux И = [I]" +| "azb2abc_aux Ј = [J]" +| "azb2abc_aux К = [K]" +| "azb2abc_aux Л = [L]" +| "azb2abc_aux Љ = [L, J]" +| "azb2abc_aux Ðœ = [M]" +| "azb2abc_aux Ð = [N]" +| "azb2abc_aux Њ = [N, J]" +| "azb2abc_aux О = [O]" +| "azb2abc_aux П = [P]" +| "azb2abc_aux Р = [R]" +| "azb2abc_aux С = [S]" +| "azb2abc_aux Т = [T]" +| "azb2abc_aux Ћ = [Ć]" +| "azb2abc_aux У = [U]" +| "azb2abc_aux Ф = [F]" +| "azb2abc_aux Ð¥ = [H]" +| "azb2abc_aux Ц = [C]" +| "azb2abc_aux Ч = [ÄŒ]" +| "azb2abc_aux Ð = [D, Ž]" +| "azb2abc_aux Ш = [Å ]" +| "azb2abc_aux azbSpc = [abcSpc]" + +primrec azb2abc :: "azbuka list \ abeceda list" +where + "azb2abc [] = []" +| "azb2abc (x # xs) = azb2abc_aux x @ azb2abc xs" + +value "azb2abc [Д, О, Б, Ð, Р, azbSpc, Д, Ð, Ð, azbSpc, С, Ð’, И, Ðœ, Ð]" +value "azb2abc [Љ, У, Б, И, Ч, И, Ц, Ð, azbSpc, Ð, Ð, azbSpc, П, О, Љ, У]" + +text{* The conversion from latin to cyrillic - + this conversion is valid in most cases but there are some exceptions *} +primrec abc2azb_aux :: "abeceda \ azbuka" +where + "abc2azb_aux A = Ð" +| "abc2azb_aux B = Б" +| "abc2azb_aux C = Ц" +| "abc2azb_aux ÄŒ = Ч" +| "abc2azb_aux Ć = Ћ" +| "abc2azb_aux D = Д" +| "abc2azb_aux E = Е" +| "abc2azb_aux F = Ф" +| "abc2azb_aux G = Г" +| "abc2azb_aux H = Ð¥" +| "abc2azb_aux I = И" +| "abc2azb_aux J = Ј" +| "abc2azb_aux K = К" +| "abc2azb_aux L = Л" +| "abc2azb_aux M = Ðœ" +| "abc2azb_aux N = Ð" +| "abc2azb_aux O = О" +| "abc2azb_aux P = П" +| "abc2azb_aux R = Р" +| "abc2azb_aux S = С" +| "abc2azb_aux Å  = Ш" +| "abc2azb_aux T = Т" +| "abc2azb_aux U = У" +| "abc2azb_aux V = Ð’" +| "abc2azb_aux Z = З" +| "abc2azb_aux Ž = Ж" +| "abc2azb_aux abcSpc = azbSpc" + +fun abc2azb :: "abeceda list \ azbuka list" +where + "abc2azb [] = []" +| "abc2azb [x] = [abc2azb_aux x]" +| "abc2azb (x1 # x2 # xs) = + (if x1 = D \ x2 = J then + Ђ # abc2azb xs + else if x1 = L \ x2 = J then + Љ # abc2azb xs + else if x1 = N \ x2 = J then + Њ # abc2azb xs + else if x1 = D \ x2 = Ž then + Ð # abc2azb xs + else + abc2azb_aux x1 # abc2azb (x2 # xs) + )" + +value "abc2azb [D, O, B, A, R, abcSpc, D, A, N, abcSpc, S, V, I, M, A]" +value "abc2azb [L, J, U, B, I, ÄŒ, I, C, A, abcSpc, N, A, abcSpc, P, O, L, J, U]" + +text{* Here are some invalid conversions *} +lemma "abc2azb [N, A, D, Ž, I, V, E, T, I] = [Ð, Ð, Ð, И, Ð’, Е, Т, И]" + by simp +text{* but it should be: ÐÐДЖИВЕТИ *} +lemma "abc2azb [I, N, J, E, K, C, I, J, A] = [И, Њ, Е, К, Ц, И, Ј, Ð]" + by simp +text{* but it should be: ИÐЈЕКЦИЈР*} + +text{* The conversion fails for all cyrrilic words that contain ÐЈ ЛЈ ДЈ ДЖ *} + + +text{* Idempotency in one direction *} +lemma [simp]: "azb2abc_aux (abc2azb_aux x) = [x]" + by (cases x) auto + +lemma [simp]: "abc2azb (Ž # xs) = Ж # abc2azb xs" + by (cases xs) auto + +lemma [simp]: "abc2azb (J # xs) = Ј # abc2azb xs" + by (cases xs) auto + +theorem "azb2abc (abc2azb x) = x" +proof (induct x) + case (Cons x1 xs) + thus ?case + proof (cases xs) + case (Cons x2 xss) + thus ?thesis + using `azb2abc (abc2azb xs) = xs` + by auto + qed simp +qed simp + +text{* Idempotency in the other direction does not hold *} +lemma "abc2azb (azb2abc [И, Ð, Ј, Е, К, Ц, И, Ј, Ð]) \ [И, Ð, Ј, Е, К, Ц, И, Ј, Ð]" + by simp +text{* It fails for all cyrrilic words that contain ÐЈ ЛЈ ДЈ ДЖ *} + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOLCF/Tools/domain/domain_axioms.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOLCF/Tools/fixrec_package.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/HOLCF/ex/Powerdomain_ex.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/ex/Powerdomain_ex.thy Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,118 @@ +(* Title: HOLCF/ex/Powerdomain_ex.thy + Author: Brian Huffman +*) + +header {* Powerdomain examples *} + +theory Powerdomain_ex +imports HOLCF +begin + +defaultsort bifinite + +subsection {* Monadic sorting example *} + +domain ordering = LT | EQ | GT + +declare ordering.rews [simp] + +definition + compare :: "int lift \ int lift \ ordering" where + "compare = (FLIFT x y. if x < y then LT else if x = y then EQ else GT)" + +definition + is_le :: "int lift \ int lift \ tr" where + "is_le = (\ x y. case compare\x\y of LT \ TT | EQ \ TT | GT \ FF)" + +definition + is_less :: "int lift \ int lift \ tr" where + "is_less = (\ x y. case compare\x\y of LT \ TT | EQ \ FF | GT \ FF)" + +definition + r1 :: "(int lift \ 'a) \ (int lift \ 'a) \ tr convex_pd" where + "r1 = (\ \x,_\ \y,_\. case compare\x\y of + LT \ {TT}\ | + EQ \ {TT, FF}\ | + GT \ {FF}\)" + +definition + r2 :: "(int lift \ 'a) \ (int lift \ 'a) \ tr convex_pd" where + "r2 = (\ \x,_\ \y,_\. {is_le\x\y, is_less\x\y}\)" + +lemma r1_r2: "r1\\x,a\\\y,b\ = (r2\\x,a\\\y,b\ :: tr convex_pd)" +apply (simp add: r1_def r2_def) +apply (simp add: is_le_def is_less_def) +apply (cases "compare\x\y" rule: ordering.casedist) +apply simp_all +done + + +subsection {* Picking a leaf from a tree *} + +domain 'a tree = + Node (lazy "'a tree") (lazy "'a tree") | + Leaf (lazy "'a") + +fixrec + mirror :: "'a tree \ 'a tree" +where + mirror_Leaf: "mirror\(Leaf\a) = Leaf\a" +| mirror_Node: "mirror\(Node\l\r) = Node\(mirror\r)\(mirror\l)" + +fixpat + mirror_strict [simp]: "mirror\\" + +fixrec + pick :: "'a tree \ 'a convex_pd" +where + pick_Leaf: "pick\(Leaf\a) = {a}\" +| pick_Node: "pick\(Node\l\r) = pick\l +\ pick\r" + +fixpat + pick_strict [simp]: "pick\\" + +lemma pick_mirror: "pick\(mirror\t) = pick\t" +by (induct t rule: tree.ind) + (simp_all add: convex_plus_ac) + +fixrec tree1 :: "int lift tree" +where "tree1 = Node\(Node\(Leaf\(Def 1))\(Leaf\(Def 2))) + \(Node\(Leaf\(Def 3))\(Leaf\(Def 4)))" + +fixrec tree2 :: "int lift tree" +where "tree2 = Node\(Node\(Leaf\(Def 1))\(Leaf\(Def 2))) + \(Node\\\(Leaf\(Def 4)))" + +fixrec tree3 :: "int lift tree" +where "tree3 = Node\(Node\(Leaf\(Def 1))\tree3) + \(Node\(Leaf\(Def 3))\(Leaf\(Def 4)))" + +declare tree1_simps tree2_simps tree3_simps [simp del] + +lemma pick_tree1: + "pick\tree1 = {Def 1, Def 2, Def 3, Def 4}\" +apply (subst tree1_simps) +apply simp +apply (simp add: convex_plus_ac) +done + +lemma pick_tree2: + "pick\tree2 = {Def 1, Def 2, \, Def 4}\" +apply (subst tree2_simps) +apply simp +apply (simp add: convex_plus_ac) +done + +lemma pick_tree3: + "pick\tree3 = {Def 1, \, Def 3, Def 4}\" +apply (subst tree3_simps) +apply simp +apply (induct rule: tree3_induct) +apply simp +apply simp +apply (simp add: convex_plus_ac) +apply simp +apply (simp add: convex_plus_ac) +done + +end diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Provers/blast.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Provers/coherent.ML --- a/src/Provers/coherent.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,233 +0,0 @@ -(* Title: Provers/coherent.ML - Author: Stefan Berghofer, TU Muenchen - Author: Marc Bezem, Institutt for Informatikk, Universitetet i Bergen - -Prover for coherent logic, see e.g. - - Marc Bezem and Thierry Coquand, Automating Coherent Logic, LPAR 2005 - -for a description of the algorithm. -*) - -signature COHERENT_DATA = -sig - val atomize_elimL: thm - val atomize_exL: thm - val atomize_conjL: thm - val atomize_disjL: thm - val operator_names: string list -end; - -signature COHERENT = -sig - val verbose: bool ref - val show_facts: bool ref - val coherent_tac: thm list -> Proof.context -> int -> tactic - val coherent_meth: thm list -> Proof.context -> Proof.method - val setup: theory -> theory -end; - -functor CoherentFun(Data: COHERENT_DATA) : COHERENT = -struct - -val verbose = ref false; - -fun message f = if !verbose then tracing (f ()) else (); - -datatype cl_prf = - ClPrf of thm * (Type.tyenv * Envir.tenv) * ((indexname * typ) * term) list * - int list * (term list * cl_prf) list; - -val is_atomic = not o exists_Const (member (op =) Data.operator_names o #1); - -local open Conv in - -fun rulify_elim_conv ct = - if is_atomic (Logic.strip_imp_concl (term_of ct)) then all_conv ct - else concl_conv (length (Logic.strip_imp_prems (term_of ct))) - (rewr_conv (symmetric Data.atomize_elimL) then_conv - MetaSimplifier.rewrite true (map symmetric - [Data.atomize_exL, Data.atomize_conjL, Data.atomize_disjL])) ct - -end; - -fun rulify_elim th = MetaSimplifier.norm_hhf (Conv.fconv_rule rulify_elim_conv th); - -(* Decompose elimination rule of the form - A1 ==> ... ==> Am ==> (!!xs1. Bs1 ==> P) ==> ... ==> (!!xsn. Bsn ==> P) ==> P -*) -fun dest_elim prop = - let - val prems = Logic.strip_imp_prems prop; - val concl = Logic.strip_imp_concl prop; - val (prems1, prems2) = - take_suffix (fn t => Logic.strip_assums_concl t = concl) prems; - in - (prems1, - if null prems2 then [([], [concl])] - else map (fn t => - (map snd (Logic.strip_params t), Logic.strip_assums_hyp t)) prems2) - end; - -fun mk_rule th = - let - val th' = rulify_elim th; - val (prems, cases) = dest_elim (prop_of th') - in (th', prems, cases) end; - -fun mk_dom ts = fold (fn t => - Typtab.map_default (fastype_of t, []) (fn us => us @ [t])) ts Typtab.empty; - -val empty_env = (Vartab.empty, Vartab.empty); - -(* Find matcher that makes conjunction valid in given state *) -fun valid_conj ctxt facts env [] = Seq.single (env, []) - | valid_conj ctxt facts env (t :: ts) = - Seq.maps (fn (u, x) => Seq.map (apsnd (cons x)) - (valid_conj ctxt facts - (Pattern.match (ProofContext.theory_of ctxt) (t, u) env) ts - handle Pattern.MATCH => Seq.empty)) - (Seq.of_list (sort (int_ord o pairself snd) (Net.unify_term facts t))); - -(* Instantiate variables that only occur free in conlusion *) -fun inst_extra_vars ctxt dom cs = - let - val vs = fold Term.add_vars (maps snd cs) []; - fun insts [] inst = Seq.single inst - | insts ((ixn, T) :: vs') inst = Seq.maps - (fn t => insts vs' (((ixn, T), t) :: inst)) - (Seq.of_list (case Typtab.lookup dom T of - NONE => error ("Unknown domain: " ^ - Syntax.string_of_typ ctxt T ^ "\nfor term(s) " ^ - commas (maps (map (Syntax.string_of_term ctxt) o snd) cs)) - | SOME ts => ts)) - in Seq.map (fn inst => - (inst, map (apsnd (map (subst_Vars (map (apfst fst) inst)))) cs)) - (insts vs []) - end; - -(* Check whether disjunction is valid in given state *) -fun is_valid_disj ctxt facts [] = false - | is_valid_disj ctxt facts ((Ts, ts) :: ds) = - let val vs = rev (map_index (fn (i, T) => Var (("x", i), T)) Ts) - in case Seq.pull (valid_conj ctxt facts empty_env - (map (fn t => subst_bounds (vs, t)) ts)) of - SOME _ => true - | NONE => is_valid_disj ctxt facts ds - end; - -val show_facts = ref false; - -fun string_of_facts ctxt s facts = space_implode "\n" - (s :: map (Syntax.string_of_term ctxt) - (map fst (sort (int_ord o pairself snd) (Net.content facts)))) ^ "\n\n"; - -fun print_facts ctxt facts = - if !show_facts then message (fn () => string_of_facts ctxt "Facts:" facts) - else (); - -fun valid ctxt rules goal dom facts nfacts nparams = - let val seq = Seq.of_list rules |> Seq.maps (fn (th, ps, cs) => - valid_conj ctxt facts empty_env ps |> Seq.maps (fn (env as (tye, _), is) => - let val cs' = map (fn (Ts, ts) => - (map (Envir.typ_subst_TVars tye) Ts, map (Envir.subst_vars env) ts)) cs - in - inst_extra_vars ctxt dom cs' |> - Seq.map_filter (fn (inst, cs'') => - if is_valid_disj ctxt facts cs'' then NONE - else SOME (th, env, inst, is, cs'')) - end)) - in - case Seq.pull seq of - NONE => (tracing (string_of_facts ctxt "Countermodel found:" facts); NONE) - | SOME ((th, env, inst, is, cs), _) => - if cs = [([], [goal])] then SOME (ClPrf (th, env, inst, is, [])) - else - (case valid_cases ctxt rules goal dom facts nfacts nparams cs of - NONE => NONE - | SOME prfs => SOME (ClPrf (th, env, inst, is, prfs))) - end - -and valid_cases ctxt rules goal dom facts nfacts nparams [] = SOME [] - | valid_cases ctxt rules goal dom facts nfacts nparams ((Ts, ts) :: ds) = - let - val _ = message (fn () => "case " ^ commas (map (Syntax.string_of_term ctxt) ts)); - val params = rev (map_index (fn (i, T) => - Free ("par" ^ string_of_int (nparams + i), T)) Ts); - val ts' = map_index (fn (i, t) => - (subst_bounds (params, t), nfacts + i)) ts; - val dom' = fold (fn (T, p) => - Typtab.map_default (T, []) (fn ps => ps @ [p])) - (Ts ~~ params) dom; - val facts' = fold (fn (t, i) => Net.insert_term op = - (t, (t, i))) ts' facts - in - case valid ctxt rules goal dom' facts' - (nfacts + length ts) (nparams + length Ts) of - NONE => NONE - | SOME prf => (case valid_cases ctxt rules goal dom facts nfacts nparams ds of - NONE => NONE - | SOME prfs => SOME ((params, prf) :: prfs)) - end; - -(** proof replaying **) - -fun thm_of_cl_prf thy goal asms (ClPrf (th, (tye, env), insts, is, prfs)) = - let - val _ = message (fn () => space_implode "\n" - ("asms:" :: map Display.string_of_thm asms) ^ "\n\n"); - val th' = Drule.implies_elim_list - (Thm.instantiate - (map (fn (ixn, (S, T)) => - (Thm.ctyp_of thy (TVar ((ixn, S))), Thm.ctyp_of thy T)) - (Vartab.dest tye), - map (fn (ixn, (T, t)) => - (Thm.cterm_of thy (Var (ixn, Envir.typ_subst_TVars tye T)), - Thm.cterm_of thy t)) (Vartab.dest env) @ - map (fn (ixnT, t) => - (Thm.cterm_of thy (Var ixnT), Thm.cterm_of thy t)) insts) th) - (map (nth asms) is); - val (_, cases) = dest_elim (prop_of th') - in - case (cases, prfs) of - ([([], [_])], []) => th' - | ([([], [_])], [([], prf)]) => thm_of_cl_prf thy goal (asms @ [th']) prf - | _ => Drule.implies_elim_list - (Thm.instantiate (Thm.match - (Drule.strip_imp_concl (cprop_of th'), goal)) th') - (map (thm_of_case_prf thy goal asms) (prfs ~~ cases)) - end - -and thm_of_case_prf thy goal asms ((params, prf), (_, asms')) = - let - val cparams = map (cterm_of thy) params; - val asms'' = map (cterm_of thy o curry subst_bounds (rev params)) asms' - in - Drule.forall_intr_list cparams (Drule.implies_intr_list asms'' - (thm_of_cl_prf thy goal (asms @ map Thm.assume asms'') prf)) - end; - - -(** external interface **) - -fun coherent_tac rules ctxt = SUBPROOF (fn {prems, concl, params, context, ...} => - rtac (rulify_elim_conv concl RS equal_elim_rule2) 1 THEN - SUBPROOF (fn {prems = prems', concl, context, ...} => - let val xs = map term_of params @ - map (fn (_, s) => Free (s, the (Variable.default_type context s))) - (Variable.fixes_of context) - in - case valid context (map mk_rule (prems' @ prems @ rules)) (term_of concl) - (mk_dom xs) Net.empty 0 0 of - NONE => no_tac - | SOME prf => - rtac (thm_of_cl_prf (ProofContext.theory_of context) concl [] prf) 1 - end) context 1) ctxt; - -fun coherent_meth rules ctxt = - Method.METHOD (fn facts => coherent_tac (facts @ rules) ctxt 1); - -val setup = Method.add_method - ("coherent", Method.thms_ctxt_args coherent_meth, "Prove coherent formula"); - -end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Provers/eqsubst.ML --- a/src/Provers/eqsubst.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,575 +0,0 @@ -(* Title: Provers/eqsubst.ML - Author: Lucas Dixon, University of Edinburgh - -A proof method to perform a substiution using an equation. -*) - -signature EQSUBST = -sig - (* a type abbreviation for match information *) - type match = - ((indexname * (sort * typ)) list (* type instantiations *) - * (indexname * (typ * term)) list) (* term instantiations *) - * (string * typ) list (* fake named type abs env *) - * (string * typ) list (* type abs env *) - * term (* outer term *) - - type searchinfo = - theory - * int (* maxidx *) - * Zipper.T (* focusterm to search under *) - - exception eqsubst_occL_exp of - string * int list * Thm.thm list * int * Thm.thm - - (* low level substitution functions *) - val apply_subst_in_asm : - int -> - Thm.thm -> - Thm.thm -> - (Thm.cterm list * int * 'a * Thm.thm) * match -> Thm.thm Seq.seq - val apply_subst_in_concl : - int -> - Thm.thm -> - Thm.cterm list * Thm.thm -> - Thm.thm -> match -> Thm.thm Seq.seq - - (* matching/unification within zippers *) - val clean_match_z : - Context.theory -> Term.term -> Zipper.T -> match option - val clean_unify_z : - Context.theory -> int -> Term.term -> Zipper.T -> match Seq.seq - - (* skipping things in seq seq's *) - - (* skipping non-empty sub-sequences but when we reach the end - of the seq, remembering how much we have left to skip. *) - datatype 'a skipseq = SkipMore of int - | SkipSeq of 'a Seq.seq Seq.seq; - - val skip_first_asm_occs_search : - ('a -> 'b -> 'c Seq.seq Seq.seq) -> - 'a -> int -> 'b -> 'c skipseq - val skip_first_occs_search : - int -> ('a -> 'b -> 'c Seq.seq Seq.seq) -> 'a -> 'b -> 'c Seq.seq - val skipto_skipseq : int -> 'a Seq.seq Seq.seq -> 'a skipseq - - (* tactics *) - val eqsubst_asm_tac : - Proof.context -> - int list -> Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq - val eqsubst_asm_tac' : - Proof.context -> - (searchinfo -> int -> Term.term -> match skipseq) -> - int -> Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq - val eqsubst_tac : - Proof.context -> - int list -> (* list of occurences to rewrite, use [0] for any *) - Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq - val eqsubst_tac' : - Proof.context -> (* proof context *) - (searchinfo -> Term.term -> match Seq.seq) (* search function *) - -> Thm.thm (* equation theorem to rewrite with *) - -> int (* subgoal number in goal theorem *) - -> Thm.thm (* goal theorem *) - -> Thm.thm Seq.seq (* rewritten goal theorem *) - - - val fakefree_badbounds : - (string * Term.typ) list -> - Term.term -> - (string * Term.typ) list * (string * Term.typ) list * Term.term - - val mk_foo_match : - (Term.term -> Term.term) -> - ('a * Term.typ) list -> Term.term -> Term.term - - (* preparing substitution *) - val prep_meta_eq : Proof.context -> Thm.thm -> Thm.thm list - val prep_concl_subst : - int -> Thm.thm -> (Thm.cterm list * Thm.thm) * searchinfo - val prep_subst_in_asm : - int -> Thm.thm -> int -> - (Thm.cterm list * int * int * Thm.thm) * searchinfo - val prep_subst_in_asms : - int -> Thm.thm -> - ((Thm.cterm list * int * int * Thm.thm) * searchinfo) list - val prep_zipper_match : - Zipper.T -> Term.term * ((string * Term.typ) list * (string * Term.typ) list * Term.term) - - (* search for substitutions *) - val valid_match_start : Zipper.T -> bool - val search_lr_all : Zipper.T -> Zipper.T Seq.seq - val search_lr_valid : (Zipper.T -> bool) -> Zipper.T -> Zipper.T Seq.seq - val searchf_lr_unify_all : - searchinfo -> Term.term -> match Seq.seq Seq.seq - val searchf_lr_unify_valid : - searchinfo -> Term.term -> match Seq.seq Seq.seq - val searchf_bt_unify_valid : - searchinfo -> Term.term -> match Seq.seq Seq.seq - - (* syntax tools *) - val ith_syntax : Args.T list -> int list * Args.T list - val options_syntax : Args.T list -> bool * Args.T list - - (* Isar level hooks *) - val eqsubst_asm_meth : Proof.context -> int list -> Thm.thm list -> Proof.method - val eqsubst_meth : Proof.context -> int list -> Thm.thm list -> Proof.method - val subst_meth : Method.src -> Proof.context -> Proof.method - val setup : theory -> theory - -end; - -structure EqSubst -: EQSUBST -= struct - -structure Z = Zipper; - -(* changes object "=" to meta "==" which prepares a given rewrite rule *) -fun prep_meta_eq ctxt = - let val (_, {mk_rews = {mk, ...}, ...}) = Simplifier.rep_ss (Simplifier.local_simpset_of ctxt) - in mk #> map Drule.zero_var_indexes end; - - - (* a type abriviation for match information *) - type match = - ((indexname * (sort * typ)) list (* type instantiations *) - * (indexname * (typ * term)) list) (* term instantiations *) - * (string * typ) list (* fake named type abs env *) - * (string * typ) list (* type abs env *) - * term (* outer term *) - - type searchinfo = - theory - * int (* maxidx *) - * Zipper.T (* focusterm to search under *) - - -(* skipping non-empty sub-sequences but when we reach the end - of the seq, remembering how much we have left to skip. *) -datatype 'a skipseq = SkipMore of int - | SkipSeq of 'a Seq.seq Seq.seq; -(* given a seqseq, skip the first m non-empty seq's, note deficit *) -fun skipto_skipseq m s = - let - fun skip_occs n sq = - case Seq.pull sq of - NONE => SkipMore n - | SOME (h,t) => - (case Seq.pull h of NONE => skip_occs n t - | SOME _ => if n <= 1 then SkipSeq (Seq.cons h t) - else skip_occs (n - 1) t) - in (skip_occs m s) end; - -(* note: outerterm is the taget with the match replaced by a bound - variable : ie: "P lhs" beocmes "%x. P x" - insts is the types of instantiations of vars in lhs - and typinsts is the type instantiations of types in the lhs - Note: Final rule is the rule lifted into the ontext of the - taget thm. *) -fun mk_foo_match mkuptermfunc Ts t = - let - val ty = Term.type_of t - val bigtype = (rev (map snd Ts)) ---> ty - fun mk_foo 0 t = t - | mk_foo i t = mk_foo (i - 1) (t $ (Bound (i - 1))) - val num_of_bnds = (length Ts) - (* foo_term = "fooabs y0 ... yn" where y's are local bounds *) - val foo_term = mk_foo num_of_bnds (Bound num_of_bnds) - in Abs("fooabs", bigtype, mkuptermfunc foo_term) end; - -(* T is outer bound vars, n is number of locally bound vars *) -(* THINK: is order of Ts correct...? or reversed? *) -fun fakefree_badbounds Ts t = - let val (FakeTs,Ts,newnames) = - List.foldr (fn ((n,ty),(FakeTs,Ts,usednames)) => - let val newname = Name.variant usednames n - in ((RWTools.mk_fake_bound_name newname,ty)::FakeTs, - (newname,ty)::Ts, - newname::usednames) end) - ([],[],[]) - Ts - in (FakeTs, Ts, Term.subst_bounds (map Free FakeTs, t)) end; - -(* before matching we need to fake the bound vars that are missing an -abstraction. In this function we additionally construct the -abstraction environment, and an outer context term (with the focus -abstracted out) for use in rewriting with RWInst.rw *) -fun prep_zipper_match z = - let - val t = Z.trm z - val c = Z.ctxt z - val Ts = Z.C.nty_ctxt c - val (FakeTs', Ts', t') = fakefree_badbounds Ts t - val absterm = mk_foo_match (Z.C.apply c) Ts' t' - in - (t', (FakeTs', Ts', absterm)) - end; - -(* Matching and Unification with exception handled *) -fun clean_match thy (a as (pat, t)) = - let val (tyenv, tenv) = Pattern.match thy a (Vartab.empty, Vartab.empty) - in SOME (Vartab.dest tyenv, Vartab.dest tenv) - end handle Pattern.MATCH => NONE; - -(* given theory, max var index, pat, tgt; returns Seq of instantiations *) -fun clean_unify thry ix (a as (pat, tgt)) = - let - (* type info will be re-derived, maybe this can be cached - for efficiency? *) - val pat_ty = Term.type_of pat; - val tgt_ty = Term.type_of tgt; - (* is it OK to ignore the type instantiation info? - or should I be using it? *) - val typs_unify = - SOME (Sign.typ_unify thry (pat_ty, tgt_ty) (Vartab.empty, ix)) - handle Type.TUNIFY => NONE; - in - case typs_unify of - SOME (typinsttab, ix2) => - let - (* is it right to throw away the flexes? - or should I be using them somehow? *) - fun mk_insts env = - (Vartab.dest (Envir.type_env env), - Envir.alist_of env); - val initenv = Envir.Envir {asol = Vartab.empty, - iTs = typinsttab, maxidx = ix2}; - val useq = Unify.smash_unifiers thry [a] initenv - handle UnequalLengths => Seq.empty - | Term.TERM _ => Seq.empty; - fun clean_unify' useq () = - (case (Seq.pull useq) of - NONE => NONE - | SOME (h,t) => SOME (mk_insts h, Seq.make (clean_unify' t))) - handle UnequalLengths => NONE - | Term.TERM _ => NONE - in - (Seq.make (clean_unify' useq)) - end - | NONE => Seq.empty - end; - -(* Matching and Unification for zippers *) -(* Note: Ts is a modified version of the original names of the outer -bound variables. New names have been introduced to make sure they are -unique w.r.t all names in the term and each other. usednames' is -oldnames + new names. *) -fun clean_match_z thy pat z = - let val (t, (FakeTs,Ts,absterm)) = prep_zipper_match z in - case clean_match thy (pat, t) of - NONE => NONE - | SOME insts => SOME (insts, FakeTs, Ts, absterm) end; -(* ix = max var index *) -fun clean_unify_z sgn ix pat z = - let val (t, (FakeTs, Ts,absterm)) = prep_zipper_match z in - Seq.map (fn insts => (insts, FakeTs, Ts, absterm)) - (clean_unify sgn ix (t, pat)) end; - - -(* FOR DEBUGGING... -type trace_subst_errT = int (* subgoal *) - * thm (* thm with all goals *) - * (Thm.cterm list (* certified free var placeholders for vars *) - * thm) (* trivial thm of goal concl *) - (* possible matches/unifiers *) - * thm (* rule *) - * (((indexname * typ) list (* type instantiations *) - * (indexname * term) list ) (* term instantiations *) - * (string * typ) list (* Type abs env *) - * term) (* outer term *); - -val trace_subst_err = (ref NONE : trace_subst_errT option ref); -val trace_subst_search = ref false; -exception trace_subst_exp of trace_subst_errT; -*) - - -fun bot_left_leaf_of (l $ r) = bot_left_leaf_of l - | bot_left_leaf_of (Abs(s,ty,t)) = bot_left_leaf_of t - | bot_left_leaf_of x = x; - -(* Avoid considering replacing terms which have a var at the head as - they always succeed trivially, and uninterestingly. *) -fun valid_match_start z = - (case bot_left_leaf_of (Z.trm z) of - Var _ => false - | _ => true); - -(* search from top, left to right, then down *) -val search_lr_all = ZipperSearch.all_bl_ur; - -(* search from top, left to right, then down *) -fun search_lr_valid validf = - let - fun sf_valid_td_lr z = - let val here = if validf z then [Z.Here z] else [] in - case Z.trm z - of _ $ _ => [Z.LookIn (Z.move_down_left z)] - @ here - @ [Z.LookIn (Z.move_down_right z)] - | Abs _ => here @ [Z.LookIn (Z.move_down_abs z)] - | _ => here - end; - in Z.lzy_search sf_valid_td_lr end; - -(* search from bottom to top, left to right *) - -fun search_bt_valid validf = - let - fun sf_valid_td_lr z = - let val here = if validf z then [Z.Here z] else [] in - case Z.trm z - of _ $ _ => [Z.LookIn (Z.move_down_left z), - Z.LookIn (Z.move_down_right z)] @ here - | Abs _ => [Z.LookIn (Z.move_down_abs z)] @ here - | _ => here - end; - in Z.lzy_search sf_valid_td_lr end; - -fun searchf_unify_gen f (sgn, maxidx, z) lhs = - Seq.map (clean_unify_z sgn maxidx lhs) - (Z.limit_apply f z); - -(* search all unifications *) -val searchf_lr_unify_all = - searchf_unify_gen search_lr_all; - -(* search only for 'valid' unifiers (non abs subterms and non vars) *) -val searchf_lr_unify_valid = - searchf_unify_gen (search_lr_valid valid_match_start); - -val searchf_bt_unify_valid = - searchf_unify_gen (search_bt_valid valid_match_start); - -(* apply a substitution in the conclusion of the theorem th *) -(* cfvs are certified free var placeholders for goal params *) -(* conclthm is a theorem of for just the conclusion *) -(* m is instantiation/match information *) -(* rule is the equation for substitution *) -fun apply_subst_in_concl i th (cfvs, conclthm) rule m = - (RWInst.rw m rule conclthm) - |> IsaND.unfix_frees cfvs - |> RWInst.beta_eta_contract - |> (fn r => Tactic.rtac r i th); - -(* substitute within the conclusion of goal i of gth, using a meta -equation rule. Note that we assume rule has var indicies zero'd *) -fun prep_concl_subst i gth = - let - val th = Thm.incr_indexes 1 gth; - val tgt_term = Thm.prop_of th; - - val sgn = Thm.theory_of_thm th; - val ctermify = Thm.cterm_of sgn; - val trivify = Thm.trivial o ctermify; - - val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term; - val cfvs = rev (map ctermify fvs); - - val conclterm = Logic.strip_imp_concl fixedbody; - val conclthm = trivify conclterm; - val maxidx = Thm.maxidx_of th; - val ft = ((Z.move_down_right (* ==> *) - o Z.move_down_left (* Trueprop *) - o Z.mktop - o Thm.prop_of) conclthm) - in - ((cfvs, conclthm), (sgn, maxidx, ft)) - end; - -(* substitute using an object or meta level equality *) -fun eqsubst_tac' ctxt searchf instepthm i th = - let - val (cvfsconclthm, searchinfo) = prep_concl_subst i th; - val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm); - fun rewrite_with_thm r = - let val (lhs,_) = Logic.dest_equals (Thm.concl_of r); - in searchf searchinfo lhs - |> Seq.maps (apply_subst_in_concl i th cvfsconclthm r) end; - in stepthms |> Seq.maps rewrite_with_thm end; - - -(* distinct subgoals *) -fun distinct_subgoals th = - the_default th (SINGLE distinct_subgoals_tac th); - -(* General substitution of multiple occurances using one of - the given theorems*) - - -exception eqsubst_occL_exp of - string * (int list) * (thm list) * int * thm; -fun skip_first_occs_search occ srchf sinfo lhs = - case (skipto_skipseq occ (srchf sinfo lhs)) of - SkipMore _ => Seq.empty - | SkipSeq ss => Seq.flat ss; - -(* The occL is a list of integers indicating which occurence -w.r.t. the search order, to rewrite. Backtracking will also find later -occurences, but all earlier ones are skipped. Thus you can use [0] to -just find all rewrites. *) - -fun eqsubst_tac ctxt occL thms i th = - let val nprems = Thm.nprems_of th in - if nprems < i then Seq.empty else - let val thmseq = (Seq.of_list thms) - fun apply_occ occ th = - thmseq |> Seq.maps - (fn r => eqsubst_tac' - ctxt - (skip_first_occs_search - occ searchf_lr_unify_valid) r - (i + ((Thm.nprems_of th) - nprems)) - th); - val sortedoccL = - Library.sort (Library.rev_order o Library.int_ord) occL; - in - Seq.map distinct_subgoals (Seq.EVERY (map apply_occ sortedoccL) th) - end - end - handle THM _ => raise eqsubst_occL_exp ("THM",occL,thms,i,th); - - -(* inthms are the given arguments in Isar, and treated as eqstep with - the first one, then the second etc *) -fun eqsubst_meth ctxt occL inthms = - Method.SIMPLE_METHOD' (eqsubst_tac ctxt occL inthms); - -(* apply a substitution inside assumption j, keeps asm in the same place *) -fun apply_subst_in_asm i th rule ((cfvs, j, ngoalprems, pth),m) = - let - val th2 = Thm.rotate_rule (j - 1) i th; (* put premice first *) - val preelimrule = - (RWInst.rw m rule pth) - |> (Seq.hd o prune_params_tac) - |> Thm.permute_prems 0 ~1 (* put old asm first *) - |> IsaND.unfix_frees cfvs (* unfix any global params *) - |> RWInst.beta_eta_contract; (* normal form *) - (* val elimrule = - preelimrule - |> Tactic.make_elim (* make into elim rule *) - |> Thm.lift_rule (th2, i); (* lift into context *) - *) - in - (* ~j because new asm starts at back, thus we subtract 1 *) - Seq.map (Thm.rotate_rule (~j) ((Thm.nprems_of rule) + i)) - (Tactic.dtac preelimrule i th2) - - (* (Thm.bicompose - false (* use unification *) - (true, (* elim resolution *) - elimrule, (2 + (Thm.nprems_of rule)) - ngoalprems) - i th2) *) - end; - - -(* prepare to substitute within the j'th premise of subgoal i of gth, -using a meta-level equation. Note that we assume rule has var indicies -zero'd. Note that we also assume that premt is the j'th premice of -subgoal i of gth. Note the repetition of work done for each -assumption, i.e. this can be made more efficient for search over -multiple assumptions. *) -fun prep_subst_in_asm i gth j = - let - val th = Thm.incr_indexes 1 gth; - val tgt_term = Thm.prop_of th; - - val sgn = Thm.theory_of_thm th; - val ctermify = Thm.cterm_of sgn; - val trivify = Thm.trivial o ctermify; - - val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term; - val cfvs = rev (map ctermify fvs); - - val asmt = nth (Logic.strip_imp_prems fixedbody) (j - 1); - val asm_nprems = length (Logic.strip_imp_prems asmt); - - val pth = trivify asmt; - val maxidx = Thm.maxidx_of th; - - val ft = ((Z.move_down_right (* trueprop *) - o Z.mktop - o Thm.prop_of) pth) - in ((cfvs, j, asm_nprems, pth), (sgn, maxidx, ft)) end; - -(* prepare subst in every possible assumption *) -fun prep_subst_in_asms i gth = - map (prep_subst_in_asm i gth) - ((fn l => Library.upto (1, length l)) - (Logic.prems_of_goal (Thm.prop_of gth) i)); - - -(* substitute in an assumption using an object or meta level equality *) -fun eqsubst_asm_tac' ctxt searchf skipocc instepthm i th = - let - val asmpreps = prep_subst_in_asms i th; - val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm); - fun rewrite_with_thm r = - let val (lhs,_) = Logic.dest_equals (Thm.concl_of r) - fun occ_search occ [] = Seq.empty - | occ_search occ ((asminfo, searchinfo)::moreasms) = - (case searchf searchinfo occ lhs of - SkipMore i => occ_search i moreasms - | SkipSeq ss => - Seq.append (Seq.map (Library.pair asminfo) (Seq.flat ss)) - (occ_search 1 moreasms)) - (* find later substs also *) - in - occ_search skipocc asmpreps |> Seq.maps (apply_subst_in_asm i th r) - end; - in stepthms |> Seq.maps rewrite_with_thm end; - - -fun skip_first_asm_occs_search searchf sinfo occ lhs = - skipto_skipseq occ (searchf sinfo lhs); - -fun eqsubst_asm_tac ctxt occL thms i th = - let val nprems = Thm.nprems_of th - in - if nprems < i then Seq.empty else - let val thmseq = (Seq.of_list thms) - fun apply_occ occK th = - thmseq |> Seq.maps - (fn r => - eqsubst_asm_tac' ctxt (skip_first_asm_occs_search - searchf_lr_unify_valid) occK r - (i + ((Thm.nprems_of th) - nprems)) - th); - val sortedoccs = - Library.sort (Library.rev_order o Library.int_ord) occL - in - Seq.map distinct_subgoals - (Seq.EVERY (map apply_occ sortedoccs) th) - end - end - handle THM _ => raise eqsubst_occL_exp ("THM",occL,thms,i,th); - -(* inthms are the given arguments in Isar, and treated as eqstep with - the first one, then the second etc *) -fun eqsubst_asm_meth ctxt occL inthms = - Method.SIMPLE_METHOD' (eqsubst_asm_tac ctxt occL inthms); - -(* syntax for options, given "(asm)" will give back true, without - gives back false *) -val options_syntax = - (Args.parens (Args.$$$ "asm") >> (K true)) || - (Scan.succeed false); - -val ith_syntax = - Scan.optional (Args.parens (Scan.repeat OuterParse.nat)) [0]; - -(* combination method that takes a flag (true indicates that subst -should be done to an assumption, false = apply to the conclusion of -the goal) as well as the theorems to use *) -fun subst_meth src = - Method.syntax ((Scan.lift options_syntax) -- (Scan.lift ith_syntax) -- Attrib.thms) src - #> (fn (((asmflag, occL), inthms), ctxt) => - (if asmflag then eqsubst_asm_meth else eqsubst_meth) ctxt occL inthms); - - -val setup = - Method.add_method ("subst", subst_meth, "single-step substitution"); - -end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Provers/project_rule.ML --- a/src/Provers/project_rule.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,63 +0,0 @@ -(* Title: Provers/project_rule.ML - ID: $Id$ - Author: Makarius - -Transform mutual rule: - HH ==> (x1:A1 --> P1 x1) & ... & (xn:An --> Pn xn) -into projection: - xi:Ai ==> HH ==> Pi xi -*) - -signature PROJECT_RULE_DATA = -sig - val conjunct1: thm - val conjunct2: thm - val mp: thm -end; - -signature PROJECT_RULE = -sig - val project: Proof.context -> int -> thm -> thm - val projects: Proof.context -> int list -> thm -> thm list - val projections: Proof.context -> thm -> thm list -end; - -functor ProjectRuleFun(Data: PROJECT_RULE_DATA): PROJECT_RULE = -struct - -fun conj1 th = th RS Data.conjunct1; -fun conj2 th = th RS Data.conjunct2; -fun imp th = th RS Data.mp; - -fun projects ctxt is raw_rule = - let - fun proj 1 th = the_default th (try conj1 th) - | proj k th = proj (k - 1) (conj2 th); - fun prems k th = - (case try imp th of - NONE => (k, th) - | SOME th' => prems (k + 1) th'); - val ((_, [rule]), ctxt') = Variable.import_thms true [raw_rule] ctxt; - fun result i = - rule - |> proj i - |> prems 0 |-> (fn k => - Thm.permute_prems 0 (~ k) - #> singleton (Variable.export ctxt' ctxt) - #> Drule.zero_var_indexes - #> RuleCases.save raw_rule - #> RuleCases.add_consumes k); - in map result is end; - -fun project ctxt i th = hd (projects ctxt [i] th); - -fun projections ctxt raw_rule = - let - fun projs k th = - (case try conj2 th of - NONE => k - | SOME th' => projs (k + 1) th'); - val ((_, [rule]), _) = Variable.import_thms true [raw_rule] ctxt; - in projects ctxt (1 upto projs 1 rule) raw_rule end; - -end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/General/binding.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/General/name_space.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/IsaMakefile diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/ROOT.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/attrib.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/class.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/code.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/constdefs.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/element.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/expression.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/find_consts.ML --- a/src/Pure/Isar/find_consts.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,120 +0,0 @@ -(* Title: find_consts.ML - Author: Timothy Bourke and Gerwin Klein, NICTA - - Hoogle-like (http://www-users.cs.york.ac.uk/~ndm/hoogle) searching by type - over constants, but matching is not fuzzy -*) - -signature FIND_CONSTS = -sig - datatype criterion = Strict of string - | Loose of string - | Name of string - - val default_criteria : (bool * criterion) list ref - - val find_consts : Proof.context -> (bool * criterion) list -> unit -end; - -structure FindConsts : FIND_CONSTS = -struct - -datatype criterion = Strict of string - | Loose of string - | Name of string; - -val default_criteria = ref [(false, Name ".sko_")]; - -fun add_tye (_, (_, t)) n = size_of_typ t + n; - -fun matches_subtype thy typat = let - val p = can (fn ty => Sign.typ_match thy (typat, ty) Vartab.empty); - - fun fs [] = false - | fs (t::ts) = f t orelse fs ts - - and f (t as Type (_, ars)) = p t orelse fs ars - | f t = p t; - in f end; - -fun check_const p (nm, (ty, _)) = if p (nm, ty) - then SOME (size_of_typ ty) - else NONE; - -fun opt_not f (c as (_, (ty, _))) = if is_some (f c) - then NONE else SOME (size_of_typ ty); - -fun filter_const (_, NONE) = NONE - | filter_const (f, (SOME (c, r))) = Option.map - (pair c o ((curry Int.min) r)) (f c); - -fun pretty_criterion (b, c) = - let - fun prfx s = if b then s else "-" ^ s; - in - (case c of - Strict pat => Pretty.str (prfx "strict: " ^ quote pat) - | Loose pat => Pretty.str (prfx (quote pat)) - | Name name => Pretty.str (prfx "name: " ^ quote name)) - end; - -fun pretty_const ctxt (nm, ty) = let - val ty' = Logic.unvarifyT ty; - in - Pretty.block [Pretty.quote (Pretty.str nm), Pretty.fbrk, - Pretty.str "::", Pretty.brk 1, - Pretty.quote (Syntax.pretty_typ ctxt ty')] - end; - -fun find_consts ctxt raw_criteria = let - val start = start_timing (); - - val thy = ProofContext.theory_of ctxt; - val low_ranking = 10000; - - fun make_pattern crit = ProofContext.read_term_pattern ctxt ("_::" ^ crit) - |> type_of; - - fun make_match (Strict arg) = - let val qty = make_pattern arg; in - fn (_, (ty, _)) => let - val tye = Sign.typ_match thy (qty, ty) Vartab.empty; - val sub_size = Vartab.fold add_tye tye 0; - in SOME sub_size end handle MATCH => NONE - end - - | make_match (Loose arg) = - check_const (matches_subtype thy (make_pattern arg) o snd) - - | make_match (Name arg) = check_const (match_string arg o fst); - - fun make_criterion (b, crit) = (if b then I else opt_not) (make_match crit); - val criteria = map make_criterion ((!default_criteria) @ raw_criteria); - - val (_, consts) = (#constants o Consts.dest o Sign.consts_of) thy; - fun eval_entry c = foldl filter_const (SOME (c, low_ranking)) criteria; - - val matches = Symtab.fold (cons o eval_entry) consts [] - |> map_filter I - |> sort (rev_order o int_ord o pairself snd) - |> map ((apsnd fst) o fst); - - val end_msg = " in " ^ - (List.nth (String.tokens Char.isSpace (end_timing start), 3)) - ^ " secs" - in - Pretty.big_list "searched for:" (map pretty_criterion raw_criteria) - :: Pretty.str "" - :: (Pretty.str o concat) - (if null matches - then ["nothing found", end_msg] - else ["found ", (string_of_int o length) matches, - " constants", end_msg, ":"]) - :: Pretty.str "" - :: map (pretty_const ctxt) matches - |> Pretty.chunks - |> Pretty.writeln - end handle ERROR s => Output.error_msg s - -end; - diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/find_theorems.ML --- a/src/Pure/Isar/find_theorems.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,378 +0,0 @@ -(* Title: Pure/Isar/find_theorems.ML - Author: Rafal Kolanski and Gerwin Klein, NICTA - -Retrieve theorems from proof context. -*) - -signature FIND_THEOREMS = -sig - val limit: int ref - val tac_limit: int ref - - datatype 'term criterion = - Name of string | Intro | Elim | Dest | Solves | Simp of 'term | - Pattern of 'term - - val find_theorems: Proof.context -> thm option -> bool -> - (bool * string criterion) list -> (Facts.ref * thm) list - - val print_theorems: Proof.context -> thm option -> int option -> bool -> - (bool * string criterion) list -> unit -end; - -structure FindTheorems: FIND_THEOREMS = -struct - -(** search criteria **) - -datatype 'term criterion = - Name of string | Intro | Elim | Dest | Solves | Simp of 'term | - Pattern of 'term; - -fun read_criterion _ (Name name) = Name name - | read_criterion _ Intro = Intro - | read_criterion _ Elim = Elim - | read_criterion _ Dest = Dest - | read_criterion _ Solves = Solves - | read_criterion ctxt (Simp str) = Simp (ProofContext.read_term_pattern ctxt str) - | read_criterion ctxt (Pattern str) = Pattern (ProofContext.read_term_pattern ctxt str); - -fun pretty_criterion ctxt (b, c) = - let - fun prfx s = if b then s else "-" ^ s; - in - (case c of - Name name => Pretty.str (prfx "name: " ^ quote name) - | Intro => Pretty.str (prfx "intro") - | Elim => Pretty.str (prfx "elim") - | Dest => Pretty.str (prfx "dest") - | Solves => Pretty.str (prfx "solves") - | Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1, - Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))] - | Pattern pat => Pretty.enclose (prfx " \"") "\"" - [Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)]) - end; - -(** search criterion filters **) - -(*generated filters are to be of the form - input: (Facts.ref * thm) - output: (p:int, s:int) option, where - NONE indicates no match - p is the primary sorting criterion - (eg. number of assumptions in the theorem) - s is the secondary sorting criterion - (eg. size of the substitution for intro, elim and dest) - when applying a set of filters to a thm, fold results in: - (biggest p, sum of all s) - currently p and s only matter for intro, elim, dest and simp filters, - otherwise the default ordering is used. -*) - - -(* matching theorems *) - -fun is_nontrivial thy = Term.is_Const o Term.head_of o ObjectLogic.drop_judgment thy; - -(*extract terms from term_src, refine them to the parts that concern us, - if po try match them against obj else vice versa. - trivial matches are ignored. - returns: smallest substitution size*) -fun is_matching_thm (extract_terms, refine_term) ctxt po obj term_src = - let - val thy = ProofContext.theory_of ctxt; - - fun matches pat = - is_nontrivial thy pat andalso - Pattern.matches thy (if po then (pat, obj) else (obj, pat)); - - fun substsize pat = - let val (_, subst) = - Pattern.match thy (if po then (pat, obj) else (obj, pat)) (Vartab.empty, Vartab.empty) - in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end; - - fun bestmatch [] = NONE - | bestmatch xs = SOME (foldr1 Int.min xs); - - val match_thm = matches o refine_term; - in - map (substsize o refine_term) (filter match_thm (extract_terms term_src)) - |> bestmatch - end; - - -(* filter_name *) - -fun filter_name str_pat (thmref, _) = - if match_string str_pat (Facts.name_of_ref thmref) - then SOME (0, 0) else NONE; - -(* filter intro/elim/dest/solves rules *) - -fun filter_dest ctxt goal (_, thm) = - let - val extract_dest = - (fn thm => if Thm.no_prems thm then [] else [Thm.full_prop_of thm], - hd o Logic.strip_imp_prems); - val prems = Logic.prems_of_goal goal 1; - - fun try_subst prem = is_matching_thm extract_dest ctxt true prem thm; - val successful = prems |> map_filter try_subst; - in - (*if possible, keep best substitution (one with smallest size)*) - (*dest rules always have assumptions, so a dest with one - assumption is as good as an intro rule with none*) - if not (null successful) - then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE - end; - -fun filter_intro ctxt goal (_, thm) = - let - val extract_intro = (single o Thm.full_prop_of, Logic.strip_imp_concl); - val concl = Logic.concl_of_goal goal 1; - val ss = is_matching_thm extract_intro ctxt true concl thm; - in - if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE - end; - -fun filter_elim ctxt goal (_, thm) = - if not (Thm.no_prems thm) then - let - val rule = Thm.full_prop_of thm; - val prems = Logic.prems_of_goal goal 1; - val goal_concl = Logic.concl_of_goal goal 1; - val rule_mp = hd (Logic.strip_imp_prems rule); - val rule_concl = Logic.strip_imp_concl rule; - fun combine t1 t2 = Const ("*combine*", dummyT --> dummyT) $ (t1 $ t2); - val rule_tree = combine rule_mp rule_concl; - fun goal_tree prem = combine prem goal_concl; - fun try_subst prem = - is_matching_thm (single, I) ctxt true (goal_tree prem) rule_tree; - val successful = prems |> map_filter try_subst; - in - (*elim rules always have assumptions, so an elim with one - assumption is as good as an intro rule with none*) - if is_nontrivial (ProofContext.theory_of ctxt) (Thm.major_prem_of thm) - andalso not (null successful) - then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE - end - else NONE - -val tac_limit = ref 5; - -fun filter_solves ctxt goal = let - val baregoal = Logic.get_goal (prop_of goal) 1; - - fun etacn thm i = Seq.take (!tac_limit) o etac thm i; - fun try_thm thm = if Thm.no_prems thm then rtac thm 1 goal - else (etacn thm THEN_ALL_NEW - (Goal.norm_hhf_tac THEN' - Method.assumption_tac ctxt)) 1 goal; - in - fn (_, thm) => if (is_some o Seq.pull o try_thm) thm - then SOME (Thm.nprems_of thm, 0) else NONE - end; - -(* filter_simp *) - -fun filter_simp ctxt t (_, thm) = - let - val (_, {mk_rews = {mk, ...}, ...}) = - Simplifier.rep_ss (Simplifier.local_simpset_of ctxt); - val extract_simp = - (map Thm.full_prop_of o mk, #1 o Logic.dest_equals o Logic.strip_imp_concl); - val ss = is_matching_thm extract_simp ctxt false t thm - in - if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE - end; - - -(* filter_pattern *) - -fun get_names t = (Term.add_const_names t []) union (Term.add_free_names t []); -fun get_thm_names (_, thm) = get_names (Thm.full_prop_of thm); - (* Including all constants and frees is only sound because - matching uses higher-order patterns. If full matching - were used, then constants that may be subject to - beta-reduction after substitution of frees should - not be included for LHS set because they could be - thrown away by the substituted function. - e.g. for (?F 1 2) do not include 1 or 2, if it were - possible for ?F to be (% x y. 3) - The largest possible set should always be included on - the RHS. *) - -fun filter_pattern ctxt pat = let - val pat_consts = get_names pat; - - fun check (t, NONE) = check (t, SOME (get_thm_names t)) - | check ((_, thm), c as SOME thm_consts) = - (if pat_consts subset_string thm_consts - andalso (Pattern.matches_subterm (ProofContext.theory_of ctxt) - (pat, Thm.full_prop_of thm)) - then SOME (0, 0) else NONE, c); - in check end; - -(* interpret criteria as filters *) - -local - -fun err_no_goal c = - error ("Current goal required for " ^ c ^ " search criterion"); - -val fix_goal = Thm.prop_of; -val fix_goalo = Option.map fix_goal; - -fun filter_crit _ _ (Name name) = apfst (filter_name name) - | filter_crit _ NONE Intro = err_no_goal "intro" - | filter_crit _ NONE Elim = err_no_goal "elim" - | filter_crit _ NONE Dest = err_no_goal "dest" - | filter_crit _ NONE Solves = err_no_goal "solves" - | filter_crit ctxt (SOME goal) Intro = apfst (filter_intro ctxt - (fix_goal goal)) - | filter_crit ctxt (SOME goal) Elim = apfst (filter_elim ctxt - (fix_goal goal)) - | filter_crit ctxt (SOME goal) Dest = apfst (filter_dest ctxt - (fix_goal goal)) - | filter_crit ctxt (SOME goal) Solves = apfst (filter_solves ctxt goal) - | filter_crit ctxt _ (Simp pat) = apfst (filter_simp ctxt pat) - | filter_crit ctxt _ (Pattern pat) = filter_pattern ctxt pat; - -fun opt_not x = if is_some x then NONE else SOME (0, 0); - -fun opt_add (SOME (a, x)) (SOME (b, y)) = SOME (Int.max (a, b), x + y : int) - | opt_add _ _ = NONE; - -fun app_filters thm = let - fun app (NONE, _, _) = NONE - | app (SOME v, consts, []) = SOME (v, thm) - | app (r, consts, f::fs) = let val (r', consts') = f (thm, consts) - in app (opt_add r r', consts', fs) end; - in app end; - -in - -fun filter_criterion ctxt opt_goal (b, c) = - (if b then I else (apfst opt_not)) o filter_crit ctxt opt_goal c; - -fun all_filters filters thms = - let - fun eval_filters thm = app_filters thm (SOME (0, 0), NONE, filters); - - (*filters return: (number of assumptions, substitution size) option, so - sort (desc. in both cases) according to number of assumptions first, - then by the substitution size*) - fun thm_ord (((p0, s0), _), ((p1, s1), _)) = - prod_ord int_ord int_ord ((p1, s1), (p0, s0)); - in map_filter eval_filters thms |> sort thm_ord |> map #2 end; - -end; - - -(* removing duplicates, preferring nicer names, roughly n log n *) - -local - -val index_ord = option_ord (K EQUAL); -val hidden_ord = bool_ord o pairself NameSpace.is_hidden; -val qual_ord = int_ord o pairself (length o NameSpace.explode); -val txt_ord = int_ord o pairself size; - -fun nicer_name (x, i) (y, j) = - (case hidden_ord (x, y) of EQUAL => - (case index_ord (i, j) of EQUAL => - (case qual_ord (x, y) of EQUAL => txt_ord (x, y) | ord => ord) - | ord => ord) - | ord => ord) <> GREATER; - -fun rem_cdups nicer xs = - let - fun rem_c rev_seen [] = rev rev_seen - | rem_c rev_seen [x] = rem_c (x :: rev_seen) [] - | rem_c rev_seen ((x as ((n, t), _)) :: (y as ((n', t'), _)) :: xs) = - if Thm.eq_thm_prop (t, t') - then rem_c rev_seen ((if nicer n n' then x else y) :: xs) - else rem_c (x :: rev_seen) (y :: xs) - in rem_c [] xs end; - -in - -fun nicer_shortest ctxt = let - val ns = ProofContext.theory_of ctxt - |> PureThy.facts_of - |> Facts.space_of; - - val len_sort = sort (int_ord o (pairself size)); - fun shorten s = (case len_sort (NameSpace.get_accesses ns s) of - [] => s - | s'::_ => s'); - - fun nicer (Facts.Named ((x, _), i)) (Facts.Named ((y, _), j)) = - nicer_name (shorten x, i) (shorten y, j) - | nicer (Facts.Fact _) (Facts.Named _) = true - | nicer (Facts.Named _) (Facts.Fact _) = false; - in nicer end; - -fun rem_thm_dups nicer xs = - xs ~~ (1 upto length xs) - |> sort (TermOrd.fast_term_ord o pairself (Thm.prop_of o #2 o #1)) - |> rem_cdups nicer - |> sort (int_ord o pairself #2) - |> map #1; - -end; - - -(* print_theorems *) - -fun all_facts_of ctxt = - maps Facts.selections - (Facts.dest_static [] (PureThy.facts_of (ProofContext.theory_of ctxt)) @ - Facts.dest_static [] (ProofContext.facts_of ctxt)); - -val limit = ref 40; - -fun find_theorems ctxt opt_goal rem_dups raw_criteria = - let - val add_prems = Seq.hd o (TRY (Method.insert_tac - (Assumption.prems_of ctxt) 1)); - val opt_goal' = Option.map add_prems opt_goal; - - val criteria = map (apsnd (read_criterion ctxt)) raw_criteria; - val filters = map (filter_criterion ctxt opt_goal') criteria; - - val raw_matches = all_filters filters (all_facts_of ctxt); - - val matches = - if rem_dups - then rem_thm_dups (nicer_shortest ctxt) raw_matches - else raw_matches; - in matches end; - -fun print_theorems ctxt opt_goal opt_limit rem_dups raw_criteria = let - val start = start_timing (); - - val criteria = map (apsnd (read_criterion ctxt)) raw_criteria; - val matches = find_theorems ctxt opt_goal rem_dups raw_criteria; - - val len = length matches; - val lim = the_default (! limit) opt_limit; - val thms = Library.drop (len - lim, matches); - - val end_msg = " in " ^ - (List.nth (String.tokens Char.isSpace (end_timing start), 3)) - ^ " secs" - in - Pretty.big_list "searched for:" (map (pretty_criterion ctxt) criteria) - :: Pretty.str "" :: - (if null thms then [Pretty.str ("nothing found" ^ end_msg)] - else - [Pretty.str ("found " ^ string_of_int len ^ " theorems" ^ - (if len <= lim then "" - else " (" ^ string_of_int lim ^ " displayed)") - ^ end_msg ^ ":"), Pretty.str ""] @ - map Display.pretty_fact thms) - |> Pretty.chunks |> Pretty.writeln - end - -end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/isar.ML --- a/src/Pure/Isar/isar.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,379 +0,0 @@ -(* Title: Pure/Isar/isar.ML - Author: Makarius - -The global Isabelle/Isar state and main read-eval-print loop. -*) - -signature ISAR = -sig - val init: unit -> unit - val exn: unit -> (exn * string) option - val state: unit -> Toplevel.state - val context: unit -> Proof.context - val goal: unit -> thm - val print: unit -> unit - val >> : Toplevel.transition -> bool - val >>> : Toplevel.transition list -> unit - val linear_undo: int -> unit - val undo: int -> unit - val kill: unit -> unit - val kill_proof: unit -> unit - val crashes: exn list ref - val toplevel_loop: {init: bool, welcome: bool, sync: bool, secure: bool} -> unit - val loop: unit -> unit - val main: unit -> unit - - type id = string - val no_id: id - val create_command: Toplevel.transition -> id - val insert_command: id -> id -> unit - val remove_command: id -> unit -end; - -structure Isar: ISAR = -struct - - -(** TTY model -- SINGLE-THREADED! **) - -(* the global state *) - -type history = (Toplevel.state * Toplevel.transition) list; - (*previous state, state transition -- regular commands only*) - -local - val global_history = ref ([]: history); - val global_state = ref Toplevel.toplevel; - val global_exn = ref (NONE: (exn * string) option); -in - -fun edit_history count f = NAMED_CRITICAL "Isar" (fn () => - let - fun edit 0 (st, hist) = (global_history := hist; global_state := st; global_exn := NONE) - | edit n (st, hist) = edit (n - 1) (f st hist); - in edit count (! global_state, ! global_history) end); - -fun state () = NAMED_CRITICAL "Isar" (fn () => ! global_state); -fun set_state state = NAMED_CRITICAL "Isar" (fn () => global_state := state); - -fun exn () = NAMED_CRITICAL "Isar" (fn () => ! global_exn); -fun set_exn exn = NAMED_CRITICAL "Isar" (fn () => global_exn := exn); - -end; - - -fun init () = edit_history 1 (K (K (Toplevel.toplevel, []))); - -fun context () = Toplevel.context_of (state ()) - handle Toplevel.UNDEF => error "Unknown context"; - -fun goal () = #2 (#2 (Proof.get_goal (Toplevel.proof_of (state ())))) - handle Toplevel.UNDEF => error "No goal present"; - -fun print () = Toplevel.print_state false (state ()); - - -(* history navigation *) - -local - -fun find_and_undo _ [] = error "Undo history exhausted" - | find_and_undo which ((prev, tr) :: hist) = - ((case Toplevel.init_of tr of SOME name => ThyInfo.kill_thy name | NONE => ()); - if which (Toplevel.name_of tr) then (prev, hist) else find_and_undo which hist); - -in - -fun linear_undo n = edit_history n (K (find_and_undo (K true))); - -fun undo n = edit_history n (fn st => fn hist => - find_and_undo (if Toplevel.is_proof st then K true else OuterKeyword.is_theory) hist); - -fun kill () = edit_history 1 (fn st => fn hist => - find_and_undo - (if Toplevel.is_proof st then OuterKeyword.is_theory else OuterKeyword.is_theory_begin) hist); - -fun kill_proof () = edit_history 1 (fn st => fn hist => - if Toplevel.is_proof st then find_and_undo OuterKeyword.is_theory hist - else raise Toplevel.UNDEF); - -end; - - -(* interactive state transformations *) - -fun op >> tr = - (case Toplevel.transition true tr (state ()) of - NONE => false - | SOME (_, SOME err) => (set_exn (SOME err); Toplevel.error_msg tr err; true) - | SOME (st', NONE) => - let - val name = Toplevel.name_of tr; - val _ = if OuterKeyword.is_theory_begin name then init () else (); - val _ = - if OuterKeyword.is_regular name - then edit_history 1 (fn st => fn hist => (st', (st, tr) :: hist)) else (); - in true end); - -fun op >>> [] = () - | op >>> (tr :: trs) = if op >> tr then op >>> trs else (); - - -(* toplevel loop *) - -val crashes = ref ([]: exn list); - -local - -fun raw_loop secure src = - let - fun check_secure () = - (if secure then warning "Secure loop -- cannot exit to ML" else (); secure); - in - (case Source.get_single (Source.set_prompt Source.default_prompt src) of - NONE => if secure then quit () else () - | SOME (tr, src') => if op >> tr orelse check_secure () then raw_loop secure src' else ()) - handle exn => - (Output.error_msg (Toplevel.exn_message exn) - handle crash => - (CRITICAL (fn () => change crashes (cons crash)); - warning "Recovering from Isar toplevel crash -- see also Isar.crashes"); - raw_loop secure src) - end; - -in - -fun toplevel_loop {init = do_init, welcome, sync, secure} = - (Context.set_thread_data NONE; - if do_init then init () else (); (* FIXME init editor model *) - if welcome then writeln (Session.welcome ()) else (); - uninterruptible (fn _ => fn () => raw_loop secure (OuterSyntax.isar sync)) ()); - -end; - -fun loop () = - toplevel_loop {init = false, welcome = false, sync = false, secure = Secure.is_secure ()}; - -fun main () = - toplevel_loop {init = true, welcome = true, sync = false, secure = Secure.is_secure ()}; - - - -(** individual toplevel commands **) - -(* unique identification *) - -type id = string; -val no_id : id = ""; - - -(* command category *) - -datatype category = Empty | Theory | Proof | Diag | Control; - -fun category_of tr = - let val name = Toplevel.name_of tr in - if name = "" then Empty - else if OuterKeyword.is_theory name then Theory - else if OuterKeyword.is_proof name then Proof - else if OuterKeyword.is_diag name then Diag - else Control - end; - -val is_theory = fn Theory => true | _ => false; -val is_proper = fn Theory => true | Proof => true | _ => false; -val is_regular = fn Control => false | _ => true; - - -(* command status *) - -datatype status = - Unprocessed | - Running | - Failed of exn * string | - Finished of Toplevel.state; - -fun status_markup Unprocessed = Markup.unprocessed - | status_markup Running = (Markup.runningN, []) - | status_markup (Failed _) = Markup.failed - | status_markup (Finished _) = Markup.finished; - -fun run int tr state = - (case Toplevel.transition int tr state of - NONE => NONE - | SOME (_, SOME err) => (Toplevel.error_msg tr err; SOME (Failed err)) - | SOME (state', NONE) => SOME (Finished state')); - - -(* datatype command *) - -datatype command = Command of - {category: category, - transition: Toplevel.transition, - status: status}; - -fun make_command (category, transition, status) = - Command {category = category, transition = transition, status = status}; - -val empty_command = - make_command (Empty, Toplevel.empty, Finished Toplevel.toplevel); - -fun map_command f (Command {category, transition, status}) = - make_command (f (category, transition, status)); - -fun map_status f = map_command (fn (category, transition, status) => - (category, transition, f status)); - - -(* global collection of identified commands *) - -fun err_dup id = sys_error ("Duplicate command " ^ quote id); -fun err_undef id = sys_error ("Unknown command " ^ quote id); - -local val global_commands = ref (Graph.empty: command Graph.T) in - -fun change_commands f = NAMED_CRITICAL "Isar" (fn () => change global_commands f) - handle Graph.DUP bad => err_dup bad | Graph.UNDEF bad => err_undef bad; - -fun get_commands () = NAMED_CRITICAL "Isar" (fn () => ! global_commands); - -end; - -fun add_edge (id1, id2) = - if id1 = no_id orelse id2 = no_id then I else Graph.add_edge (id1, id2); - - -fun init_commands () = change_commands (K Graph.empty); - -fun the_command id = - let val Command cmd = - if id = no_id then empty_command - else (Graph.get_node (get_commands ()) id handle Graph.UNDEF bad => err_undef bad) - in cmd end; - -fun prev_command id = - if id = no_id then no_id - else - (case Graph.imm_preds (get_commands ()) id handle Graph.UNDEF bad => err_undef bad of - [] => no_id - | [prev] => prev - | _ => sys_error ("Non-linear command dependency " ^ quote id)); - -fun next_commands id = - if id = no_id then [] - else Graph.imm_succs (get_commands ()) id handle Graph.UNDEF bad => err_undef bad; - -fun descendant_commands ids = - Graph.all_succs (get_commands ()) (distinct (op =) (filter_out (fn id => id = no_id) ids)) - handle Graph.UNDEF bad => err_undef bad; - - -(* maintain status *) - -fun report_status markup id = Toplevel.status (#transition (the_command id)) markup; - -fun update_status status id = change_commands (Graph.map_node id (map_status (K status))); - -fun report_update_status status id = - change_commands (Graph.map_node id (map_status (fn old_status => - let val markup = status_markup status - in if markup <> status_markup old_status then report_status markup id else (); status end))); - - -(* create and dispose commands *) - -fun create_command raw_tr = - let - val (id, tr) = - (case Toplevel.get_id raw_tr of - SOME id => (id, raw_tr) - | NONE => - let val id = - if ! Toplevel.debug then "isabelle:" ^ Toplevel.name_of raw_tr ^ serial_string () - else "isabelle:" ^ serial_string () - in (id, Toplevel.put_id id raw_tr) end); - - val cmd = make_command (category_of tr, tr, Unprocessed); - val _ = change_commands (Graph.new_node (id, cmd)); - in id end; - -fun dispose_commands ids = - let - val desc = descendant_commands ids; - val _ = List.app (report_status Markup.disposed) desc; - val _ = change_commands (Graph.del_nodes desc); - in () end; - - -(* final state *) - -fun the_state id = - (case the_command id of - {status = Finished state, ...} => state - | {transition, ...} => error ("Unfinished command " ^ Toplevel.str_of transition)); - - - -(** editor model **) - -(* run commands *) - -fun try_run id = - (case try the_state (prev_command id) of - NONE => () - | SOME state => - (case run true (#transition (the_command id)) state of - NONE => () - | SOME status => report_update_status status id)); - -fun rerun_commands ids = - (List.app (report_update_status Unprocessed) ids; List.app try_run ids); - - -(* modify document *) - -fun insert_command prev id = NAMED_CRITICAL "Isar" (fn () => - let - val nexts = next_commands prev; - val _ = change_commands - (fold (fn next => Graph.del_edge (prev, next)) nexts #> add_edge (prev, id) #> - fold (fn next => Graph.add_edge (id, next)) nexts); - in descendant_commands [id] end) |> rerun_commands; - -fun remove_command id = NAMED_CRITICAL "Isar" (fn () => - let - val prev = prev_command id; - val nexts = next_commands id; - val _ = change_commands - (fold (fn next => Graph.del_edge (id, next)) nexts #> - fold (fn next => add_edge (prev, next)) nexts); - in descendant_commands nexts end) |> rerun_commands; - - -(* concrete syntax *) - -local - -structure P = OuterParse; -val op >> = Scan.>>; - -in - -val _ = - OuterSyntax.internal_command "Isar.command" - (P.string -- P.string >> (fn (id, text) => - Toplevel.imperative (fn () => - ignore (create_command (OuterSyntax.prepare_command (Position.id id) text))))); - -val _ = - OuterSyntax.internal_command "Isar.insert" - (P.string -- P.string >> (fn (prev, id) => - Toplevel.imperative (fn () => insert_command prev id))); - -val _ = - OuterSyntax.internal_command "Isar.remove" - (P.string >> (fn id => Toplevel.imperative (fn () => remove_command id))); - -end; - -end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/isar_cmd.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/isar_syn.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/local_defs.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/method.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/obtain.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/proof_context.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/session.ML --- a/src/Pure/Isar/session.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,105 +0,0 @@ -(* Title: Pure/Isar/session.ML - Author: Markus Wenzel, TU Muenchen - -Session management -- maintain state of logic images. -*) - -signature SESSION = -sig - val id: unit -> string list - val name: unit -> string - val welcome: unit -> string - val use_dir: string -> bool -> string list -> bool -> bool -> string -> bool -> string list -> - string -> string -> bool * string -> string -> int -> bool -> int -> int -> bool -> unit - val finish: unit -> unit -end; - -structure Session: SESSION = -struct - - -(* session state *) - -val session = ref ([Context.PureN]: string list); -val session_path = ref ([]: string list); -val session_finished = ref false; -val remote_path = ref (NONE: Url.T option); - - -(* access path *) - -fun id () = ! session; -fun path () = ! session_path; - -fun str_of [] = Context.PureN - | str_of elems = space_implode "/" elems; - -fun name () = "Isabelle/" ^ str_of (path ()); - -fun welcome () = - if Distribution.is_official then - "Welcome to " ^ name () ^ " (" ^ Distribution.version ^ ")" - else - "Unofficial version of " ^ name () ^ " (" ^ Distribution.version ^ ")" ^ - (if Distribution.changelog <> "" then "\nSee also " ^ Distribution.changelog else ""); - - -(* add_path *) - -fun add_path reset s = - let val sess = ! session @ [s] in - (case duplicates (op =) sess of - [] => (session := sess; session_path := ((if reset then [] else ! session_path) @ [s])) - | dups => error ("Duplicate session identifiers " ^ commas_quote dups ^ " in " ^ str_of sess)) - end; - - -(* init *) - -fun init reset parent name = - if not (member (op =) (! session) parent) orelse not (! session_finished) then - error ("Unfinished parent session " ^ quote parent ^ " for " ^ quote name) - else (add_path reset name; session_finished := false); - - -(* finish *) - -fun finish () = - (Output.accumulated_time (); - ThyInfo.finish (); - Present.finish (); - Future.shutdown (); - session_finished := true); - - -(* use_dir *) - -fun get_rpath rpath = - (if rpath = "" then () else - if is_some (! remote_path) then - error "Path for remote theory browsing information may only be set once" - else - remote_path := SOME (Url.explode rpath); - (! remote_path, rpath <> "")); - -fun dumping (_, "") = NONE - | dumping (cp, path) = SOME (cp, Path.explode path); - -fun use_dir root build modes reset info doc doc_graph doc_versions - parent name dump rpath level verbose max_threads trace_threads parallel_proofs = - ((fn () => - (init reset parent name; - Present.init build info doc doc_graph doc_versions (path ()) name - (dumping dump) (get_rpath rpath) verbose (map ThyInfo.get_theory (ThyInfo.get_names ())); - use root; - finish ())) - |> setmp_noncritical Proofterm.proofs level - |> setmp_noncritical print_mode (modes @ print_mode_value ()) - |> setmp_noncritical Goal.parallel_proofs parallel_proofs - |> setmp_noncritical Multithreading.trace trace_threads - |> setmp_noncritical Multithreading.max_threads - (if Multithreading.available then max_threads - else (if max_threads = 1 then () else warning "Multithreading support unavailable"; 1))) () - handle exn => (Output.error_msg (Toplevel.exn_message exn); exit 1); - -end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Isar/theory_target.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/ML-Systems/alice.ML --- a/src/Pure/ML-Systems/alice.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,220 +0,0 @@ -(* Title: Pure/ML-Systems/alice.ML - -Compatibility file for Alice 1.4. - -NOTE: there is no wrapper script; may run it interactively as follows: - -$ cd Isabelle/src/Pure -$ env ALICE_JIT_MODE=0 ISABELLE_HOME=$(cd ../..; pwd) alice -- val ml_system = "alice"; -- use "ML-Systems/exn.ML"; -- use "ML-Systems/universal.ML"; -- use "ML-Systems/multithreading.ML"; -- use "ML-Systems/time_limit.ML"; -- use "ML-Systems/alice.ML"; -- use "ROOT.ML"; -- Session.finish (); -*) - -val ml_system_fix_ints = false; - -fun forget_structure _ = (); - -fun exit 0 = (OS.Process.exit OS.Process.success): unit - | exit _ = OS.Process.exit OS.Process.failure; - - -(** ML system related **) - -(*low-level pointer equality*) -fun pointer_eq (_: 'a, _: 'a) = false; - - -(* integer compatibility -- downgraded IntInf *) - -structure Time = -struct - open Time; - val fromMilliseconds = Time.fromMilliseconds o IntInf.fromInt; - val fromSeconds = Time.fromSeconds o IntInf.fromInt; -end; - -structure IntInf = -struct - fun divMod (x, y) = (x div y, x mod y); - open Int; -end; - - -(* restore old-style character / string functions *) - -exception Ord; -fun ord "" = raise Ord - | ord s = Char.ord (String.sub (s, 0)); - -val chr = String.str o chr; -val explode = map String.str o String.explode; -val implode = String.concat; - - -(* Poly/ML emulation *) - -fun quit () = exit 0; - -fun get_print_depth () = ! Print.depth; -fun print_depth n = Print.depth := n; - - -(* compiler-independent timing functions *) - -structure Timer = -struct - open Timer; - type cpu_timer = unit; - fun startCPUTimer () = (); - fun checkCPUTimer () = {sys = Time.zeroTime, usr = Time.zeroTime}; - fun checkGCTime () = Time.zeroTime; -end; - -fun start_timing () = - let val CPUtimer = Timer.startCPUTimer(); - val time = Timer.checkCPUTimer(CPUtimer) - in (CPUtimer,time) end; - -fun end_timing (CPUtimer, {sys,usr}) = - let open Time (*...for Time.toString, Time.+ and Time.- *) - val {sys=sys2,usr=usr2} = Timer.checkCPUTimer(CPUtimer) - in "User " ^ toString (usr2-usr) ^ - " All "^ toString (sys2-sys + usr2-usr) ^ - " secs" - handle Time => "" - end; - -fun check_timer timer = - let - val {sys, usr} = Timer.checkCPUTimer timer; - val gc = Timer.checkGCTime timer; (* FIXME already included in usr? *) - in (sys, usr, gc) end; - - -(*prompts*) -fun ml_prompts p1 p2 = (); - -(*dummy implementation*) -fun profile (n: int) f x = f x; - -(*dummy implementation*) -fun exception_trace f = f (); - -(*dummy implementation*) -fun print x = x; - - -(* toplevel pretty printing (see also Pure/pure_setup.ML) *) - -fun make_pp path pprint = (path, pprint); -fun install_pp (path, pp) = (); - - -(* ML command execution *) - -fun use_text _ _ _ _ _ txt = (Compiler.eval txt; ()); -fun use_file _ _ _ _ name = use name; - - - -(** interrupts **) - -exception Interrupt; - -fun interruptible f x = f x; -fun uninterruptible f x = f (fn (g: 'c -> 'd) => g) x; - - -(* basis library fixes *) - -structure TextIO = -struct - open TextIO; - fun inputLine is = TextIO.inputLine is - handle IO.Io _ => raise Interrupt; -end; - - - -(** OS related **) - -structure OS = -struct - open OS; - structure FileSys = - struct - open FileSys; - fun tmpName () = - let val name = FileSys.tmpName () in - if String.isSuffix "\000" name - then String.substring (name, 0, size name - 1) - else name - end; - end; -end; - -val cd = OS.FileSys.chDir; -val pwd = OS.FileSys.getDir; - -local - -fun read_file name = - let val is = TextIO.openIn name - in Exn.release (Exn.capture TextIO.inputAll is before TextIO.closeIn is) end; - -fun write_file name txt = - let val os = TextIO.openOut name - in Exn.release (Exn.capture TextIO.output (os, txt) before TextIO.closeOut os) end; - -in - -fun system_out script = - let - val script_name = OS.FileSys.tmpName (); - val _ = write_file script_name script; - - val output_name = OS.FileSys.tmpName (); - - val status = - OS.Process.system ("perl -w \"$ISABELLE_HOME/lib/scripts/system.pl\" nogroup " ^ - script_name ^ " /dev/null " ^ output_name); - val rc = if OS.Process.isSuccess status then 0 else 1; - - val output = read_file output_name handle IO.Io _ => ""; - val _ = OS.FileSys.remove script_name handle OS.SysErr _ => (); - val _ = OS.FileSys.remove output_name handle OS.SysErr _ => (); - in (output, rc) end; - -end; - -structure OS = -struct - open OS; - structure FileSys = - struct - fun fileId name = - (case system_out ("perl -e '@_ = stat(q:" ^ name ^ ":); print $_[1]'") of - ("", _) => raise Fail "OS.FileSys.fileId" (* FIXME IO.Io!? *) - | (s, _) => (case Int.fromString s of NONE => raise Fail "OS.FileSys.fileId" | SOME i => i)); - val compare = Int.compare; - fun fullPath name = - (case system_out ("FILE='" ^ name ^ - "' && cd \"$(dirname \"$FILE\")\" && echo -n \"$(pwd -P)/$(basename \"$FILE\")\"") of - ("", _) => raise SysErr ("Bad file", NONE) - | (s, _) => s); - open FileSys; - end; -end; - -fun process_id () = raise Fail "process_id undefined"; - -fun getenv var = - (case OS.Process.getEnv var of - NONE => "" - | SOME txt => txt); diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Proof/reconstruct.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/System/isabelle_process.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/System/isabelle_process.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,138 @@ +(* Title: Pure/System/isabelle_process.ML + Author: Makarius + +Isabelle process wrapper -- interaction via external program. + +General format of process output: + + (1) unmarked stdout/stderr, no line structure (output should be + processed immediately as it arrives); + + (2) properly marked-up messages, e.g. for writeln channel + + "\002A" ^ props ^ "\002,\n" ^ text ^ "\002.\n" + + where the props consist of name=value lines terminated by "\002,\n" + each, and the remaining text is any number of lines (output is + supposed to be processed in one piece); + + (3) special init message holds "pid" and "session" property; + + (4) message content is encoded in YXML format. +*) + +signature ISABELLE_PROCESS = +sig + val isabelle_processN: string + val init: string -> unit +end; + +structure IsabelleProcess: ISABELLE_PROCESS = +struct + +(* print modes *) + +val isabelle_processN = "isabelle_process"; + +val _ = Output.add_mode isabelle_processN Output.default_output Output.default_escape; +val _ = Markup.add_mode isabelle_processN YXML.output_markup; + + +(* message markup *) + +fun special ch = Symbol.STX ^ ch; +val special_sep = special ","; +val special_end = special "."; + +local + +fun clean_string bad str = + if exists_string (member (op =) bad) str then + translate_string (fn c => if member (op =) bad c then Symbol.DEL else c) str + else str; + +fun message_props props = + let val clean = clean_string [Symbol.STX, "\n", "\r"] + in implode (map (fn (x, y) => clean x ^ "=" ^ clean y ^ special_sep ^ "\n") props) end; + +fun message_pos trees = trees |> get_first + (fn XML.Elem (name, atts, ts) => + if name = Markup.positionN then SOME (Position.of_properties atts) + else message_pos ts + | _ => NONE); + +fun output out_stream s = NAMED_CRITICAL "IO" (fn () => + (TextIO.output (out_stream, s); TextIO.output (out_stream, "\n"))); + +in + +fun message _ _ "" = () + | message out_stream ch body = + let + val pos = the_default Position.none (message_pos (YXML.parse_body body)); + val props = + Position.properties_of (Position.thread_data ()) + |> Position.default_properties pos; + val txt = clean_string [Symbol.STX] body; + in output out_stream (special ch ^ message_props props ^ txt ^ special_end) end; + +fun init_message out_stream = + let + val pid = (Markup.pidN, process_id ()); + val session = (Markup.sessionN, List.last (Session.id ()) handle List.Empty => "unknown"); + val text = Session.welcome (); + in output out_stream (special "A" ^ message_props [pid, session] ^ text ^ special_end) end; + +end; + + +(* channels *) + +local + +fun auto_flush stream = + let + val _ = TextIO.StreamIO.setBufferMode (TextIO.getOutstream stream, IO.BLOCK_BUF); + fun loop () = + (OS.Process.sleep (Time.fromMilliseconds 50); try TextIO.flushOut stream; loop ()); + in loop end; + +in + +fun setup_channels out = + let + val out_stream = + if out = "-" then TextIO.stdOut + else + let + val path = File.platform_path (Path.explode out); + val out_stream = TextIO.openOut path; (*fifo blocks until reader is ready*) + val _ = OS.FileSys.remove path; (*prevent alien access, indicate writer is ready*) + val _ = SimpleThread.fork false (auto_flush TextIO.stdOut); + in out_stream end; + val _ = SimpleThread.fork false (auto_flush out_stream); + val _ = SimpleThread.fork false (auto_flush TextIO.stdErr); + in + Output.status_fn := message out_stream "B"; + Output.writeln_fn := message out_stream "C"; + Output.priority_fn := message out_stream "D"; + Output.tracing_fn := message out_stream "E"; + Output.warning_fn := message out_stream "F"; + Output.error_fn := message out_stream "G"; + Output.debug_fn := message out_stream "H"; + Output.prompt_fn := ignore; + out_stream + end; + +end; + + +(* init *) + +fun init out = + (change print_mode (update (op =) isabelle_processN); + setup_channels out |> init_message; + OuterKeyword.report (); + Isar.toplevel_loop {init = true, welcome = false, sync = true, secure = true}); + +end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/System/isabelle_process.scala --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/System/isabelle_process.scala Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,435 @@ +/* Title: Pure/System/isabelle_process.ML + Author: Makarius + Options: :folding=explicit:collapseFolds=1: + +Isabelle process management -- always reactive due to multi-threaded I/O. +*/ + +package isabelle + +import java.util.concurrent.LinkedBlockingQueue +import java.io.{BufferedReader, BufferedWriter, InputStreamReader, OutputStreamWriter, + InputStream, OutputStream, IOException} + + +object IsabelleProcess { + + /* results */ + + object Kind extends Enumeration { + //{{{ values and codes + // internal system notification + val SYSTEM = Value("SYSTEM") + // Posix channels/events + val STDIN = Value("STDIN") + val STDOUT = Value("STDOUT") + val SIGNAL = Value("SIGNAL") + val EXIT = Value("EXIT") + // Isabelle messages + val INIT = Value("INIT") + val STATUS = Value("STATUS") + val WRITELN = Value("WRITELN") + val PRIORITY = Value("PRIORITY") + val TRACING = Value("TRACING") + val WARNING = Value("WARNING") + val ERROR = Value("ERROR") + val DEBUG = Value("DEBUG") + // messages codes + val code = Map( + ('A' : Int) -> Kind.INIT, + ('B' : Int) -> Kind.STATUS, + ('C' : Int) -> Kind.WRITELN, + ('D' : Int) -> Kind.PRIORITY, + ('E' : Int) -> Kind.TRACING, + ('F' : Int) -> Kind.WARNING, + ('G' : Int) -> Kind.ERROR, + ('H' : Int) -> Kind.DEBUG, + ('0' : Int) -> Kind.SYSTEM, + ('1' : Int) -> Kind.STDIN, + ('2' : Int) -> Kind.STDOUT, + ('3' : Int) -> Kind.SIGNAL, + ('4' : Int) -> Kind.EXIT) + // message markup + val markup = Map( + Kind.INIT -> Markup.INIT, + Kind.STATUS -> Markup.STATUS, + Kind.WRITELN -> Markup.WRITELN, + Kind.PRIORITY -> Markup.PRIORITY, + Kind.TRACING -> Markup.TRACING, + Kind.WARNING -> Markup.WARNING, + Kind.ERROR -> Markup.ERROR, + Kind.DEBUG -> Markup.DEBUG, + Kind.SYSTEM -> Markup.SYSTEM, + Kind.STDIN -> Markup.STDIN, + Kind.STDOUT -> Markup.STDOUT, + Kind.SIGNAL -> Markup.SIGNAL, + Kind.EXIT -> Markup.EXIT) + //}}} + def is_raw(kind: Value) = + kind == STDOUT + def is_control(kind: Value) = + kind == SYSTEM || + kind == SIGNAL || + kind == EXIT + def is_system(kind: Value) = + kind == SYSTEM || + kind == STDIN || + kind == SIGNAL || + kind == EXIT || + kind == STATUS + } + + class Result(val kind: Kind.Value, val props: List[(String, String)], val result: String) { + override def toString = { + val trees = YXML.parse_body_failsafe(result) + val res = + if (kind == Kind.STATUS) trees.map(_.toString).mkString + else trees.flatMap(XML.content(_).mkString).mkString + if (props.isEmpty) + kind.toString + " [[" + res + "]]" + else + kind.toString + " " + + (for ((x, y) <- props) yield x + "=" + y).mkString("{", ",", "}") + " [[" + res + "]]" + } + def is_raw = Kind.is_raw(kind) + def is_control = Kind.is_control(kind) + def is_system = Kind.is_system(kind) + } + + def parse_message(isabelle_system: IsabelleSystem, result: Result) = + { + XML.Elem(Markup.MESSAGE, (Markup.CLASS, Kind.markup(result.kind)) :: result.props, + YXML.parse_body_failsafe(isabelle_system.symbols.decode(result.result))) + } +} + + +class IsabelleProcess(isabelle_system: IsabelleSystem, + results: EventBus[IsabelleProcess.Result], args: String*) +{ + import IsabelleProcess._ + + + /* demo constructor */ + + def this(args: String*) = + this(new IsabelleSystem, new EventBus[IsabelleProcess.Result] + Console.println, args: _*) + + + /* process information */ + + @volatile private var proc: Process = null + @volatile private var closing = false + @volatile private var pid: String = null + @volatile private var the_session: String = null + def session = the_session + + + /* results */ + + def parse_message(result: Result): XML.Tree = + IsabelleProcess.parse_message(isabelle_system, result) + + private val result_queue = new LinkedBlockingQueue[Result] + + private def put_result(kind: Kind.Value, props: List[(String, String)], result: String) + { + if (kind == Kind.INIT) { + val map = Map(props: _*) + if (map.isDefinedAt(Markup.PID)) pid = map(Markup.PID) + if (map.isDefinedAt(Markup.SESSION)) the_session = map(Markup.SESSION) + } + result_queue.put(new Result(kind, props, result)) + } + + private class ResultThread extends Thread("isabelle: results") { + override def run() = { + var finished = false + while (!finished) { + val result = + try { result_queue.take } + catch { case _: NullPointerException => null } + + if (result != null) { + results.event(result) + if (result.kind == Kind.EXIT) finished = true + } + else finished = true + } + } + } + + + /* signals */ + + def interrupt() = synchronized { + if (proc == null) error("Cannot interrupt Isabelle: no process") + if (pid == null) put_result(Kind.SYSTEM, Nil, "Cannot interrupt: unknown pid") + else { + try { + if (isabelle_system.execute(true, "kill", "-INT", pid).waitFor == 0) + put_result(Kind.SIGNAL, Nil, "INT") + else + put_result(Kind.SYSTEM, Nil, "Cannot interrupt: kill command failed") + } + catch { case e: IOException => error("Cannot interrupt Isabelle: " + e.getMessage) } + } + } + + def kill() = synchronized { + if (proc == 0) error("Cannot kill Isabelle: no process") + else { + try_close() + Thread.sleep(500) + put_result(Kind.SIGNAL, Nil, "KILL") + proc.destroy + proc = null + pid = null + } + } + + + /* output being piped into the process */ + + private val output = new LinkedBlockingQueue[String] + + private def output_raw(text: String) = synchronized { + if (proc == null) error("Cannot output to Isabelle: no process") + if (closing) error("Cannot output to Isabelle: already closing") + output.put(text) + } + + def output_sync(text: String) = + output_raw(" \\<^sync>\n; " + text + " \\<^sync>;\n") + + + def command(text: String) = + output_sync("Isabelle.command " + IsabelleSyntax.encode_string(text)) + + def command(props: List[(String, String)], text: String) = + output_sync("Isabelle.command " + IsabelleSyntax.encode_properties(props) + " " + + IsabelleSyntax.encode_string(text)) + + def ML(text: String) = + output_sync("ML_val " + IsabelleSyntax.encode_string(text)) + + def close() = synchronized { // FIXME watchdog/timeout + output_raw("\u0000") + closing = true + } + + def try_close() = synchronized { + if (proc != null && !closing) { + try { close() } + catch { case _: RuntimeException => } + } + } + + + /* stdin */ + + private class StdinThread(out_stream: OutputStream) extends Thread("isabelle: stdin") { + override def run() = { + val writer = new BufferedWriter(new OutputStreamWriter(out_stream, isabelle_system.charset)) + var finished = false + while (!finished) { + try { + //{{{ + val s = output.take + if (s == "\u0000") { + writer.close + finished = true + } + else { + put_result(Kind.STDIN, Nil, s) + writer.write(s) + writer.flush + } + //}}} + } + catch { + case e: IOException => put_result(Kind.SYSTEM, Nil, "Stdin thread: " + e.getMessage) + } + } + put_result(Kind.SYSTEM, Nil, "Stdin thread terminated") + } + } + + + /* stdout */ + + private class StdoutThread(in_stream: InputStream) extends Thread("isabelle: stdout") { + override def run() = { + val reader = new BufferedReader(new InputStreamReader(in_stream, isabelle_system.charset)) + var result = new StringBuilder(100) + + var finished = false + while (!finished) { + try { + //{{{ + var c = -1 + var done = false + while (!done && (result.length == 0 || reader.ready)) { + c = reader.read + if (c >= 0) result.append(c.asInstanceOf[Char]) + else done = true + } + if (result.length > 0) { + put_result(Kind.STDOUT, Nil, result.toString) + result.length = 0 + } + else { + reader.close + finished = true + try_close() + } + //}}} + } + catch { + case e: IOException => put_result(Kind.SYSTEM, Nil, "Stdout thread: " + e.getMessage) + } + } + put_result(Kind.SYSTEM, Nil, "Stdout thread terminated") + } + } + + + /* messages */ + + private class MessageThread(fifo: String) extends Thread("isabelle: messages") { + override def run() = { + val reader = isabelle_system.fifo_reader(fifo) + var kind: Kind.Value = null + var props: List[(String, String)] = Nil + var result = new StringBuilder + + var finished = false + while (!finished) { + try { + if (kind == null) { + //{{{ Char mode -- resync + var c = -1 + do { + c = reader.read + if (c >= 0 && c != 2) result.append(c.asInstanceOf[Char]) + } while (c >= 0 && c != 2) + + if (result.length > 0) { + put_result(Kind.SYSTEM, Nil, "Malformed message:\n" + result.toString) + result.length = 0 + } + if (c < 0) { + reader.close + finished = true + try_close() + } + else { + c = reader.read + if (Kind.code.isDefinedAt(c)) kind = Kind.code(c) + else kind = null + } + //}}} + } + else { + //{{{ Line mode + val line = reader.readLine + if (line == null) { + reader.close + finished = true + try_close() + } + else { + val len = line.length + // property + if (line.endsWith("\u0002,")) { + val i = line.indexOf('=') + if (i > 0) { + val name = line.substring(0, i) + val value = line.substring(i + 1, len - 2) + props = (name, value) :: props + } + } + // last text line + else if (line.endsWith("\u0002.")) { + result.append(line.substring(0, len - 2)) + put_result(kind, props.reverse, result.toString) + kind = null + props = Nil + result.length = 0 + } + // text line + else { + result.append(line) + result.append('\n') + } + } + //}}} + } + } + catch { + case e: IOException => put_result(Kind.SYSTEM, Nil, "Message thread: " + e.getMessage) + } + } + put_result(Kind.SYSTEM, Nil, "Message thread terminated") + } + } + + + + /** main **/ + + { + /* isabelle version */ + + { + val (msg, rc) = isabelle_system.isabelle_tool("version") + if (rc != 0) error("Version check failed -- bad Isabelle installation:\n" + msg) + put_result(Kind.SYSTEM, Nil, msg) + } + + + /* messages */ + + val message_fifo = isabelle_system.mk_fifo() + def rm_fifo() = isabelle_system.rm_fifo(message_fifo) + + val message_thread = new MessageThread(message_fifo) + message_thread.start + + new ResultThread().start + + + /* exec process */ + + try { + val cmdline = + List(isabelle_system.getenv_strict("ISABELLE_PROCESS"), "-W", message_fifo) ++ args + proc = isabelle_system.execute(true, cmdline: _*) + } + catch { + case e: IOException => + rm_fifo() + error("Failed to execute Isabelle process: " + e.getMessage) + } + + + /* stdin/stdout */ + + new StdinThread(proc.getOutputStream).start + new StdoutThread(proc.getInputStream).start + + + /* exit */ + + new Thread("isabelle: exit") { + override def run() = { + val rc = proc.waitFor() + Thread.sleep(300) + put_result(Kind.SYSTEM, Nil, "Exit thread terminated") + put_result(Kind.EXIT, Nil, Integer.toString(rc)) + rm_fifo() + } + }.start + + } +} diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/System/isabelle_system.scala --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/System/isabelle_system.scala Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,158 @@ +/* Title: Pure/System/isabelle_system.scala + Author: Makarius + +Isabelle system support -- basic Cygwin/Posix compatibility. +*/ + +package isabelle + +import java.util.regex.{Pattern, Matcher} +import java.io.{BufferedReader, InputStreamReader, FileInputStream, File, IOException} + +import scala.io.Source + + +class IsabelleSystem { + + val charset = "UTF-8" + + + /* Isabelle environment settings */ + + private val environment = System.getenv + + def getenv(name: String) = { + val value = environment.get(if (name == "HOME") "HOME_JVM" else name) + if (value != null) value else "" + } + + def getenv_strict(name: String) = { + val value = environment.get(name) + if (value != "") value else error("Undefined environment variable: " + name) + } + + val is_cygwin = Pattern.matches(".*-cygwin", getenv_strict("ML_PLATFORM")) + + + /* file path specifications */ + + private val cygdrive_pattern = Pattern.compile("/cygdrive/([a-zA-Z])($|/.*)") + + def platform_path(source_path: String) = { + val result_path = new StringBuilder + + def init(path: String) = { + val cygdrive = cygdrive_pattern.matcher(path) + if (cygdrive.matches) { + result_path.length = 0 + result_path.append(cygdrive.group(1)) + result_path.append(":") + result_path.append(File.separator) + cygdrive.group(2) + } + else if (path.startsWith("/")) { + result_path.length = 0 + result_path.append(getenv_strict("ISABELLE_ROOT_JVM")) + path.substring(1) + } + else path + } + def append(path: String) = { + for (p <- init(path).split("/")) { + if (p != "") { + val len = result_path.length + if (len > 0 && result_path(len - 1) != File.separatorChar) + result_path.append(File.separator) + result_path.append(p) + } + } + } + for (p <- init(source_path).split("/")) { + if (p.startsWith("$")) append(getenv_strict(p.substring(1))) + else if (p == "~") append(getenv_strict("HOME")) + else if (p == "~~") append(getenv_strict("ISABELLE_HOME")) + else append(p) + } + result_path.toString + } + + def platform_file(path: String) = + new File(platform_path(path)) + + + /* processes */ + + def execute(redirect: Boolean, args: String*): Process = { + val cmdline = new java.util.LinkedList[String] + if (is_cygwin) cmdline.add(platform_path("/bin/env")) + for (s <- args) cmdline.add(s) + + val proc = new ProcessBuilder(cmdline) + proc.environment.clear + proc.environment.putAll(environment) + proc.redirectErrorStream(redirect) + proc.start + } + + + /* Isabelle tools (non-interactive) */ + + def isabelle_tool(args: String*) = { + val proc = + try { execute(true, (List(getenv_strict("ISABELLE_TOOL")) ++ args): _*) } + catch { case e: IOException => error(e.getMessage) } + proc.getOutputStream.close + val output = Source.fromInputStream(proc.getInputStream, charset).mkString + val rc = proc.waitFor + (output, rc) + } + + + /* named pipes */ + + def mk_fifo() = { + val (result, rc) = isabelle_tool("mkfifo") + if (rc == 0) result.trim + else error(result) + } + + def rm_fifo(fifo: String) = { + val (result, rc) = isabelle_tool("rmfifo", fifo) + if (rc != 0) error(result) + } + + def fifo_reader(fifo: String) = { + // blocks until writer is ready + val stream = + if (is_cygwin) execute(false, "cat", fifo).getInputStream + else new FileInputStream(fifo) + new BufferedReader(new InputStreamReader(stream, charset)) + } + + + /* find logics */ + + def find_logics() = { + val ml_ident = getenv_strict("ML_IDENTIFIER") + var logics: Set[String] = Set() + for (dir <- getenv_strict("ISABELLE_PATH").split(":")) { + val files = platform_file(dir + "/" + ml_ident).listFiles() + if (files != null) { + for (file <- files if file.isFile) logics += file.getName + } + } + logics.toList.sort(_ < _) + } + + + /* symbols */ + + private def read_symbols(path: String) = { + val file = new File(platform_path(path)) + if (file.canRead) Source.fromFile(file).getLines + else Iterator.empty + } + val symbols = new Symbol.Interpretation( + read_symbols("$ISABELLE_HOME/etc/symbols") ++ + read_symbols("$ISABELLE_HOME_USER/etc/symbols")) +} diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/System/isar.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/System/isar.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,415 @@ +(* Title: Pure/System/isar.ML + Author: Makarius + +The global Isabelle/Isar state and main read-eval-print loop. +*) + +signature ISAR = +sig + val init: unit -> unit + val exn: unit -> (exn * string) option + val state: unit -> Toplevel.state + val context: unit -> Proof.context + val goal: unit -> thm + val print: unit -> unit + val >> : Toplevel.transition -> bool + val >>> : Toplevel.transition list -> unit + val linear_undo: int -> unit + val undo: int -> unit + val kill: unit -> unit + val kill_proof: unit -> unit + val crashes: exn list ref + val toplevel_loop: {init: bool, welcome: bool, sync: bool, secure: bool} -> unit + val loop: unit -> unit + val main: unit -> unit + + type id = string + val no_id: id + val create_command: Toplevel.transition -> id + val insert_command: id -> id -> unit + val remove_command: id -> unit +end; + +structure Isar: ISAR = +struct + + +(** TTY model -- SINGLE-THREADED! **) + +(* the global state *) + +type history = (Toplevel.state * Toplevel.transition) list; + (*previous state, state transition -- regular commands only*) + +local + val global_history = ref ([]: history); + val global_state = ref Toplevel.toplevel; + val global_exn = ref (NONE: (exn * string) option); +in + +fun edit_history count f = NAMED_CRITICAL "Isar" (fn () => + let + fun edit 0 (st, hist) = (global_history := hist; global_state := st; global_exn := NONE) + | edit n (st, hist) = edit (n - 1) (f st hist); + in edit count (! global_state, ! global_history) end); + +fun state () = NAMED_CRITICAL "Isar" (fn () => ! global_state); +fun set_state state = NAMED_CRITICAL "Isar" (fn () => global_state := state); + +fun exn () = NAMED_CRITICAL "Isar" (fn () => ! global_exn); +fun set_exn exn = NAMED_CRITICAL "Isar" (fn () => global_exn := exn); + +end; + + +fun init () = edit_history 1 (K (K (Toplevel.toplevel, []))); + +fun context () = Toplevel.context_of (state ()) + handle Toplevel.UNDEF => error "Unknown context"; + +fun goal () = #2 (#2 (Proof.get_goal (Toplevel.proof_of (state ())))) + handle Toplevel.UNDEF => error "No goal present"; + +fun print () = Toplevel.print_state false (state ()); + + +(* history navigation *) + +local + +fun find_and_undo _ [] = error "Undo history exhausted" + | find_and_undo which ((prev, tr) :: hist) = + ((case Toplevel.init_of tr of SOME name => ThyInfo.kill_thy name | NONE => ()); + if which (Toplevel.name_of tr) then (prev, hist) else find_and_undo which hist); + +in + +fun linear_undo n = edit_history n (K (find_and_undo (K true))); + +fun undo n = edit_history n (fn st => fn hist => + find_and_undo (if Toplevel.is_proof st then K true else OuterKeyword.is_theory) hist); + +fun kill () = edit_history 1 (fn st => fn hist => + find_and_undo + (if Toplevel.is_proof st then OuterKeyword.is_theory else OuterKeyword.is_theory_begin) hist); + +fun kill_proof () = edit_history 1 (fn st => fn hist => + if Toplevel.is_proof st then find_and_undo OuterKeyword.is_theory hist + else raise Toplevel.UNDEF); + +end; + + +(* interactive state transformations *) + +fun op >> tr = + (case Toplevel.transition true tr (state ()) of + NONE => false + | SOME (_, SOME err) => (set_exn (SOME err); Toplevel.error_msg tr err; true) + | SOME (st', NONE) => + let + val name = Toplevel.name_of tr; + val _ = if OuterKeyword.is_theory_begin name then init () else (); + val _ = + if OuterKeyword.is_regular name + then edit_history 1 (fn st => fn hist => (st', (st, tr) :: hist)) else (); + in true end); + +fun op >>> [] = () + | op >>> (tr :: trs) = if op >> tr then op >>> trs else (); + + +(* toplevel loop *) + +val crashes = ref ([]: exn list); + +local + +fun raw_loop secure src = + let + fun check_secure () = + (if secure then warning "Secure loop -- cannot exit to ML" else (); secure); + in + (case Source.get_single (Source.set_prompt Source.default_prompt src) of + NONE => if secure then quit () else () + | SOME (tr, src') => if op >> tr orelse check_secure () then raw_loop secure src' else ()) + handle exn => + (Output.error_msg (Toplevel.exn_message exn) + handle crash => + (CRITICAL (fn () => change crashes (cons crash)); + warning "Recovering from Isar toplevel crash -- see also Isar.crashes"); + raw_loop secure src) + end; + +in + +fun toplevel_loop {init = do_init, welcome, sync, secure} = + (Context.set_thread_data NONE; + if do_init then init () else (); (* FIXME init editor model *) + if welcome then writeln (Session.welcome ()) else (); + uninterruptible (fn _ => fn () => raw_loop secure (OuterSyntax.isar sync)) ()); + +end; + +fun loop () = + toplevel_loop {init = false, welcome = false, sync = false, secure = Secure.is_secure ()}; + +fun main () = + toplevel_loop {init = true, welcome = true, sync = false, secure = Secure.is_secure ()}; + + + +(** individual toplevel commands **) + +(* unique identification *) + +type id = string; +val no_id : id = ""; + + +(* command category *) + +datatype category = Empty | Theory | Proof | Diag | Control; + +fun category_of tr = + let val name = Toplevel.name_of tr in + if name = "" then Empty + else if OuterKeyword.is_theory name then Theory + else if OuterKeyword.is_proof name then Proof + else if OuterKeyword.is_diag name then Diag + else Control + end; + +val is_theory = fn Theory => true | _ => false; +val is_proper = fn Theory => true | Proof => true | _ => false; +val is_regular = fn Control => false | _ => true; + + +(* command status *) + +datatype status = + Unprocessed | + Running | + Failed of exn * string | + Finished of Toplevel.state; + +fun status_markup Unprocessed = Markup.unprocessed + | status_markup Running = (Markup.runningN, []) + | status_markup (Failed _) = Markup.failed + | status_markup (Finished _) = Markup.finished; + +fun run int tr state = + (case Toplevel.transition int tr state of + NONE => NONE + | SOME (_, SOME err) => (Toplevel.error_msg tr err; SOME (Failed err)) + | SOME (state', NONE) => SOME (Finished state')); + + +(* datatype command *) + +datatype command = Command of + {category: category, + transition: Toplevel.transition, + status: status}; + +fun make_command (category, transition, status) = + Command {category = category, transition = transition, status = status}; + +val empty_command = + make_command (Empty, Toplevel.empty, Finished Toplevel.toplevel); + +fun map_command f (Command {category, transition, status}) = + make_command (f (category, transition, status)); + +fun map_status f = map_command (fn (category, transition, status) => + (category, transition, f status)); + + +(* global collection of identified commands *) + +fun err_dup id = sys_error ("Duplicate command " ^ quote id); +fun err_undef id = sys_error ("Unknown command " ^ quote id); + +local val global_commands = ref (Graph.empty: command Graph.T) in + +fun change_commands f = NAMED_CRITICAL "Isar" (fn () => change global_commands f) + handle Graph.DUP bad => err_dup bad | Graph.UNDEF bad => err_undef bad; + +fun get_commands () = NAMED_CRITICAL "Isar" (fn () => ! global_commands); + +end; + +fun add_edge (id1, id2) = + if id1 = no_id orelse id2 = no_id then I else Graph.add_edge (id1, id2); + + +fun init_commands () = change_commands (K Graph.empty); + +fun the_command id = + let val Command cmd = + if id = no_id then empty_command + else (Graph.get_node (get_commands ()) id handle Graph.UNDEF bad => err_undef bad) + in cmd end; + +fun prev_command id = + if id = no_id then no_id + else + (case Graph.imm_preds (get_commands ()) id handle Graph.UNDEF bad => err_undef bad of + [] => no_id + | [prev] => prev + | _ => sys_error ("Non-linear command dependency " ^ quote id)); + +fun next_commands id = + if id = no_id then [] + else Graph.imm_succs (get_commands ()) id handle Graph.UNDEF bad => err_undef bad; + +fun descendant_commands ids = + Graph.all_succs (get_commands ()) (distinct (op =) (filter_out (fn id => id = no_id) ids)) + handle Graph.UNDEF bad => err_undef bad; + + +(* maintain status *) + +fun report_status markup id = Toplevel.status (#transition (the_command id)) markup; + +fun update_status status id = change_commands (Graph.map_node id (map_status (K status))); + +fun report_update_status status id = + change_commands (Graph.map_node id (map_status (fn old_status => + let val markup = status_markup status + in if markup <> status_markup old_status then report_status markup id else (); status end))); + + +(* create and dispose commands *) + +fun create_command raw_tr = + let + val (id, tr) = + (case Toplevel.get_id raw_tr of + SOME id => (id, raw_tr) + | NONE => + let val id = + if ! Toplevel.debug then "isabelle:" ^ Toplevel.name_of raw_tr ^ serial_string () + else "isabelle:" ^ serial_string () + in (id, Toplevel.put_id id raw_tr) end); + + val cmd = make_command (category_of tr, tr, Unprocessed); + val _ = change_commands (Graph.new_node (id, cmd)); + in id end; + +fun dispose_commands ids = + let + val desc = descendant_commands ids; + val _ = List.app (report_status Markup.disposed) desc; + val _ = change_commands (Graph.del_nodes desc); + in () end; + + +(* final state *) + +fun the_state id = + (case the_command id of + {status = Finished state, ...} => state + | {transition, ...} => error ("Unfinished command " ^ Toplevel.str_of transition)); + + + +(** editor model **) + +(* run commands *) + +fun try_run id = + (case try the_state (prev_command id) of + NONE => () + | SOME state => + (case run true (#transition (the_command id)) state of + NONE => () + | SOME status => report_update_status status id)); + +fun rerun_commands ids = + (List.app (report_update_status Unprocessed) ids; List.app try_run ids); + + +(* modify document *) + +fun insert_command prev id = NAMED_CRITICAL "Isar" (fn () => + let + val nexts = next_commands prev; + val _ = change_commands + (fold (fn next => Graph.del_edge (prev, next)) nexts #> add_edge (prev, id) #> + fold (fn next => Graph.add_edge (id, next)) nexts); + in descendant_commands [id] end) |> rerun_commands; + +fun remove_command id = NAMED_CRITICAL "Isar" (fn () => + let + val prev = prev_command id; + val nexts = next_commands id; + val _ = change_commands + (fold (fn next => Graph.del_edge (id, next)) nexts #> + fold (fn next => add_edge (prev, next)) nexts); + in descendant_commands nexts end) |> rerun_commands; + + + +(** command syntax **) + +local + +structure P = OuterParse and K = OuterKeyword; +val op >> = Scan.>>; + +in + +(* global history *) + +val _ = + OuterSyntax.improper_command "init_toplevel" "init toplevel point-of-interest" K.control + (Scan.succeed (Toplevel.no_timing o Toplevel.imperative init)); + +val _ = + OuterSyntax.improper_command "linear_undo" "undo commands" K.control + (Scan.optional P.nat 1 >> + (fn n => Toplevel.no_timing o Toplevel.imperative (fn () => linear_undo n))); + +val _ = + OuterSyntax.improper_command "undo" "undo commands (skipping closed proofs)" K.control + (Scan.optional P.nat 1 >> + (fn n => Toplevel.no_timing o Toplevel.imperative (fn () => undo n))); + +val _ = + OuterSyntax.improper_command "undos_proof" "undo commands (skipping closed proofs)" K.control + (Scan.optional P.nat 1 >> (fn n => Toplevel.no_timing o + Toplevel.keep (fn state => + if Toplevel.is_proof state then (undo n; print ()) else raise Toplevel.UNDEF))); + +val _ = + OuterSyntax.improper_command "cannot_undo" "partial undo -- Proof General legacy" K.control + (P.name >> + (fn "end" => Toplevel.no_timing o Toplevel.imperative (fn () => undo 1) + | txt => Toplevel.imperative (fn () => error ("Cannot undo " ^ quote txt)))); + +val _ = + OuterSyntax.improper_command "kill" "kill partial proof or theory development" K.control + (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill)); + + +(* editor model *) + +val _ = + OuterSyntax.internal_command "Isar.command" + (P.string -- P.string >> (fn (id, text) => + Toplevel.imperative (fn () => + ignore (create_command (OuterSyntax.prepare_command (Position.id id) text))))); + +val _ = + OuterSyntax.internal_command "Isar.insert" + (P.string -- P.string >> (fn (prev, id) => + Toplevel.imperative (fn () => insert_command prev id))); + +val _ = + OuterSyntax.internal_command "Isar.remove" + (P.string >> (fn id => Toplevel.imperative (fn () => remove_command id))); + +end; + +end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/System/session.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/System/session.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,112 @@ +(* Title: Pure/System/session.ML + Author: Markus Wenzel, TU Muenchen + +Session management -- maintain state of logic images. +*) + +signature SESSION = +sig + val id: unit -> string list + val name: unit -> string + val welcome: unit -> string + val use_dir: string -> bool -> string list -> bool -> bool -> string -> bool -> string list -> + string -> string -> bool * string -> string -> int -> bool -> int -> int -> bool -> unit + val finish: unit -> unit +end; + +structure Session: SESSION = +struct + + +(* session state *) + +val session = ref ([Context.PureN]: string list); +val session_path = ref ([]: string list); +val session_finished = ref false; +val remote_path = ref (NONE: Url.T option); + + +(* access path *) + +fun id () = ! session; +fun path () = ! session_path; + +fun str_of [] = Context.PureN + | str_of elems = space_implode "/" elems; + +fun name () = "Isabelle/" ^ str_of (path ()); + + +(* welcome *) + +fun welcome () = + if Distribution.is_official then + "Welcome to " ^ name () ^ " (" ^ Distribution.version ^ ")" + else + "Unofficial version of " ^ name () ^ " (" ^ Distribution.version ^ ")" ^ + (if Distribution.changelog <> "" then "\nSee also " ^ Distribution.changelog else ""); + +val _ = + OuterSyntax.improper_command "welcome" "print welcome message" OuterKeyword.diag + (Scan.succeed (Toplevel.no_timing o Toplevel.imperative (writeln o welcome))); + + +(* add_path *) + +fun add_path reset s = + let val sess = ! session @ [s] in + (case duplicates (op =) sess of + [] => (session := sess; session_path := ((if reset then [] else ! session_path) @ [s])) + | dups => error ("Duplicate session identifiers " ^ commas_quote dups ^ " in " ^ str_of sess)) + end; + + +(* init *) + +fun init reset parent name = + if not (member (op =) (! session) parent) orelse not (! session_finished) then + error ("Unfinished parent session " ^ quote parent ^ " for " ^ quote name) + else (add_path reset name; session_finished := false); + + +(* finish *) + +fun finish () = + (Output.accumulated_time (); + ThyInfo.finish (); + Present.finish (); + Future.shutdown (); + session_finished := true); + + +(* use_dir *) + +fun get_rpath rpath = + (if rpath = "" then () else + if is_some (! remote_path) then + error "Path for remote theory browsing information may only be set once" + else + remote_path := SOME (Url.explode rpath); + (! remote_path, rpath <> "")); + +fun dumping (_, "") = NONE + | dumping (cp, path) = SOME (cp, Path.explode path); + +fun use_dir root build modes reset info doc doc_graph doc_versions + parent name dump rpath level verbose max_threads trace_threads parallel_proofs = + ((fn () => + (init reset parent name; + Present.init build info doc doc_graph doc_versions (path ()) name + (dumping dump) (get_rpath rpath) verbose (map ThyInfo.get_theory (ThyInfo.get_names ())); + use root; + finish ())) + |> setmp_noncritical Proofterm.proofs level + |> setmp_noncritical print_mode (modes @ print_mode_value ()) + |> setmp_noncritical Goal.parallel_proofs parallel_proofs + |> setmp_noncritical Multithreading.trace trace_threads + |> setmp_noncritical Multithreading.max_threads + (if Multithreading.available then max_threads + else (if max_threads = 1 then () else warning "Multithreading support unavailable"; 1))) () + handle exn => (Output.error_msg (Toplevel.exn_message exn); exit 1); + +end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Tools/ROOT.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Tools/find_consts.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Tools/find_consts.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,169 @@ +(* Title: Pure/Tools/find_consts.ML + Author: Timothy Bourke and Gerwin Klein, NICTA + +Hoogle-like (http://www-users.cs.york.ac.uk/~ndm/hoogle) searching by +type over constants, but matching is not fuzzy. +*) + +signature FIND_CONSTS = +sig + datatype criterion = + Strict of string + | Loose of string + | Name of string + + val find_consts : Proof.context -> (bool * criterion) list -> unit +end; + +structure FindConsts : FIND_CONSTS = +struct + +(* search criteria *) + +datatype criterion = + Strict of string + | Loose of string + | Name of string; + +(* matching types/consts *) + +fun add_tye (_, (_, t)) n = Term.size_of_typ t + n; + +fun matches_subtype thy typat = + let + val p = can (fn ty => Sign.typ_match thy (typat, ty) Vartab.empty); + + fun fs [] = false + | fs (t :: ts) = f t orelse fs ts + + and f (t as Type (_, ars)) = p t orelse fs ars + | f t = p t; + in f end; + +fun check_const p (nm, (ty, _)) = + if p (nm, ty) + then SOME (Term.size_of_typ ty) + else NONE; + +fun opt_not f (c as (_, (ty, _))) = + if is_some (f c) + then NONE else SOME (Term.size_of_typ ty); + +fun filter_const _ NONE = NONE + | filter_const f (SOME (c, r)) = + Option.map (pair c o (curry Int.min r)) (f c); + + +(* pretty results *) + +fun pretty_criterion (b, c) = + let + fun prfx s = if b then s else "-" ^ s; + in + (case c of + Strict pat => Pretty.str (prfx "strict: " ^ quote pat) + | Loose pat => Pretty.str (prfx (quote pat)) + | Name name => Pretty.str (prfx "name: " ^ quote name)) + end; + +fun pretty_const ctxt (nm, ty) = + let + val ty' = Logic.unvarifyT ty; + in + Pretty.block + [Pretty.quote (Pretty.str nm), Pretty.fbrk, + Pretty.str "::", Pretty.brk 1, + Pretty.quote (Syntax.pretty_typ ctxt ty')] + end; + +(* find_consts *) + +fun find_consts ctxt raw_criteria = + let + val start = start_timing (); + + val thy = ProofContext.theory_of ctxt; + val low_ranking = 10000; + + fun not_internal consts (nm, _) = + if member (op =) (Consts.the_tags consts nm) Markup.property_internal + then NONE else SOME low_ranking; + + fun make_pattern crit = + let + val raw_T = Syntax.parse_typ ctxt crit; + val t = Syntax.check_term + (ProofContext.set_mode ProofContext.mode_pattern ctxt) + (Term.dummy_pattern raw_T); + in Term.type_of t end; + + fun make_match (Strict arg) = + let val qty = make_pattern arg; in + fn (_, (ty, _)) => + let + val tye = Sign.typ_match thy (qty, ty) Vartab.empty; + val sub_size = Vartab.fold add_tye tye 0; + in SOME sub_size end handle MATCH => NONE + end + + | make_match (Loose arg) = + check_const (matches_subtype thy (make_pattern arg) o snd) + + | make_match (Name arg) = check_const (match_string arg o fst); + + fun make_criterion (b, crit) = (if b then I else opt_not) (make_match crit); + val criteria = map make_criterion raw_criteria; + + val consts = Sign.consts_of thy; + val (_, consts_tab) = (#constants o Consts.dest) consts; + fun eval_entry c = fold filter_const (not_internal consts::criteria) + (SOME (c, low_ranking)); + + val matches = + Symtab.fold (cons o eval_entry) consts_tab [] + |> map_filter I + |> sort (rev_order o int_ord o pairself snd) + |> map ((apsnd fst) o fst); + + val end_msg = " in " ^ Time.toString (#all (end_timing start)) ^ " secs"; + in + Pretty.big_list "searched for:" (map pretty_criterion raw_criteria) + :: Pretty.str "" + :: (Pretty.str o concat) + (if null matches + then ["nothing found", end_msg] + else ["found ", (string_of_int o length) matches, + " constants", end_msg, ":"]) + :: Pretty.str "" + :: map (pretty_const ctxt) matches + |> Pretty.chunks + |> Pretty.writeln + end; + + +(* command syntax *) + +fun find_consts_cmd spec = + Toplevel.unknown_theory o Toplevel.keep (fn state => + find_consts (Proof.context_of (Toplevel.enter_proof_body state)) spec); + +local + +structure P = OuterParse and K = OuterKeyword; + +val criterion = + P.reserved "strict" |-- P.!!! (P.$$$ ":" |-- P.xname) >> Strict || + P.reserved "name" |-- P.!!! (P.$$$ ":" |-- P.xname) >> Name || + P.xname >> Loose; + +in + +val _ = + OuterSyntax.improper_command "find_consts" "search constants by type pattern" K.diag + (Scan.repeat (((Scan.option P.minus >> is_none) -- criterion)) + >> (Toplevel.no_timing oo find_consts_cmd)); + +end; + +end; + diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Tools/find_theorems.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Tools/find_theorems.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,424 @@ +(* Title: Pure/Tools/find_theorems.ML + Author: Rafal Kolanski and Gerwin Klein, NICTA + +Retrieve theorems from proof context. +*) + +signature FIND_THEOREMS = +sig + datatype 'term criterion = + Name of string | Intro | Elim | Dest | Solves | Simp of 'term | + Pattern of 'term + val tac_limit: int ref + val limit: int ref + val find_theorems: Proof.context -> thm option -> bool -> + (bool * string criterion) list -> (Facts.ref * thm) list + val pretty_thm: Proof.context -> Facts.ref * thm -> Pretty.T + val print_theorems: Proof.context -> thm option -> int option -> bool -> + (bool * string criterion) list -> unit +end; + +structure FindTheorems: FIND_THEOREMS = +struct + +(** search criteria **) + +datatype 'term criterion = + Name of string | Intro | Elim | Dest | Solves | Simp of 'term | + Pattern of 'term; + +fun read_criterion _ (Name name) = Name name + | read_criterion _ Intro = Intro + | read_criterion _ Elim = Elim + | read_criterion _ Dest = Dest + | read_criterion _ Solves = Solves + | read_criterion ctxt (Simp str) = Simp (ProofContext.read_term_pattern ctxt str) + | read_criterion ctxt (Pattern str) = Pattern (ProofContext.read_term_pattern ctxt str); + +fun pretty_criterion ctxt (b, c) = + let + fun prfx s = if b then s else "-" ^ s; + in + (case c of + Name name => Pretty.str (prfx "name: " ^ quote name) + | Intro => Pretty.str (prfx "intro") + | Elim => Pretty.str (prfx "elim") + | Dest => Pretty.str (prfx "dest") + | Solves => Pretty.str (prfx "solves") + | Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1, + Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))] + | Pattern pat => Pretty.enclose (prfx " \"") "\"" + [Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)]) + end; + + + +(** search criterion filters **) + +(*generated filters are to be of the form + input: (Facts.ref * thm) + output: (p:int, s:int) option, where + NONE indicates no match + p is the primary sorting criterion + (eg. number of assumptions in the theorem) + s is the secondary sorting criterion + (eg. size of the substitution for intro, elim and dest) + when applying a set of filters to a thm, fold results in: + (biggest p, sum of all s) + currently p and s only matter for intro, elim, dest and simp filters, + otherwise the default ordering is used. +*) + + +(* matching theorems *) + +fun is_nontrivial thy = Term.is_Const o Term.head_of o ObjectLogic.drop_judgment thy; + +(*extract terms from term_src, refine them to the parts that concern us, + if po try match them against obj else vice versa. + trivial matches are ignored. + returns: smallest substitution size*) +fun is_matching_thm (extract_terms, refine_term) ctxt po obj term_src = + let + val thy = ProofContext.theory_of ctxt; + + fun matches pat = + is_nontrivial thy pat andalso + Pattern.matches thy (if po then (pat, obj) else (obj, pat)); + + fun substsize pat = + let val (_, subst) = + Pattern.match thy (if po then (pat, obj) else (obj, pat)) (Vartab.empty, Vartab.empty) + in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end; + + fun bestmatch [] = NONE + | bestmatch xs = SOME (foldr1 Int.min xs); + + val match_thm = matches o refine_term; + in + map (substsize o refine_term) (filter match_thm (extract_terms term_src)) + |> bestmatch + end; + + +(* filter_name *) + +fun filter_name str_pat (thmref, _) = + if match_string str_pat (Facts.name_of_ref thmref) + then SOME (0, 0) else NONE; + + +(* filter intro/elim/dest/solves rules *) + +fun filter_dest ctxt goal (_, thm) = + let + val extract_dest = + (fn thm => if Thm.no_prems thm then [] else [Thm.full_prop_of thm], + hd o Logic.strip_imp_prems); + val prems = Logic.prems_of_goal goal 1; + + fun try_subst prem = is_matching_thm extract_dest ctxt true prem thm; + val successful = prems |> map_filter try_subst; + in + (*if possible, keep best substitution (one with smallest size)*) + (*dest rules always have assumptions, so a dest with one + assumption is as good as an intro rule with none*) + if not (null successful) + then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE + end; + +fun filter_intro ctxt goal (_, thm) = + let + val extract_intro = (single o Thm.full_prop_of, Logic.strip_imp_concl); + val concl = Logic.concl_of_goal goal 1; + val ss = is_matching_thm extract_intro ctxt true concl thm; + in + if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE + end; + +fun filter_elim ctxt goal (_, thm) = + if not (Thm.no_prems thm) then + let + val rule = Thm.full_prop_of thm; + val prems = Logic.prems_of_goal goal 1; + val goal_concl = Logic.concl_of_goal goal 1; + val rule_mp = hd (Logic.strip_imp_prems rule); + val rule_concl = Logic.strip_imp_concl rule; + fun combine t1 t2 = Const ("*combine*", dummyT --> dummyT) $ (t1 $ t2); + val rule_tree = combine rule_mp rule_concl; + fun goal_tree prem = combine prem goal_concl; + fun try_subst prem = + is_matching_thm (single, I) ctxt true (goal_tree prem) rule_tree; + val successful = prems |> map_filter try_subst; + in + (*elim rules always have assumptions, so an elim with one + assumption is as good as an intro rule with none*) + if is_nontrivial (ProofContext.theory_of ctxt) (Thm.major_prem_of thm) + andalso not (null successful) + then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE + end + else NONE + +val tac_limit = ref 5; + +fun filter_solves ctxt goal = + let + val baregoal = Logic.get_goal (Thm.prop_of goal) 1; + + fun etacn thm i = Seq.take (! tac_limit) o etac thm i; + fun try_thm thm = + if Thm.no_prems thm then rtac thm 1 goal + else (etacn thm THEN_ALL_NEW + (Goal.norm_hhf_tac THEN' Method.assm_tac ctxt)) 1 goal; + in + fn (_, thm) => + if (is_some o Seq.pull o try_thm) thm + then SOME (Thm.nprems_of thm, 0) else NONE + end; + + +(* filter_simp *) + +fun filter_simp ctxt t (_, thm) = + let + val (_, {mk_rews = {mk, ...}, ...}) = + Simplifier.rep_ss (Simplifier.local_simpset_of ctxt); + val extract_simp = + (map Thm.full_prop_of o mk, #1 o Logic.dest_equals o Logic.strip_imp_concl); + val ss = is_matching_thm extract_simp ctxt false t thm + in + if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE + end; + + +(* filter_pattern *) + +fun get_names t = (Term.add_const_names t []) union (Term.add_free_names t []); +fun get_thm_names (_, thm) = get_names (Thm.full_prop_of thm); + +(*Including all constants and frees is only sound because + matching uses higher-order patterns. If full matching + were used, then constants that may be subject to + beta-reduction after substitution of frees should + not be included for LHS set because they could be + thrown away by the substituted function. + e.g. for (?F 1 2) do not include 1 or 2, if it were + possible for ?F to be (% x y. 3) + The largest possible set should always be included on + the RHS.*) + +fun filter_pattern ctxt pat = + let + val pat_consts = get_names pat; + + fun check (t, NONE) = check (t, SOME (get_thm_names t)) + | check ((_, thm), c as SOME thm_consts) = + (if pat_consts subset_string thm_consts + andalso (Pattern.matches_subterm (ProofContext.theory_of ctxt) + (pat, Thm.full_prop_of thm)) + then SOME (0, 0) else NONE, c); + in check end; + + +(* interpret criteria as filters *) + +local + +fun err_no_goal c = + error ("Current goal required for " ^ c ^ " search criterion"); + +val fix_goal = Thm.prop_of; +val fix_goalo = Option.map fix_goal; + +fun filter_crit _ _ (Name name) = apfst (filter_name name) + | filter_crit _ NONE Intro = err_no_goal "intro" + | filter_crit _ NONE Elim = err_no_goal "elim" + | filter_crit _ NONE Dest = err_no_goal "dest" + | filter_crit _ NONE Solves = err_no_goal "solves" + | filter_crit ctxt (SOME goal) Intro = apfst (filter_intro ctxt (fix_goal goal)) + | filter_crit ctxt (SOME goal) Elim = apfst (filter_elim ctxt (fix_goal goal)) + | filter_crit ctxt (SOME goal) Dest = apfst (filter_dest ctxt (fix_goal goal)) + | filter_crit ctxt (SOME goal) Solves = apfst (filter_solves ctxt goal) + | filter_crit ctxt _ (Simp pat) = apfst (filter_simp ctxt pat) + | filter_crit ctxt _ (Pattern pat) = filter_pattern ctxt pat; + +fun opt_not x = if is_some x then NONE else SOME (0, 0); + +fun opt_add (SOME (a, x)) (SOME (b, y)) = SOME (Int.max (a, b), x + y : int) + | opt_add _ _ = NONE; + +fun app_filters thm = + let + fun app (NONE, _, _) = NONE + | app (SOME v, consts, []) = SOME (v, thm) + | app (r, consts, f :: fs) = + let val (r', consts') = f (thm, consts) + in app (opt_add r r', consts', fs) end; + in app end; + +in + +fun filter_criterion ctxt opt_goal (b, c) = + (if b then I else (apfst opt_not)) o filter_crit ctxt opt_goal c; + +fun all_filters filters thms = + let + fun eval_filters thm = app_filters thm (SOME (0, 0), NONE, filters); + + (*filters return: (number of assumptions, substitution size) option, so + sort (desc. in both cases) according to number of assumptions first, + then by the substitution size*) + fun thm_ord (((p0, s0), _), ((p1, s1), _)) = + prod_ord int_ord int_ord ((p1, s1), (p0, s0)); + in map_filter eval_filters thms |> sort thm_ord |> map #2 end; + +end; + + +(* removing duplicates, preferring nicer names, roughly n log n *) + +local + +val index_ord = option_ord (K EQUAL); +val hidden_ord = bool_ord o pairself NameSpace.is_hidden; +val qual_ord = int_ord o pairself (length o NameSpace.explode); +val txt_ord = int_ord o pairself size; + +fun nicer_name (x, i) (y, j) = + (case hidden_ord (x, y) of EQUAL => + (case index_ord (i, j) of EQUAL => + (case qual_ord (x, y) of EQUAL => txt_ord (x, y) | ord => ord) + | ord => ord) + | ord => ord) <> GREATER; + +fun rem_cdups nicer xs = + let + fun rem_c rev_seen [] = rev rev_seen + | rem_c rev_seen [x] = rem_c (x :: rev_seen) [] + | rem_c rev_seen ((x as ((n, t), _)) :: (y as ((n', t'), _)) :: xs) = + if Thm.eq_thm_prop (t, t') + then rem_c rev_seen ((if nicer n n' then x else y) :: xs) + else rem_c (x :: rev_seen) (y :: xs) + in rem_c [] xs end; + +in + +fun nicer_shortest ctxt = + let + (* FIXME global name space!? *) + val space = Facts.space_of (PureThy.facts_of (ProofContext.theory_of ctxt)); + + val shorten = + NameSpace.extern_flags {long_names = false, short_names = false, unique_names = false} space; + + fun nicer (Facts.Named ((x, _), i)) (Facts.Named ((y, _), j)) = + nicer_name (shorten x, i) (shorten y, j) + | nicer (Facts.Fact _) (Facts.Named _) = true + | nicer (Facts.Named _) (Facts.Fact _) = false; + in nicer end; + +fun rem_thm_dups nicer xs = + xs ~~ (1 upto length xs) + |> sort (TermOrd.fast_term_ord o pairself (Thm.prop_of o #2 o #1)) + |> rem_cdups nicer + |> sort (int_ord o pairself #2) + |> map #1; + +end; + + +(* print_theorems *) + +fun all_facts_of ctxt = + maps Facts.selections + (Facts.dest_static [] (PureThy.facts_of (ProofContext.theory_of ctxt)) @ + Facts.dest_static [] (ProofContext.facts_of ctxt)); + +val limit = ref 40; + +fun find_theorems ctxt opt_goal rem_dups raw_criteria = + let + val add_prems = Seq.hd o (TRY (Method.insert_tac (Assumption.prems_of ctxt) 1)); + val opt_goal' = Option.map add_prems opt_goal; + + val criteria = map (apsnd (read_criterion ctxt)) raw_criteria; + val filters = map (filter_criterion ctxt opt_goal') criteria; + + val raw_matches = all_filters filters (all_facts_of ctxt); + + val matches = + if rem_dups + then rem_thm_dups (nicer_shortest ctxt) raw_matches + else raw_matches; + in matches end; + + +fun pretty_thm ctxt (thmref, thm) = Pretty.block + [Pretty.str (Facts.string_of_ref thmref), Pretty.str ":", Pretty.brk 1, + ProofContext.pretty_thm ctxt thm]; + +fun print_theorems ctxt opt_goal opt_limit rem_dups raw_criteria = + let + val start = start_timing (); + + val criteria = map (apsnd (read_criterion ctxt)) raw_criteria; + val matches = find_theorems ctxt opt_goal rem_dups raw_criteria; + + val len = length matches; + val lim = the_default (! limit) opt_limit; + val thms = Library.drop (len - lim, matches); + + val end_msg = " in " ^ Time.toString (#all (end_timing start)) ^ " secs"; + in + Pretty.big_list "searched for:" (map (pretty_criterion ctxt) criteria) + :: Pretty.str "" :: + (if null thms then [Pretty.str ("nothing found" ^ end_msg)] + else + [Pretty.str ("found " ^ string_of_int len ^ " theorems" ^ + (if len <= lim then "" + else " (" ^ string_of_int lim ^ " displayed)") + ^ end_msg ^ ":"), Pretty.str ""] @ + map (pretty_thm ctxt) thms) + |> Pretty.chunks |> Pretty.writeln + end; + + + +(** command syntax **) + +fun find_theorems_cmd ((opt_lim, rem_dups), spec) = + Toplevel.unknown_theory o Toplevel.keep (fn state => + let + val proof_state = Toplevel.enter_proof_body state; + val ctxt = Proof.context_of proof_state; + val opt_goal = try Proof.get_goal proof_state |> Option.map (#2 o #2); + in print_theorems ctxt opt_goal opt_lim rem_dups spec end); + +local + +structure P = OuterParse and K = OuterKeyword; + +val criterion = + P.reserved "name" |-- P.!!! (P.$$$ ":" |-- P.xname) >> Name || + P.reserved "intro" >> K Intro || + P.reserved "elim" >> K Elim || + P.reserved "dest" >> K Dest || + P.reserved "solves" >> K Solves || + P.reserved "simp" |-- P.!!! (P.$$$ ":" |-- P.term) >> Simp || + P.term >> Pattern; + +val options = + Scan.optional + (P.$$$ "(" |-- + P.!!! (Scan.option P.nat -- Scan.optional (P.reserved "with_dups" >> K false) true + --| P.$$$ ")")) (NONE, true); +in + +val _ = + OuterSyntax.improper_command "find_theorems" "print theorems meeting specified criteria" K.diag + (options -- Scan.repeat (((Scan.option P.minus >> is_none) -- criterion)) + >> (Toplevel.no_timing oo find_theorems_cmd)); + +end; + +end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Tools/isabelle_process.ML --- a/src/Pure/Tools/isabelle_process.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,138 +0,0 @@ -(* Title: Pure/Tools/isabelle_process.ML - Author: Makarius - -Isabelle process wrapper -- interaction via external program. - -General format of process output: - - (1) unmarked stdout/stderr, no line structure (output should be - processed immediately as it arrives); - - (2) properly marked-up messages, e.g. for writeln channel - - "\002A" ^ props ^ "\002,\n" ^ text ^ "\002.\n" - - where the props consist of name=value lines terminated by "\002,\n" - each, and the remaining text is any number of lines (output is - supposed to be processed in one piece); - - (3) special init message holds "pid" and "session" property; - - (4) message content is encoded in YXML format. -*) - -signature ISABELLE_PROCESS = -sig - val isabelle_processN: string - val init: string -> unit -end; - -structure IsabelleProcess: ISABELLE_PROCESS = -struct - -(* print modes *) - -val isabelle_processN = "isabelle_process"; - -val _ = Output.add_mode isabelle_processN Output.default_output Output.default_escape; -val _ = Markup.add_mode isabelle_processN YXML.output_markup; - - -(* message markup *) - -fun special ch = Symbol.STX ^ ch; -val special_sep = special ","; -val special_end = special "."; - -local - -fun clean_string bad str = - if exists_string (member (op =) bad) str then - translate_string (fn c => if member (op =) bad c then Symbol.DEL else c) str - else str; - -fun message_props props = - let val clean = clean_string [Symbol.STX, "\n", "\r"] - in implode (map (fn (x, y) => clean x ^ "=" ^ clean y ^ special_sep ^ "\n") props) end; - -fun message_pos trees = trees |> get_first - (fn XML.Elem (name, atts, ts) => - if name = Markup.positionN then SOME (Position.of_properties atts) - else message_pos ts - | _ => NONE); - -fun output out_stream s = NAMED_CRITICAL "IO" (fn () => - (TextIO.output (out_stream, s); TextIO.output (out_stream, "\n"))); - -in - -fun message _ _ "" = () - | message out_stream ch body = - let - val pos = the_default Position.none (message_pos (YXML.parse_body body)); - val props = - Position.properties_of (Position.thread_data ()) - |> Position.default_properties pos; - val txt = clean_string [Symbol.STX] body; - in output out_stream (special ch ^ message_props props ^ txt ^ special_end) end; - -fun init_message out_stream = - let - val pid = (Markup.pidN, process_id ()); - val session = (Markup.sessionN, List.last (Session.id ()) handle List.Empty => "unknown"); - val text = Session.welcome (); - in output out_stream (special "A" ^ message_props [pid, session] ^ text ^ special_end) end; - -end; - - -(* channels *) - -local - -fun auto_flush stream = - let - val _ = TextIO.StreamIO.setBufferMode (TextIO.getOutstream stream, IO.BLOCK_BUF); - fun loop () = - (OS.Process.sleep (Time.fromMilliseconds 50); try TextIO.flushOut stream; loop ()); - in loop end; - -in - -fun setup_channels out = - let - val out_stream = - if out = "-" then TextIO.stdOut - else - let - val path = File.platform_path (Path.explode out); - val out_stream = TextIO.openOut path; (*fifo blocks until reader is ready*) - val _ = OS.FileSys.remove path; (*prevent alien access, indicate writer is ready*) - val _ = SimpleThread.fork false (auto_flush TextIO.stdOut); - in out_stream end; - val _ = SimpleThread.fork false (auto_flush out_stream); - val _ = SimpleThread.fork false (auto_flush TextIO.stdErr); - in - Output.status_fn := message out_stream "B"; - Output.writeln_fn := message out_stream "C"; - Output.priority_fn := message out_stream "D"; - Output.tracing_fn := message out_stream "E"; - Output.warning_fn := message out_stream "F"; - Output.error_fn := message out_stream "G"; - Output.debug_fn := message out_stream "H"; - Output.prompt_fn := ignore; - out_stream - end; - -end; - - -(* init *) - -fun init out = - (change print_mode (update (op =) isabelle_processN); - setup_channels out |> init_message; - OuterKeyword.report (); - Isar.toplevel_loop {init = true, welcome = false, sync = true, secure = true}); - -end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Tools/isabelle_process.scala --- a/src/Pure/Tools/isabelle_process.scala Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,435 +0,0 @@ -/* Title: Pure/Tools/isabelle_process.ML - Author: Makarius - Options: :folding=explicit:collapseFolds=1: - -Isabelle process management -- always reactive due to multi-threaded I/O. -*/ - -package isabelle - -import java.util.concurrent.LinkedBlockingQueue -import java.io.{BufferedReader, BufferedWriter, InputStreamReader, OutputStreamWriter, - InputStream, OutputStream, IOException} - - -object IsabelleProcess { - - /* results */ - - object Kind extends Enumeration { - //{{{ values and codes - // internal system notification - val SYSTEM = Value("SYSTEM") - // Posix channels/events - val STDIN = Value("STDIN") - val STDOUT = Value("STDOUT") - val SIGNAL = Value("SIGNAL") - val EXIT = Value("EXIT") - // Isabelle messages - val INIT = Value("INIT") - val STATUS = Value("STATUS") - val WRITELN = Value("WRITELN") - val PRIORITY = Value("PRIORITY") - val TRACING = Value("TRACING") - val WARNING = Value("WARNING") - val ERROR = Value("ERROR") - val DEBUG = Value("DEBUG") - // messages codes - val code = Map( - ('A' : Int) -> Kind.INIT, - ('B' : Int) -> Kind.STATUS, - ('C' : Int) -> Kind.WRITELN, - ('D' : Int) -> Kind.PRIORITY, - ('E' : Int) -> Kind.TRACING, - ('F' : Int) -> Kind.WARNING, - ('G' : Int) -> Kind.ERROR, - ('H' : Int) -> Kind.DEBUG, - ('0' : Int) -> Kind.SYSTEM, - ('1' : Int) -> Kind.STDIN, - ('2' : Int) -> Kind.STDOUT, - ('3' : Int) -> Kind.SIGNAL, - ('4' : Int) -> Kind.EXIT) - // message markup - val markup = Map( - Kind.INIT -> Markup.INIT, - Kind.STATUS -> Markup.STATUS, - Kind.WRITELN -> Markup.WRITELN, - Kind.PRIORITY -> Markup.PRIORITY, - Kind.TRACING -> Markup.TRACING, - Kind.WARNING -> Markup.WARNING, - Kind.ERROR -> Markup.ERROR, - Kind.DEBUG -> Markup.DEBUG, - Kind.SYSTEM -> Markup.SYSTEM, - Kind.STDIN -> Markup.STDIN, - Kind.STDOUT -> Markup.STDOUT, - Kind.SIGNAL -> Markup.SIGNAL, - Kind.EXIT -> Markup.EXIT) - //}}} - def is_raw(kind: Value) = - kind == STDOUT - def is_control(kind: Value) = - kind == SYSTEM || - kind == SIGNAL || - kind == EXIT - def is_system(kind: Value) = - kind == SYSTEM || - kind == STDIN || - kind == SIGNAL || - kind == EXIT || - kind == STATUS - } - - class Result(val kind: Kind.Value, val props: List[(String, String)], val result: String) { - override def toString = { - val trees = YXML.parse_body_failsafe(result) - val res = - if (kind == Kind.STATUS) trees.map(_.toString).mkString - else trees.flatMap(XML.content(_).mkString).mkString - if (props.isEmpty) - kind.toString + " [[" + res + "]]" - else - kind.toString + " " + - (for ((x, y) <- props) yield x + "=" + y).mkString("{", ",", "}") + " [[" + res + "]]" - } - def is_raw = Kind.is_raw(kind) - def is_control = Kind.is_control(kind) - def is_system = Kind.is_system(kind) - } - - def parse_message(isabelle_system: IsabelleSystem, result: Result) = - { - XML.Elem(Markup.MESSAGE, (Markup.CLASS, Kind.markup(result.kind)) :: result.props, - YXML.parse_body_failsafe(isabelle_system.symbols.decode(result.result))) - } -} - - -class IsabelleProcess(isabelle_system: IsabelleSystem, - results: EventBus[IsabelleProcess.Result], args: String*) -{ - import IsabelleProcess._ - - - /* demo constructor */ - - def this(args: String*) = - this(new IsabelleSystem, new EventBus[IsabelleProcess.Result] + Console.println, args: _*) - - - /* process information */ - - @volatile private var proc: Process = null - @volatile private var closing = false - @volatile private var pid: String = null - @volatile private var the_session: String = null - def session = the_session - - - /* results */ - - def parse_message(result: Result): XML.Tree = - IsabelleProcess.parse_message(isabelle_system, result) - - private val result_queue = new LinkedBlockingQueue[Result] - - private def put_result(kind: Kind.Value, props: List[(String, String)], result: String) - { - if (kind == Kind.INIT) { - val map = Map(props: _*) - if (map.isDefinedAt(Markup.PID)) pid = map(Markup.PID) - if (map.isDefinedAt(Markup.SESSION)) the_session = map(Markup.SESSION) - } - result_queue.put(new Result(kind, props, result)) - } - - private class ResultThread extends Thread("isabelle: results") { - override def run() = { - var finished = false - while (!finished) { - val result = - try { result_queue.take } - catch { case _: NullPointerException => null } - - if (result != null) { - results.event(result) - if (result.kind == Kind.EXIT) finished = true - } - else finished = true - } - } - } - - - /* signals */ - - def interrupt() = synchronized { - if (proc == null) error("Cannot interrupt Isabelle: no process") - if (pid == null) put_result(Kind.SYSTEM, Nil, "Cannot interrupt: unknown pid") - else { - try { - if (isabelle_system.execute(true, "kill", "-INT", pid).waitFor == 0) - put_result(Kind.SIGNAL, Nil, "INT") - else - put_result(Kind.SYSTEM, Nil, "Cannot interrupt: kill command failed") - } - catch { case e: IOException => error("Cannot interrupt Isabelle: " + e.getMessage) } - } - } - - def kill() = synchronized { - if (proc == 0) error("Cannot kill Isabelle: no process") - else { - try_close() - Thread.sleep(500) - put_result(Kind.SIGNAL, Nil, "KILL") - proc.destroy - proc = null - pid = null - } - } - - - /* output being piped into the process */ - - private val output = new LinkedBlockingQueue[String] - - private def output_raw(text: String) = synchronized { - if (proc == null) error("Cannot output to Isabelle: no process") - if (closing) error("Cannot output to Isabelle: already closing") - output.put(text) - } - - def output_sync(text: String) = - output_raw(" \\<^sync>\n; " + text + " \\<^sync>;\n") - - - def command(text: String) = - output_sync("Isabelle.command " + IsabelleSyntax.encode_string(text)) - - def command(props: List[(String, String)], text: String) = - output_sync("Isabelle.command " + IsabelleSyntax.encode_properties(props) + " " + - IsabelleSyntax.encode_string(text)) - - def ML(text: String) = - output_sync("ML_val " + IsabelleSyntax.encode_string(text)) - - def close() = synchronized { // FIXME watchdog/timeout - output_raw("\u0000") - closing = true - } - - def try_close() = synchronized { - if (proc != null && !closing) { - try { close() } - catch { case _: RuntimeException => } - } - } - - - /* stdin */ - - private class StdinThread(out_stream: OutputStream) extends Thread("isabelle: stdin") { - override def run() = { - val writer = new BufferedWriter(new OutputStreamWriter(out_stream, isabelle_system.charset)) - var finished = false - while (!finished) { - try { - //{{{ - val s = output.take - if (s == "\u0000") { - writer.close - finished = true - } - else { - put_result(Kind.STDIN, Nil, s) - writer.write(s) - writer.flush - } - //}}} - } - catch { - case e: IOException => put_result(Kind.SYSTEM, Nil, "Stdin thread: " + e.getMessage) - } - } - put_result(Kind.SYSTEM, Nil, "Stdin thread terminated") - } - } - - - /* stdout */ - - private class StdoutThread(in_stream: InputStream) extends Thread("isabelle: stdout") { - override def run() = { - val reader = new BufferedReader(new InputStreamReader(in_stream, isabelle_system.charset)) - var result = new StringBuilder(100) - - var finished = false - while (!finished) { - try { - //{{{ - var c = -1 - var done = false - while (!done && (result.length == 0 || reader.ready)) { - c = reader.read - if (c >= 0) result.append(c.asInstanceOf[Char]) - else done = true - } - if (result.length > 0) { - put_result(Kind.STDOUT, Nil, result.toString) - result.length = 0 - } - else { - reader.close - finished = true - try_close() - } - //}}} - } - catch { - case e: IOException => put_result(Kind.SYSTEM, Nil, "Stdout thread: " + e.getMessage) - } - } - put_result(Kind.SYSTEM, Nil, "Stdout thread terminated") - } - } - - - /* messages */ - - private class MessageThread(fifo: String) extends Thread("isabelle: messages") { - override def run() = { - val reader = isabelle_system.fifo_reader(fifo) - var kind: Kind.Value = null - var props: List[(String, String)] = Nil - var result = new StringBuilder - - var finished = false - while (!finished) { - try { - if (kind == null) { - //{{{ Char mode -- resync - var c = -1 - do { - c = reader.read - if (c >= 0 && c != 2) result.append(c.asInstanceOf[Char]) - } while (c >= 0 && c != 2) - - if (result.length > 0) { - put_result(Kind.SYSTEM, Nil, "Malformed message:\n" + result.toString) - result.length = 0 - } - if (c < 0) { - reader.close - finished = true - try_close() - } - else { - c = reader.read - if (Kind.code.isDefinedAt(c)) kind = Kind.code(c) - else kind = null - } - //}}} - } - else { - //{{{ Line mode - val line = reader.readLine - if (line == null) { - reader.close - finished = true - try_close() - } - else { - val len = line.length - // property - if (line.endsWith("\u0002,")) { - val i = line.indexOf('=') - if (i > 0) { - val name = line.substring(0, i) - val value = line.substring(i + 1, len - 2) - props = (name, value) :: props - } - } - // last text line - else if (line.endsWith("\u0002.")) { - result.append(line.substring(0, len - 2)) - put_result(kind, props.reverse, result.toString) - kind = null - props = Nil - result.length = 0 - } - // text line - else { - result.append(line) - result.append('\n') - } - } - //}}} - } - } - catch { - case e: IOException => put_result(Kind.SYSTEM, Nil, "Message thread: " + e.getMessage) - } - } - put_result(Kind.SYSTEM, Nil, "Message thread terminated") - } - } - - - - /** main **/ - - { - /* isabelle version */ - - { - val (msg, rc) = isabelle_system.isabelle_tool("version") - if (rc != 0) error("Version check failed -- bad Isabelle installation:\n" + msg) - put_result(Kind.SYSTEM, Nil, msg) - } - - - /* messages */ - - val message_fifo = isabelle_system.mk_fifo() - def rm_fifo() = isabelle_system.rm_fifo(message_fifo) - - val message_thread = new MessageThread(message_fifo) - message_thread.start - - new ResultThread().start - - - /* exec process */ - - try { - val cmdline = - List(isabelle_system.getenv_strict("ISABELLE_PROCESS"), "-W", message_fifo) ++ args - proc = isabelle_system.execute(true, cmdline: _*) - } - catch { - case e: IOException => - rm_fifo() - error("Failed to execute Isabelle process: " + e.getMessage) - } - - - /* stdin/stdout */ - - new StdinThread(proc.getOutputStream).start - new StdoutThread(proc.getInputStream).start - - - /* exit */ - - new Thread("isabelle: exit") { - override def run() = { - val rc = proc.waitFor() - Thread.sleep(300) - put_result(Kind.SYSTEM, Nil, "Exit thread terminated") - put_result(Kind.EXIT, Nil, Integer.toString(rc)) - rm_fifo() - } - }.start - - } -} diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/Tools/isabelle_system.scala --- a/src/Pure/Tools/isabelle_system.scala Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,158 +0,0 @@ -/* Title: Pure/Tools/isabelle_system.scala - Author: Makarius - -Isabelle system support -- basic Cygwin/Posix compatibility. -*/ - -package isabelle - -import java.util.regex.{Pattern, Matcher} -import java.io.{BufferedReader, InputStreamReader, FileInputStream, File, IOException} - -import scala.io.Source - - -class IsabelleSystem { - - val charset = "UTF-8" - - - /* Isabelle environment settings */ - - private val environment = System.getenv - - def getenv(name: String) = { - val value = environment.get(if (name == "HOME") "HOME_JVM" else name) - if (value != null) value else "" - } - - def getenv_strict(name: String) = { - val value = environment.get(name) - if (value != "") value else error("Undefined environment variable: " + name) - } - - val is_cygwin = Pattern.matches(".*-cygwin", getenv_strict("ML_PLATFORM")) - - - /* file path specifications */ - - private val cygdrive_pattern = Pattern.compile("/cygdrive/([a-zA-Z])($|/.*)") - - def platform_path(source_path: String) = { - val result_path = new StringBuilder - - def init(path: String) = { - val cygdrive = cygdrive_pattern.matcher(path) - if (cygdrive.matches) { - result_path.length = 0 - result_path.append(cygdrive.group(1)) - result_path.append(":") - result_path.append(File.separator) - cygdrive.group(2) - } - else if (path.startsWith("/")) { - result_path.length = 0 - result_path.append(getenv_strict("ISABELLE_ROOT_JVM")) - path.substring(1) - } - else path - } - def append(path: String) = { - for (p <- init(path).split("/")) { - if (p != "") { - val len = result_path.length - if (len > 0 && result_path(len - 1) != File.separatorChar) - result_path.append(File.separator) - result_path.append(p) - } - } - } - for (p <- init(source_path).split("/")) { - if (p.startsWith("$")) append(getenv_strict(p.substring(1))) - else if (p == "~") append(getenv_strict("HOME")) - else if (p == "~~") append(getenv_strict("ISABELLE_HOME")) - else append(p) - } - result_path.toString - } - - def platform_file(path: String) = - new File(platform_path(path)) - - - /* processes */ - - def execute(redirect: Boolean, args: String*): Process = { - val cmdline = new java.util.LinkedList[String] - if (is_cygwin) cmdline.add(platform_path("/bin/env")) - for (s <- args) cmdline.add(s) - - val proc = new ProcessBuilder(cmdline) - proc.environment.clear - proc.environment.putAll(environment) - proc.redirectErrorStream(redirect) - proc.start - } - - - /* Isabelle tools (non-interactive) */ - - def isabelle_tool(args: String*) = { - val proc = - try { execute(true, (List(getenv_strict("ISABELLE_TOOL")) ++ args): _*) } - catch { case e: IOException => error(e.getMessage) } - proc.getOutputStream.close - val output = Source.fromInputStream(proc.getInputStream, charset).mkString - val rc = proc.waitFor - (output, rc) - } - - - /* named pipes */ - - def mk_fifo() = { - val (result, rc) = isabelle_tool("mkfifo") - if (rc == 0) result.trim - else error(result) - } - - def rm_fifo(fifo: String) = { - val (result, rc) = isabelle_tool("rmfifo", fifo) - if (rc != 0) error(result) - } - - def fifo_reader(fifo: String) = { - // blocks until writer is ready - val stream = - if (is_cygwin) execute(false, "cat", fifo).getInputStream - else new FileInputStream(fifo) - new BufferedReader(new InputStreamReader(stream, charset)) - } - - - /* find logics */ - - def find_logics() = { - val ml_ident = getenv_strict("ML_IDENTIFIER") - var logics: Set[String] = Set() - for (dir <- getenv_strict("ISABELLE_PATH").split(":")) { - val files = platform_file(dir + "/" + ml_ident).listFiles() - if (files != null) { - for (file <- files if file.isFile) logics += file.getName - } - } - logics.toList.sort(_ < _) - } - - - /* symbols */ - - private def read_symbols(path: String) = { - val file = new File(platform_path(path)) - if (file.canRead) Source.fromFile(file).getLines - else Iterator.empty - } - val symbols = new Symbol.Interpretation( - read_symbols("$ISABELLE_HOME/etc/symbols") ++ - read_symbols("$ISABELLE_HOME_USER/etc/symbols")) -} diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/pure_thy.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/sign.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/sorts.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Pure/term.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Tools/auto_solve.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Tools/code/code_funcgr.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Tools/code/code_funcgr_new.ML --- a/src/Tools/code/code_funcgr_new.ML Wed Mar 04 11:05:02 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,414 +0,0 @@ -(* Title: Tools/code/code_funcgr.ML - ID: $Id$ - Author: Florian Haftmann, TU Muenchen - -Retrieving, well-sorting and structuring defining equations in graph -with explicit dependencies. -*) - -signature CODE_FUNCGR = -sig - type T - val eqns: T -> string -> (thm * bool) list - val typ: T -> string -> (string * sort) list * typ - val all: T -> string list - val pretty: theory -> T -> Pretty.T - val make: theory -> string list - -> ((sort -> sort) * Sorts.algebra) * T - val eval_conv: theory - -> (term -> term * (((sort -> sort) * Sorts.algebra) -> T -> thm)) -> cterm -> thm - val eval_term: theory - -> (term -> term * (((sort -> sort) * Sorts.algebra) -> T -> 'a)) -> term -> 'a -end - -structure Code_Funcgr : CODE_FUNCGR = -struct - -(** the graph type **) - -type T = (((string * sort) list * typ) * (thm * bool) list) Graph.T; - -fun eqns funcgr = - these o Option.map snd o try (Graph.get_node funcgr); - -fun typ funcgr = - fst o Graph.get_node funcgr; - -fun all funcgr = Graph.keys funcgr; - -fun pretty thy funcgr = - AList.make (snd o Graph.get_node funcgr) (Graph.keys funcgr) - |> (map o apfst) (Code_Unit.string_of_const thy) - |> sort (string_ord o pairself fst) - |> map (fn (s, thms) => - (Pretty.block o Pretty.fbreaks) ( - Pretty.str s - :: map (Display.pretty_thm o fst) thms - )) - |> Pretty.chunks; - - -(** generic combinators **) - -fun fold_consts f thms = - thms - |> maps (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of) - |> (fold o fold_aterms) (fn Const c => f c | _ => I); - -fun consts_of (const, []) = [] - | consts_of (const, thms as _ :: _) = - let - fun the_const (c, _) = if c = const then I else insert (op =) c - in fold_consts the_const (map fst thms) [] end; - - -(** graph algorithm **) - -(* some nonsense -- FIXME *) - -fun lhs_rhss_of thy c = - let - val eqns = Code.these_eqns thy c - |> burrow_fst (Code_Unit.norm_args thy) - |> burrow_fst (Code_Unit.norm_varnames thy Code_Name.purify_tvar Code_Name.purify_var); - val (lhs, _) = case eqns of [] => Code.default_typscheme thy c - | ((thm, _) :: _) => (snd o Code_Unit.head_eqn thy) thm; - val rhss = fold_consts (fn (c, ty) => - insert (op =) (c, Sign.const_typargs thy (c, Logic.unvarifyT ty))) (map fst eqns) []; - in (lhs, rhss) end; - -fun inst_params thy tyco class = - map (fn (c, _) => AxClass.param_of_inst thy (c, tyco)) - ((#params o AxClass.get_info thy) class); - -fun complete_proper_sort thy sort = - Sign.complete_sort thy sort |> filter (can (AxClass.get_info thy)); - -fun minimal_proper_sort thy sort = - complete_proper_sort thy sort |> Sign.minimize_sort thy; - -fun dicts_of thy algebra (T, sort) = - let - fun class_relation (x, _) _ = x; - fun type_constructor tyco xs class = - inst_params thy tyco class @ (maps o maps) fst xs; - fun type_variable (TFree (_, sort)) = map (pair []) sort; - in - flat (Sorts.of_sort_derivation (Syntax.pp_global thy) algebra - { class_relation = class_relation, type_constructor = type_constructor, - type_variable = type_variable } (T, minimal_proper_sort thy sort) - handle Sorts.CLASS_ERROR _ => [] (*permissive!*)) - end; - - -(* data structures *) - -datatype const = Fun of string | Inst of class * string; - -fun const_ord (Fun c1, Fun c2) = fast_string_ord (c1, c2) - | const_ord (Inst class_tyco1, Inst class_tyco2) = - prod_ord fast_string_ord fast_string_ord (class_tyco1, class_tyco2) - | const_ord (Fun _, Inst _) = LESS - | const_ord (Inst _, Fun _) = GREATER; - -type var = const * int; - -structure Vargraph = - GraphFun(type key = var val ord = prod_ord const_ord int_ord); - -datatype styp = Tyco of string * styp list | Var of var; - -type vardeps = const list * ((string * styp list) list * class list) Vargraph.T; - - -(* computing instantiations -- FIXME does not consider existing things *) - -fun add_classes thy c_k new_classes vardeps = - let - val _ = tracing "add_classes"; - val (styps, old_classes) = Vargraph.get_node (snd vardeps) c_k; - val diff_classes = new_classes |> subtract (op =) old_classes; - in if null diff_classes then vardeps - else let - val c_ks = Vargraph.imm_succs (snd vardeps) c_k |> insert (op =) c_k; - in - vardeps - |> (apsnd o Vargraph.map_node c_k o apsnd) (append diff_classes) - |> fold (fn styp => fold (add_typmatch_inst thy styp) new_classes) styps - |> fold (fn c_k => add_classes thy c_k diff_classes) c_ks - end end -and add_styp thy c_k tyco_styps vardeps = - let - val _ = tracing "add_styp"; - val (old_styps, classes) = Vargraph.get_node (snd vardeps) c_k; - in if member (op =) old_styps tyco_styps then vardeps - else - vardeps - |> (apsnd o Vargraph.map_node c_k o apfst) (cons tyco_styps) - |> fold (add_typmatch_inst thy tyco_styps) classes - end -and add_dep thy c_k c_k' vardeps = - let - val _ = tracing ("add_dep " ^ makestring c_k ^ " -> " ^ makestring c_k'); - val (_, classes) = Vargraph.get_node (snd vardeps) c_k; - in - vardeps - |> add_classes thy c_k' classes - |> apsnd (Vargraph.add_edge (c_k, c_k')) - end -and add_typmatch_inst thy (tyco, styps) class vardeps = if can (Sign.arity_sorts thy tyco) [class] - then vardeps - |> tap (fn _ => tracing "add_typmatch_inst") - |> assert thy (Inst (class, tyco)) - |> fold_index (fn (k, styp) => - add_typmatch thy styp (Inst (class, tyco), k)) styps - else vardeps (*permissive!*) -and add_typmatch thy (Var c_k') c_k vardeps = - vardeps - |> tap (fn _ => tracing "add_typmatch (Inst)") - |> add_dep thy c_k c_k' - | add_typmatch thy (Tyco tyco_styps) c_k vardeps = - vardeps - |> tap (fn _ => tracing "add_typmatch (Tyco)") - |> add_styp thy c_k tyco_styps -and add_inst thy (class, tyco) vardeps = - let - val _ = tracing ("add_inst " ^ tyco ^ "::" ^ class); - val superclasses = complete_proper_sort thy - (Sign.super_classes thy class); - val classess = map (complete_proper_sort thy) - (Sign.arity_sorts thy tyco [class]); - val inst_params = inst_params thy tyco class; - in - vardeps - |> fold (fn superclass => assert thy (Inst (superclass, tyco))) superclasses - |> fold (assert thy o Fun) inst_params - |> fold_index (fn (k, classes) => - apsnd (Vargraph.default_node ((Inst (class, tyco), k), ([] ,[]))) - #> add_classes thy (Inst (class, tyco), k) classes - #> fold (fn superclass => - add_dep thy (Inst (superclass, tyco), k) - (Inst (class, tyco), k)) superclasses - #> fold (fn inst_param => - add_dep thy (Fun inst_param, k) - (Inst (class, tyco), k) - ) inst_params - ) classess - end -and add_const thy c vardeps = - let - val _ = tracing "add_const"; - val (lhs, rhss) = lhs_rhss_of thy c; - fun styp_of (Type (tyco, tys)) = Tyco (tyco, map styp_of tys) - | styp_of (TFree (v, _)) = Var (Fun c, find_index (fn (v', _) => v = v') lhs); - val rhss' = (map o apsnd o map) styp_of rhss; - in - vardeps - |> fold_index (fn (k, (_, sort)) => - apsnd (Vargraph.default_node ((Fun c, k), ([] ,[]))) - #> add_classes thy (Fun c, k) (complete_proper_sort thy sort)) lhs - |> fold (assert thy o Fun o fst) rhss' - |> fold (fn (c', styps) => fold_index (fn (k', styp) => - add_typmatch thy styp (Fun c', k')) styps) rhss' - end -and assert thy c (vardeps as (asserted, _)) = - if member (op =) asserted c then vardeps - else case c - of Fun const => vardeps |> apfst (cons c) |> add_const thy const - | Inst inst => vardeps |> apfst (cons c) |> add_inst thy inst; - - -(* applying instantiations *) - -fun algebra_of thy vardeps = - let - val pp = Syntax.pp_global thy; - val thy_algebra = Sign.classes_of thy; - val is_proper = can (AxClass.get_info thy); - val arities = Vargraph.fold (fn ((Fun _, _), _) => I - | ((Inst (class, tyco), k), ((_, classes), _)) => - AList.map_default (op =) - ((tyco, class), replicate (Sign.arity_number thy tyco) []) - (nth_map k (K classes))) vardeps []; - val classrels = Sorts.classrels_of thy_algebra - |> filter (is_proper o fst) - |> (map o apsnd) (filter is_proper); - fun add_arity (tyco, class) = case AList.lookup (op =) arities (tyco, class) - of SOME sorts => Sorts.add_arities pp (tyco, [(class, sorts)]) - | NONE => if Sign.arity_number thy tyco = 0 - then (tracing (tyco ^ "::" ^ class); Sorts.add_arities pp (tyco, [(class, [])])) - else I; - val instances = Sorts.instances_of thy_algebra - |> filter (is_proper o snd) - in - Sorts.empty_algebra - |> fold (Sorts.add_class pp) classrels - |> fold add_arity instances - end; - -fun add_eqs thy algebra vardeps c gr = - let - val eqns = Code.these_eqns thy c - |> burrow_fst (Code_Unit.norm_args thy) - |> burrow_fst (Code_Unit.norm_varnames thy Code_Name.purify_tvar Code_Name.purify_var); - val (vs, _) = case eqns of [] => Code.default_typscheme thy c - | ((thm, _) :: _) => (snd o Code_Unit.head_eqn thy) thm; - val inst = Vartab.empty |> fold_index (fn (k, (v, _)) => - Vartab.update ((v, 0), snd (Vargraph.get_node vardeps (Fun c, k)))) vs; - val eqns' = eqns - |> (map o apfst) (Code_Unit.inst_thm thy inst); - val tyscm = case eqns' of [] => Code.default_typscheme thy c - | ((thm, _) :: _) => (snd o Code_Unit.head_eqn thy) thm; - val _ = tracing ("tyscm " ^ makestring (map snd (fst tyscm))); - val rhss = fold_consts (fn (c, ty) => - insert (op =) (c, Sign.const_typargs thy (c, Logic.unvarifyT ty))) (map fst eqns') []; - in - gr - |> Graph.new_node (c, (tyscm, eqns')) - |> fold (fn (c', Ts) => ensure_eqs_dep thy algebra vardeps c c' - #-> (fn (vs, _) => - fold2 (ensure_match thy algebra vardeps c) Ts (map snd vs))) rhss - |> pair tyscm - end -and ensure_match thy algebra vardeps c T sort gr = - gr - |> fold (fn c' => ensure_eqs_dep thy algebra vardeps c c' #> snd) - (dicts_of thy algebra (T, sort)) -and ensure_eqs_dep thy algebra vardeps c c' gr = - gr - |> ensure_eqs thy algebra vardeps c' - ||> Graph.add_edge (c, c') -and ensure_eqs thy algebra vardeps c gr = - case try (Graph.get_node gr) c - of SOME (tyscm, _) => (tyscm, gr) - | NONE => add_eqs thy algebra vardeps c gr; - -fun extend_graph thy cs gr = - let - val _ = tracing ("extending with " ^ commas cs); - val _ = tracing "obtaining instantiations"; - val (_, vardeps) = fold (assert thy o Fun) cs ([], Vargraph.empty) - val _ = tracing "obtaining algebra"; - val algebra = algebra_of thy vardeps; - val _ = tracing "obtaining equations"; - val (_, gr) = fold_map (ensure_eqs thy algebra vardeps) cs gr; - val _ = tracing "sort projection"; - val minimal_proper_sort = fn sort => sort - |> Sorts.complete_sort (Sign.classes_of thy) - |> filter (can (AxClass.get_info thy)) - |> Sorts.minimize_sort algebra; - in ((minimal_proper_sort, algebra), gr) end; - - -(** retrieval interfaces **) - -fun proto_eval thy cterm_of evaluator_lift evaluator proto_ct funcgr = - let - val ct = cterm_of proto_ct; - val _ = Sign.no_vars (Syntax.pp_global thy) (Thm.term_of ct); - val _ = Term.fold_types (Type.no_tvars #> K I) (Thm.term_of ct) (); - fun consts_of t = - fold_aterms (fn Const c_ty => cons c_ty | _ => I) t []; - val thm = Code.preprocess_conv thy ct; - val ct' = Thm.rhs_of thm; - val t' = Thm.term_of ct'; - val consts = map fst (consts_of t'); - val (algebra', funcgr') = extend_graph thy consts funcgr; - val (t'', evaluator_funcgr) = evaluator t'; - val consts' = consts_of t''; - val const_matches = fold (fn (c, ty) => - insert (op =) (Sign.const_typargs thy (c, Logic.unvarifyT ty), c)) consts' []; - val typ_matches = maps (fn (tys, c) => tys ~~ map snd (fst (fst (Graph.get_node funcgr' c)))) - const_matches; - val dicts = maps (dicts_of thy (snd algebra')) typ_matches; - val (algebra'', funcgr'') = extend_graph thy dicts funcgr'; - in (evaluator_lift (evaluator_funcgr algebra'') thm funcgr'', funcgr'') end; - -fun proto_eval_conv thy = - let - fun evaluator_lift evaluator thm1 funcgr = - let - val thm2 = evaluator funcgr; - val thm3 = Code.postprocess_conv thy (Thm.rhs_of thm2); - in - Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ => - error ("could not construct evaluation proof:\n" - ^ (cat_lines o map Display.string_of_thm) [thm1, thm2, thm3]) - end; - in proto_eval thy I evaluator_lift end; - -fun proto_eval_term thy = - let - fun evaluator_lift evaluator _ funcgr = evaluator funcgr; - in proto_eval thy (Thm.cterm_of thy) evaluator_lift end; - -structure Funcgr = CodeDataFun -( - type T = T; - val empty = Graph.empty; - fun purge _ cs funcgr = - Graph.del_nodes ((Graph.all_preds funcgr - o filter (can (Graph.get_node funcgr))) cs) funcgr; -); - -fun make thy = Funcgr.change_yield thy o extend_graph thy; - -fun eval_conv thy f = - fst o Funcgr.change_yield thy o proto_eval_conv thy f; - -fun eval_term thy f = - fst o Funcgr.change_yield thy o proto_eval_term thy f; - - -(** diagnostic commands **) - -fun code_depgr thy consts = - let - val (_, gr) = make thy consts; - val select = Graph.all_succs gr consts; - in - gr - |> not (null consts) ? Graph.subgraph (member (op =) select) - |> Graph.map_nodes ((apsnd o map o apfst) (AxClass.overload thy)) - end; - -fun code_thms thy = Pretty.writeln o pretty thy o code_depgr thy; - -fun code_deps thy consts = - let - val gr = code_depgr thy consts; - fun mk_entry (const, (_, (_, parents))) = - let - val name = Code_Unit.string_of_const thy const; - val nameparents = map (Code_Unit.string_of_const thy) parents; - in { name = name, ID = name, dir = "", unfold = true, - path = "", parents = nameparents } - end; - val prgr = Graph.fold ((fn x => fn xs => xs @ [x]) o mk_entry) gr []; - in Present.display_graph prgr end; - -local - -structure P = OuterParse -and K = OuterKeyword - -fun code_thms_cmd thy = code_thms thy o op @ o Code_Name.read_const_exprs thy; -fun code_deps_cmd thy = code_deps thy o op @ o Code_Name.read_const_exprs thy; - -in - -val _ = - OuterSyntax.improper_command "code_thms" "print system of defining equations for code" OuterKeyword.diag - (Scan.repeat P.term_group - >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory - o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of))); - -val _ = - OuterSyntax.improper_command "code_deps" "visualize dependencies of defining equations for code" OuterKeyword.diag - (Scan.repeat P.term_group - >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory - o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of))); - -end; - -end; (*struct*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Tools/code/code_target.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Tools/code/code_thingol.ML diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Tools/code/code_wellsorted.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/code/code_wellsorted.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,390 @@ +(* Title: Tools/code/code_wellsorted.ML + Author: Florian Haftmann, TU Muenchen + +Producing well-sorted systems of code equations in a graph +with explicit dependencies -- the Waisenhaus algorithm. +*) + +signature CODE_WELLSORTED = +sig + type T + val eqns: T -> string -> (thm * bool) list + val typ: T -> string -> (string * sort) list * typ + val all: T -> string list + val pretty: theory -> T -> Pretty.T + val make: theory -> string list + -> ((sort -> sort) * Sorts.algebra) * T + val eval_conv: theory + -> (term -> term * (((sort -> sort) * Sorts.algebra) -> T -> thm)) -> cterm -> thm + val eval_term: theory + -> (term -> term * (((sort -> sort) * Sorts.algebra) -> T -> 'a)) -> term -> 'a +end + +structure Code_Wellsorted : CODE_WELLSORTED = +struct + +(** the equation graph type **) + +type T = (((string * sort) list * typ) * (thm * bool) list) Graph.T; + +fun eqns eqngr = these o Option.map snd o try (Graph.get_node eqngr); +fun typ eqngr = fst o Graph.get_node eqngr; +fun all eqngr = Graph.keys eqngr; + +fun pretty thy eqngr = + AList.make (snd o Graph.get_node eqngr) (Graph.keys eqngr) + |> (map o apfst) (Code_Unit.string_of_const thy) + |> sort (string_ord o pairself fst) + |> map (fn (s, thms) => + (Pretty.block o Pretty.fbreaks) ( + Pretty.str s + :: map (Display.pretty_thm o fst) thms + )) + |> Pretty.chunks; + + +(** the Waisenhaus algorithm **) + +(* auxiliary *) + +fun complete_proper_sort thy = + Sign.complete_sort thy #> filter (can (AxClass.get_info thy)); + +fun inst_params thy tyco = + map (fn (c, _) => AxClass.param_of_inst thy (c, tyco)) + o maps (#params o AxClass.get_info thy); + +fun consts_of thy eqns = [] |> (fold o fold o fold_aterms) + (fn Const (c, ty) => insert (op =) (c, Sign.const_typargs thy (c, Logic.unvarifyT ty)) | _ => I) + (map (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of o fst) eqns); + +fun tyscm_rhss_of thy c eqns = + let + val tyscm = case eqns of [] => Code.default_typscheme thy c + | ((thm, _) :: _) => (snd o Code_Unit.head_eqn thy) thm; + val rhss = consts_of thy eqns; + in (tyscm, rhss) end; + + +(* data structures *) + +datatype const = Fun of string | Inst of class * string; + +fun const_ord (Fun c1, Fun c2) = fast_string_ord (c1, c2) + | const_ord (Inst class_tyco1, Inst class_tyco2) = + prod_ord fast_string_ord fast_string_ord (class_tyco1, class_tyco2) + | const_ord (Fun _, Inst _) = LESS + | const_ord (Inst _, Fun _) = GREATER; + +type var = const * int; + +structure Vargraph = + GraphFun(type key = var val ord = prod_ord const_ord int_ord); + +datatype styp = Tyco of string * styp list | Var of var | Free; + +fun styp_of c_lhs (Type (tyco, tys)) = Tyco (tyco, map (styp_of c_lhs) tys) + | styp_of c_lhs (TFree (v, _)) = case c_lhs + of SOME (c, lhs) => Var (Fun c, find_index (fn (v', _) => v = v') lhs) + | NONE => Free; + +type vardeps_data = ((string * styp list) list * class list) Vargraph.T + * (((string * sort) list * (thm * bool) list) Symtab.table + * (class * string) list); + +val empty_vardeps_data : vardeps_data = + (Vargraph.empty, (Symtab.empty, [])); + +(* retrieving equations and instances from the background context *) + +fun obtain_eqns thy eqngr c = + case try (Graph.get_node eqngr) c + of SOME ((lhs, _), eqns) => ((lhs, []), []) + | NONE => let + val eqns = Code.these_eqns thy c + |> burrow_fst (Code_Unit.norm_args thy) + |> burrow_fst (Code_Unit.norm_varnames thy Code_Name.purify_tvar Code_Name.purify_var); + val ((lhs, _), rhss) = tyscm_rhss_of thy c eqns; + in ((lhs, rhss), eqns) end; + +fun obtain_instance thy arities (inst as (class, tyco)) = + case AList.lookup (op =) arities inst + of SOME classess => (classess, ([], [])) + | NONE => let + val all_classes = complete_proper_sort thy [class]; + val superclasses = remove (op =) class all_classes + val classess = map (complete_proper_sort thy) + (Sign.arity_sorts thy tyco [class]); + val inst_params = inst_params thy tyco all_classes; + in (classess, (superclasses, inst_params)) end; + + +(* computing instantiations *) + +fun add_classes thy arities eqngr c_k new_classes vardeps_data = + let + val (styps, old_classes) = Vargraph.get_node (fst vardeps_data) c_k; + val diff_classes = new_classes |> subtract (op =) old_classes; + in if null diff_classes then vardeps_data + else let + val c_ks = Vargraph.imm_succs (fst vardeps_data) c_k |> insert (op =) c_k; + in + vardeps_data + |> (apfst o Vargraph.map_node c_k o apsnd) (append diff_classes) + |> fold (fn styp => fold (assert_typmatch_inst thy arities eqngr styp) new_classes) styps + |> fold (fn c_k => add_classes thy arities eqngr c_k diff_classes) c_ks + end end +and add_styp thy arities eqngr c_k tyco_styps vardeps_data = + let + val (old_styps, classes) = Vargraph.get_node (fst vardeps_data) c_k; + in if member (op =) old_styps tyco_styps then vardeps_data + else + vardeps_data + |> (apfst o Vargraph.map_node c_k o apfst) (cons tyco_styps) + |> fold (assert_typmatch_inst thy arities eqngr tyco_styps) classes + end +and add_dep thy arities eqngr c_k c_k' vardeps_data = + let + val (_, classes) = Vargraph.get_node (fst vardeps_data) c_k; + in + vardeps_data + |> add_classes thy arities eqngr c_k' classes + |> apfst (Vargraph.add_edge (c_k, c_k')) + end +and assert_typmatch_inst thy arities eqngr (tyco, styps) class vardeps_data = + if can (Sign.arity_sorts thy tyco) [class] + then vardeps_data + |> assert_inst thy arities eqngr (class, tyco) + |> fold_index (fn (k, styp) => + assert_typmatch thy arities eqngr styp (Inst (class, tyco), k)) styps + else vardeps_data (*permissive!*) +and assert_inst thy arities eqngr (inst as (class, tyco)) (vardeps_data as (_, (_, insts))) = + if member (op =) insts inst then vardeps_data + else let + val (classess, (superclasses, inst_params)) = + obtain_instance thy arities inst; + in + vardeps_data + |> (apsnd o apsnd) (insert (op =) inst) + |> fold_index (fn (k, _) => + apfst (Vargraph.new_node ((Inst (class, tyco), k), ([] ,[])))) classess + |> fold (fn superclass => assert_inst thy arities eqngr (superclass, tyco)) superclasses + |> fold (assert_fun thy arities eqngr) inst_params + |> fold_index (fn (k, classes) => + add_classes thy arities eqngr (Inst (class, tyco), k) classes + #> fold (fn superclass => + add_dep thy arities eqngr (Inst (superclass, tyco), k) + (Inst (class, tyco), k)) superclasses + #> fold (fn inst_param => + add_dep thy arities eqngr (Fun inst_param, k) + (Inst (class, tyco), k) + ) inst_params + ) classess + end +and assert_typmatch thy arities eqngr (Tyco tyco_styps) c_k vardeps_data = + vardeps_data + |> add_styp thy arities eqngr c_k tyco_styps + | assert_typmatch thy arities eqngr (Var c_k') c_k vardeps_data = + vardeps_data + |> add_dep thy arities eqngr c_k c_k' + | assert_typmatch thy arities eqngr Free c_k vardeps_data = + vardeps_data +and assert_rhs thy arities eqngr (c', styps) vardeps_data = + vardeps_data + |> assert_fun thy arities eqngr c' + |> fold_index (fn (k, styp) => + assert_typmatch thy arities eqngr styp (Fun c', k)) styps +and assert_fun thy arities eqngr c (vardeps_data as (_, (eqntab, _))) = + if Symtab.defined eqntab c then vardeps_data + else let + val ((lhs, rhss), eqns) = obtain_eqns thy eqngr c; + val rhss' = (map o apsnd o map) (styp_of (SOME (c, lhs))) rhss; + in + vardeps_data + |> (apsnd o apfst) (Symtab.update_new (c, (lhs, eqns))) + |> fold_index (fn (k, _) => + apfst (Vargraph.new_node ((Fun c, k), ([] ,[])))) lhs + |> fold_index (fn (k, (_, sort)) => + add_classes thy arities eqngr (Fun c, k) (complete_proper_sort thy sort)) lhs + |> fold (assert_rhs thy arities eqngr) rhss' + end; + + +(* applying instantiations *) + +fun dicts_of thy (proj_sort, algebra) (T, sort) = + let + fun class_relation (x, _) _ = x; + fun type_constructor tyco xs class = + inst_params thy tyco (Sorts.complete_sort algebra [class]) + @ (maps o maps) fst xs; + fun type_variable (TFree (_, sort)) = map (pair []) (proj_sort sort); + in + flat (Sorts.of_sort_derivation (Syntax.pp_global thy) algebra + { class_relation = class_relation, type_constructor = type_constructor, + type_variable = type_variable } (T, proj_sort sort) + handle Sorts.CLASS_ERROR _ => [] (*permissive!*)) + end; + +fun add_arity thy vardeps (class, tyco) = + AList.default (op =) + ((class, tyco), map (fn k => (snd o Vargraph.get_node vardeps) (Inst (class, tyco), k)) + (0 upto Sign.arity_number thy tyco - 1)); + +fun add_eqs thy (proj_sort, algebra) vardeps + (c, (proto_lhs, proto_eqns)) (rhss, eqngr) = + if can (Graph.get_node eqngr) c then (rhss, eqngr) + else let + val lhs = map_index (fn (k, (v, _)) => + (v, snd (Vargraph.get_node vardeps (Fun c, k)))) proto_lhs; + val inst_tab = Vartab.empty |> fold (fn (v, sort) => + Vartab.update ((v, 0), sort)) lhs; + val eqns = proto_eqns + |> (map o apfst) (Code_Unit.inst_thm thy inst_tab); + val (tyscm, rhss') = tyscm_rhss_of thy c eqns; + val eqngr' = Graph.new_node (c, (tyscm, eqns)) eqngr; + in (map (pair c) rhss' @ rhss, eqngr') end; + +fun extend_arities_eqngr thy cs cs_rhss (arities, eqngr) = + let + val cs_rhss' = (map o apsnd o map) (styp_of NONE) cs_rhss; + val (vardeps, (eqntab, insts)) = empty_vardeps_data + |> fold (assert_fun thy arities eqngr) cs + |> fold (assert_rhs thy arities eqngr) cs_rhss'; + val arities' = fold (add_arity thy vardeps) insts arities; + val pp = Syntax.pp_global thy; + val is_proper_class = can (AxClass.get_info thy); + val (proj_sort, algebra) = Sorts.subalgebra pp is_proper_class + (AList.lookup (op =) arities') (Sign.classes_of thy); + val (rhss, eqngr') = Symtab.fold + (add_eqs thy (proj_sort, algebra) vardeps) eqntab ([], eqngr); + fun deps_of (c, rhs) = c :: + maps (dicts_of thy (proj_sort, algebra)) + (rhs ~~ (map snd o fst o fst o Graph.get_node eqngr') c); + val eqngr'' = fold (fn (c, rhs) => fold + (curry Graph.add_edge c) (deps_of rhs)) rhss eqngr'; + in ((proj_sort, algebra), (arities', eqngr'')) end; + + +(** retrieval interfaces **) + +fun proto_eval thy cterm_of evaluator_lift evaluator proto_ct arities_eqngr = + let + val ct = cterm_of proto_ct; + val _ = Sign.no_vars (Syntax.pp_global thy) (Thm.term_of ct); + val _ = Term.fold_types (Type.no_tvars #> K I) (Thm.term_of ct) (); + fun consts_of t = + fold_aterms (fn Const c_ty => cons c_ty | _ => I) t []; + val thm = Code.preprocess_conv thy ct; + val ct' = Thm.rhs_of thm; + val t' = Thm.term_of ct'; + val (t'', evaluator_eqngr) = evaluator t'; + val consts = map fst (consts_of t'); + val consts' = consts_of t''; + val const_matches' = fold (fn (c, ty) => + insert (op =) (c, Sign.const_typargs thy (c, ty))) consts' []; + val (algebra', arities_eqngr') = + extend_arities_eqngr thy consts const_matches' arities_eqngr; + in + (evaluator_lift (evaluator_eqngr algebra') thm (snd arities_eqngr'), + arities_eqngr') + end; + +fun proto_eval_conv thy = + let + fun evaluator_lift evaluator thm1 eqngr = + let + val thm2 = evaluator eqngr; + val thm3 = Code.postprocess_conv thy (Thm.rhs_of thm2); + in + Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ => + error ("could not construct evaluation proof:\n" + ^ (cat_lines o map Display.string_of_thm) [thm1, thm2, thm3]) + end; + in proto_eval thy I evaluator_lift end; + +fun proto_eval_term thy = + let + fun evaluator_lift evaluator _ eqngr = evaluator eqngr; + in proto_eval thy (Thm.cterm_of thy) evaluator_lift end; + +structure Wellsorted = CodeDataFun +( + type T = ((string * class) * sort list) list * T; + val empty = ([], Graph.empty); + fun purge thy cs (arities, eqngr) = + let + val del_cs = ((Graph.all_preds eqngr + o filter (can (Graph.get_node eqngr))) cs); + val del_arities = del_cs + |> map_filter (AxClass.inst_of_param thy) + |> maps (fn (c, tyco) => + (map (rpair tyco) o Sign.complete_sort thy o the_list + o AxClass.class_of_param thy) c); + val arities' = fold (AList.delete (op =)) del_arities arities; + val eqngr' = Graph.del_nodes del_cs eqngr; + in (arities', eqngr') end; +); + +fun make thy cs = apsnd snd + (Wellsorted.change_yield thy (extend_arities_eqngr thy cs [])); + +fun eval_conv thy f = + fst o Wellsorted.change_yield thy o proto_eval_conv thy f; + +fun eval_term thy f = + fst o Wellsorted.change_yield thy o proto_eval_term thy f; + + +(** diagnostic commands **) + +fun code_depgr thy consts = + let + val (_, eqngr) = make thy consts; + val select = Graph.all_succs eqngr consts; + in + eqngr + |> not (null consts) ? Graph.subgraph (member (op =) select) + |> Graph.map_nodes ((apsnd o map o apfst) (AxClass.overload thy)) + end; + +fun code_thms thy = Pretty.writeln o pretty thy o code_depgr thy; + +fun code_deps thy consts = + let + val eqngr = code_depgr thy consts; + fun mk_entry (const, (_, (_, parents))) = + let + val name = Code_Unit.string_of_const thy const; + val nameparents = map (Code_Unit.string_of_const thy) parents; + in { name = name, ID = name, dir = "", unfold = true, + path = "", parents = nameparents } + end; + val prgr = Graph.fold ((fn x => fn xs => xs @ [x]) o mk_entry) eqngr []; + in Present.display_graph prgr end; + +local + +structure P = OuterParse +and K = OuterKeyword + +fun code_thms_cmd thy = code_thms thy o op @ o Code_Name.read_const_exprs thy; +fun code_deps_cmd thy = code_deps thy o op @ o Code_Name.read_const_exprs thy; + +in + +val _ = + OuterSyntax.improper_command "code_thms" "print system of code equations for code" OuterKeyword.diag + (Scan.repeat P.term_group + >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory + o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of))); + +val _ = + OuterSyntax.improper_command "code_deps" "visualize dependencies of code equations for code" OuterKeyword.diag + (Scan.repeat P.term_group + >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory + o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of))); + +end; + +end; (*struct*) diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Tools/coherent.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/coherent.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,233 @@ +(* Title: Tools/coherent.ML + Author: Stefan Berghofer, TU Muenchen + Author: Marc Bezem, Institutt for Informatikk, Universitetet i Bergen + +Prover for coherent logic, see e.g. + + Marc Bezem and Thierry Coquand, Automating Coherent Logic, LPAR 2005 + +for a description of the algorithm. +*) + +signature COHERENT_DATA = +sig + val atomize_elimL: thm + val atomize_exL: thm + val atomize_conjL: thm + val atomize_disjL: thm + val operator_names: string list +end; + +signature COHERENT = +sig + val verbose: bool ref + val show_facts: bool ref + val coherent_tac: thm list -> Proof.context -> int -> tactic + val coherent_meth: thm list -> Proof.context -> Proof.method + val setup: theory -> theory +end; + +functor CoherentFun(Data: COHERENT_DATA) : COHERENT = +struct + +val verbose = ref false; + +fun message f = if !verbose then tracing (f ()) else (); + +datatype cl_prf = + ClPrf of thm * (Type.tyenv * Envir.tenv) * ((indexname * typ) * term) list * + int list * (term list * cl_prf) list; + +val is_atomic = not o exists_Const (member (op =) Data.operator_names o #1); + +local open Conv in + +fun rulify_elim_conv ct = + if is_atomic (Logic.strip_imp_concl (term_of ct)) then all_conv ct + else concl_conv (length (Logic.strip_imp_prems (term_of ct))) + (rewr_conv (symmetric Data.atomize_elimL) then_conv + MetaSimplifier.rewrite true (map symmetric + [Data.atomize_exL, Data.atomize_conjL, Data.atomize_disjL])) ct + +end; + +fun rulify_elim th = MetaSimplifier.norm_hhf (Conv.fconv_rule rulify_elim_conv th); + +(* Decompose elimination rule of the form + A1 ==> ... ==> Am ==> (!!xs1. Bs1 ==> P) ==> ... ==> (!!xsn. Bsn ==> P) ==> P +*) +fun dest_elim prop = + let + val prems = Logic.strip_imp_prems prop; + val concl = Logic.strip_imp_concl prop; + val (prems1, prems2) = + take_suffix (fn t => Logic.strip_assums_concl t = concl) prems; + in + (prems1, + if null prems2 then [([], [concl])] + else map (fn t => + (map snd (Logic.strip_params t), Logic.strip_assums_hyp t)) prems2) + end; + +fun mk_rule th = + let + val th' = rulify_elim th; + val (prems, cases) = dest_elim (prop_of th') + in (th', prems, cases) end; + +fun mk_dom ts = fold (fn t => + Typtab.map_default (fastype_of t, []) (fn us => us @ [t])) ts Typtab.empty; + +val empty_env = (Vartab.empty, Vartab.empty); + +(* Find matcher that makes conjunction valid in given state *) +fun valid_conj ctxt facts env [] = Seq.single (env, []) + | valid_conj ctxt facts env (t :: ts) = + Seq.maps (fn (u, x) => Seq.map (apsnd (cons x)) + (valid_conj ctxt facts + (Pattern.match (ProofContext.theory_of ctxt) (t, u) env) ts + handle Pattern.MATCH => Seq.empty)) + (Seq.of_list (sort (int_ord o pairself snd) (Net.unify_term facts t))); + +(* Instantiate variables that only occur free in conlusion *) +fun inst_extra_vars ctxt dom cs = + let + val vs = fold Term.add_vars (maps snd cs) []; + fun insts [] inst = Seq.single inst + | insts ((ixn, T) :: vs') inst = Seq.maps + (fn t => insts vs' (((ixn, T), t) :: inst)) + (Seq.of_list (case Typtab.lookup dom T of + NONE => error ("Unknown domain: " ^ + Syntax.string_of_typ ctxt T ^ "\nfor term(s) " ^ + commas (maps (map (Syntax.string_of_term ctxt) o snd) cs)) + | SOME ts => ts)) + in Seq.map (fn inst => + (inst, map (apsnd (map (subst_Vars (map (apfst fst) inst)))) cs)) + (insts vs []) + end; + +(* Check whether disjunction is valid in given state *) +fun is_valid_disj ctxt facts [] = false + | is_valid_disj ctxt facts ((Ts, ts) :: ds) = + let val vs = rev (map_index (fn (i, T) => Var (("x", i), T)) Ts) + in case Seq.pull (valid_conj ctxt facts empty_env + (map (fn t => subst_bounds (vs, t)) ts)) of + SOME _ => true + | NONE => is_valid_disj ctxt facts ds + end; + +val show_facts = ref false; + +fun string_of_facts ctxt s facts = space_implode "\n" + (s :: map (Syntax.string_of_term ctxt) + (map fst (sort (int_ord o pairself snd) (Net.content facts)))) ^ "\n\n"; + +fun print_facts ctxt facts = + if !show_facts then message (fn () => string_of_facts ctxt "Facts:" facts) + else (); + +fun valid ctxt rules goal dom facts nfacts nparams = + let val seq = Seq.of_list rules |> Seq.maps (fn (th, ps, cs) => + valid_conj ctxt facts empty_env ps |> Seq.maps (fn (env as (tye, _), is) => + let val cs' = map (fn (Ts, ts) => + (map (Envir.typ_subst_TVars tye) Ts, map (Envir.subst_vars env) ts)) cs + in + inst_extra_vars ctxt dom cs' |> + Seq.map_filter (fn (inst, cs'') => + if is_valid_disj ctxt facts cs'' then NONE + else SOME (th, env, inst, is, cs'')) + end)) + in + case Seq.pull seq of + NONE => (tracing (string_of_facts ctxt "Countermodel found:" facts); NONE) + | SOME ((th, env, inst, is, cs), _) => + if cs = [([], [goal])] then SOME (ClPrf (th, env, inst, is, [])) + else + (case valid_cases ctxt rules goal dom facts nfacts nparams cs of + NONE => NONE + | SOME prfs => SOME (ClPrf (th, env, inst, is, prfs))) + end + +and valid_cases ctxt rules goal dom facts nfacts nparams [] = SOME [] + | valid_cases ctxt rules goal dom facts nfacts nparams ((Ts, ts) :: ds) = + let + val _ = message (fn () => "case " ^ commas (map (Syntax.string_of_term ctxt) ts)); + val params = rev (map_index (fn (i, T) => + Free ("par" ^ string_of_int (nparams + i), T)) Ts); + val ts' = map_index (fn (i, t) => + (subst_bounds (params, t), nfacts + i)) ts; + val dom' = fold (fn (T, p) => + Typtab.map_default (T, []) (fn ps => ps @ [p])) + (Ts ~~ params) dom; + val facts' = fold (fn (t, i) => Net.insert_term op = + (t, (t, i))) ts' facts + in + case valid ctxt rules goal dom' facts' + (nfacts + length ts) (nparams + length Ts) of + NONE => NONE + | SOME prf => (case valid_cases ctxt rules goal dom facts nfacts nparams ds of + NONE => NONE + | SOME prfs => SOME ((params, prf) :: prfs)) + end; + +(** proof replaying **) + +fun thm_of_cl_prf thy goal asms (ClPrf (th, (tye, env), insts, is, prfs)) = + let + val _ = message (fn () => space_implode "\n" + ("asms:" :: map Display.string_of_thm asms) ^ "\n\n"); + val th' = Drule.implies_elim_list + (Thm.instantiate + (map (fn (ixn, (S, T)) => + (Thm.ctyp_of thy (TVar ((ixn, S))), Thm.ctyp_of thy T)) + (Vartab.dest tye), + map (fn (ixn, (T, t)) => + (Thm.cterm_of thy (Var (ixn, Envir.typ_subst_TVars tye T)), + Thm.cterm_of thy t)) (Vartab.dest env) @ + map (fn (ixnT, t) => + (Thm.cterm_of thy (Var ixnT), Thm.cterm_of thy t)) insts) th) + (map (nth asms) is); + val (_, cases) = dest_elim (prop_of th') + in + case (cases, prfs) of + ([([], [_])], []) => th' + | ([([], [_])], [([], prf)]) => thm_of_cl_prf thy goal (asms @ [th']) prf + | _ => Drule.implies_elim_list + (Thm.instantiate (Thm.match + (Drule.strip_imp_concl (cprop_of th'), goal)) th') + (map (thm_of_case_prf thy goal asms) (prfs ~~ cases)) + end + +and thm_of_case_prf thy goal asms ((params, prf), (_, asms')) = + let + val cparams = map (cterm_of thy) params; + val asms'' = map (cterm_of thy o curry subst_bounds (rev params)) asms' + in + Drule.forall_intr_list cparams (Drule.implies_intr_list asms'' + (thm_of_cl_prf thy goal (asms @ map Thm.assume asms'') prf)) + end; + + +(** external interface **) + +fun coherent_tac rules ctxt = SUBPROOF (fn {prems, concl, params, context, ...} => + rtac (rulify_elim_conv concl RS equal_elim_rule2) 1 THEN + SUBPROOF (fn {prems = prems', concl, context, ...} => + let val xs = map term_of params @ + map (fn (_, s) => Free (s, the (Variable.default_type context s))) + (Variable.fixes_of context) + in + case valid context (map mk_rule (prems' @ prems @ rules)) (term_of concl) + (mk_dom xs) Net.empty 0 0 of + NONE => no_tac + | SOME prf => + rtac (thm_of_cl_prf (ProofContext.theory_of context) concl [] prf) 1 + end) context 1) ctxt; + +fun coherent_meth rules ctxt = + Method.METHOD (fn facts => coherent_tac (facts @ rules) ctxt 1); + +val setup = Method.add_method + ("coherent", Method.thms_ctxt_args coherent_meth, "prove coherent formula"); + +end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Tools/eqsubst.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/eqsubst.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,575 @@ +(* Title: Tools/eqsubst.ML + Author: Lucas Dixon, University of Edinburgh + +A proof method to perform a substiution using an equation. +*) + +signature EQSUBST = +sig + (* a type abbreviation for match information *) + type match = + ((indexname * (sort * typ)) list (* type instantiations *) + * (indexname * (typ * term)) list) (* term instantiations *) + * (string * typ) list (* fake named type abs env *) + * (string * typ) list (* type abs env *) + * term (* outer term *) + + type searchinfo = + theory + * int (* maxidx *) + * Zipper.T (* focusterm to search under *) + + exception eqsubst_occL_exp of + string * int list * Thm.thm list * int * Thm.thm + + (* low level substitution functions *) + val apply_subst_in_asm : + int -> + Thm.thm -> + Thm.thm -> + (Thm.cterm list * int * 'a * Thm.thm) * match -> Thm.thm Seq.seq + val apply_subst_in_concl : + int -> + Thm.thm -> + Thm.cterm list * Thm.thm -> + Thm.thm -> match -> Thm.thm Seq.seq + + (* matching/unification within zippers *) + val clean_match_z : + Context.theory -> Term.term -> Zipper.T -> match option + val clean_unify_z : + Context.theory -> int -> Term.term -> Zipper.T -> match Seq.seq + + (* skipping things in seq seq's *) + + (* skipping non-empty sub-sequences but when we reach the end + of the seq, remembering how much we have left to skip. *) + datatype 'a skipseq = SkipMore of int + | SkipSeq of 'a Seq.seq Seq.seq; + + val skip_first_asm_occs_search : + ('a -> 'b -> 'c Seq.seq Seq.seq) -> + 'a -> int -> 'b -> 'c skipseq + val skip_first_occs_search : + int -> ('a -> 'b -> 'c Seq.seq Seq.seq) -> 'a -> 'b -> 'c Seq.seq + val skipto_skipseq : int -> 'a Seq.seq Seq.seq -> 'a skipseq + + (* tactics *) + val eqsubst_asm_tac : + Proof.context -> + int list -> Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq + val eqsubst_asm_tac' : + Proof.context -> + (searchinfo -> int -> Term.term -> match skipseq) -> + int -> Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq + val eqsubst_tac : + Proof.context -> + int list -> (* list of occurences to rewrite, use [0] for any *) + Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq + val eqsubst_tac' : + Proof.context -> (* proof context *) + (searchinfo -> Term.term -> match Seq.seq) (* search function *) + -> Thm.thm (* equation theorem to rewrite with *) + -> int (* subgoal number in goal theorem *) + -> Thm.thm (* goal theorem *) + -> Thm.thm Seq.seq (* rewritten goal theorem *) + + + val fakefree_badbounds : + (string * Term.typ) list -> + Term.term -> + (string * Term.typ) list * (string * Term.typ) list * Term.term + + val mk_foo_match : + (Term.term -> Term.term) -> + ('a * Term.typ) list -> Term.term -> Term.term + + (* preparing substitution *) + val prep_meta_eq : Proof.context -> Thm.thm -> Thm.thm list + val prep_concl_subst : + int -> Thm.thm -> (Thm.cterm list * Thm.thm) * searchinfo + val prep_subst_in_asm : + int -> Thm.thm -> int -> + (Thm.cterm list * int * int * Thm.thm) * searchinfo + val prep_subst_in_asms : + int -> Thm.thm -> + ((Thm.cterm list * int * int * Thm.thm) * searchinfo) list + val prep_zipper_match : + Zipper.T -> Term.term * ((string * Term.typ) list * (string * Term.typ) list * Term.term) + + (* search for substitutions *) + val valid_match_start : Zipper.T -> bool + val search_lr_all : Zipper.T -> Zipper.T Seq.seq + val search_lr_valid : (Zipper.T -> bool) -> Zipper.T -> Zipper.T Seq.seq + val searchf_lr_unify_all : + searchinfo -> Term.term -> match Seq.seq Seq.seq + val searchf_lr_unify_valid : + searchinfo -> Term.term -> match Seq.seq Seq.seq + val searchf_bt_unify_valid : + searchinfo -> Term.term -> match Seq.seq Seq.seq + + (* syntax tools *) + val ith_syntax : Args.T list -> int list * Args.T list + val options_syntax : Args.T list -> bool * Args.T list + + (* Isar level hooks *) + val eqsubst_asm_meth : Proof.context -> int list -> Thm.thm list -> Proof.method + val eqsubst_meth : Proof.context -> int list -> Thm.thm list -> Proof.method + val subst_meth : Method.src -> Proof.context -> Proof.method + val setup : theory -> theory + +end; + +structure EqSubst +: EQSUBST += struct + +structure Z = Zipper; + +(* changes object "=" to meta "==" which prepares a given rewrite rule *) +fun prep_meta_eq ctxt = + let val (_, {mk_rews = {mk, ...}, ...}) = Simplifier.rep_ss (Simplifier.local_simpset_of ctxt) + in mk #> map Drule.zero_var_indexes end; + + + (* a type abriviation for match information *) + type match = + ((indexname * (sort * typ)) list (* type instantiations *) + * (indexname * (typ * term)) list) (* term instantiations *) + * (string * typ) list (* fake named type abs env *) + * (string * typ) list (* type abs env *) + * term (* outer term *) + + type searchinfo = + theory + * int (* maxidx *) + * Zipper.T (* focusterm to search under *) + + +(* skipping non-empty sub-sequences but when we reach the end + of the seq, remembering how much we have left to skip. *) +datatype 'a skipseq = SkipMore of int + | SkipSeq of 'a Seq.seq Seq.seq; +(* given a seqseq, skip the first m non-empty seq's, note deficit *) +fun skipto_skipseq m s = + let + fun skip_occs n sq = + case Seq.pull sq of + NONE => SkipMore n + | SOME (h,t) => + (case Seq.pull h of NONE => skip_occs n t + | SOME _ => if n <= 1 then SkipSeq (Seq.cons h t) + else skip_occs (n - 1) t) + in (skip_occs m s) end; + +(* note: outerterm is the taget with the match replaced by a bound + variable : ie: "P lhs" beocmes "%x. P x" + insts is the types of instantiations of vars in lhs + and typinsts is the type instantiations of types in the lhs + Note: Final rule is the rule lifted into the ontext of the + taget thm. *) +fun mk_foo_match mkuptermfunc Ts t = + let + val ty = Term.type_of t + val bigtype = (rev (map snd Ts)) ---> ty + fun mk_foo 0 t = t + | mk_foo i t = mk_foo (i - 1) (t $ (Bound (i - 1))) + val num_of_bnds = (length Ts) + (* foo_term = "fooabs y0 ... yn" where y's are local bounds *) + val foo_term = mk_foo num_of_bnds (Bound num_of_bnds) + in Abs("fooabs", bigtype, mkuptermfunc foo_term) end; + +(* T is outer bound vars, n is number of locally bound vars *) +(* THINK: is order of Ts correct...? or reversed? *) +fun fakefree_badbounds Ts t = + let val (FakeTs,Ts,newnames) = + List.foldr (fn ((n,ty),(FakeTs,Ts,usednames)) => + let val newname = Name.variant usednames n + in ((RWTools.mk_fake_bound_name newname,ty)::FakeTs, + (newname,ty)::Ts, + newname::usednames) end) + ([],[],[]) + Ts + in (FakeTs, Ts, Term.subst_bounds (map Free FakeTs, t)) end; + +(* before matching we need to fake the bound vars that are missing an +abstraction. In this function we additionally construct the +abstraction environment, and an outer context term (with the focus +abstracted out) for use in rewriting with RWInst.rw *) +fun prep_zipper_match z = + let + val t = Z.trm z + val c = Z.ctxt z + val Ts = Z.C.nty_ctxt c + val (FakeTs', Ts', t') = fakefree_badbounds Ts t + val absterm = mk_foo_match (Z.C.apply c) Ts' t' + in + (t', (FakeTs', Ts', absterm)) + end; + +(* Matching and Unification with exception handled *) +fun clean_match thy (a as (pat, t)) = + let val (tyenv, tenv) = Pattern.match thy a (Vartab.empty, Vartab.empty) + in SOME (Vartab.dest tyenv, Vartab.dest tenv) + end handle Pattern.MATCH => NONE; + +(* given theory, max var index, pat, tgt; returns Seq of instantiations *) +fun clean_unify thry ix (a as (pat, tgt)) = + let + (* type info will be re-derived, maybe this can be cached + for efficiency? *) + val pat_ty = Term.type_of pat; + val tgt_ty = Term.type_of tgt; + (* is it OK to ignore the type instantiation info? + or should I be using it? *) + val typs_unify = + SOME (Sign.typ_unify thry (pat_ty, tgt_ty) (Vartab.empty, ix)) + handle Type.TUNIFY => NONE; + in + case typs_unify of + SOME (typinsttab, ix2) => + let + (* is it right to throw away the flexes? + or should I be using them somehow? *) + fun mk_insts env = + (Vartab.dest (Envir.type_env env), + Envir.alist_of env); + val initenv = Envir.Envir {asol = Vartab.empty, + iTs = typinsttab, maxidx = ix2}; + val useq = Unify.smash_unifiers thry [a] initenv + handle UnequalLengths => Seq.empty + | Term.TERM _ => Seq.empty; + fun clean_unify' useq () = + (case (Seq.pull useq) of + NONE => NONE + | SOME (h,t) => SOME (mk_insts h, Seq.make (clean_unify' t))) + handle UnequalLengths => NONE + | Term.TERM _ => NONE + in + (Seq.make (clean_unify' useq)) + end + | NONE => Seq.empty + end; + +(* Matching and Unification for zippers *) +(* Note: Ts is a modified version of the original names of the outer +bound variables. New names have been introduced to make sure they are +unique w.r.t all names in the term and each other. usednames' is +oldnames + new names. *) +fun clean_match_z thy pat z = + let val (t, (FakeTs,Ts,absterm)) = prep_zipper_match z in + case clean_match thy (pat, t) of + NONE => NONE + | SOME insts => SOME (insts, FakeTs, Ts, absterm) end; +(* ix = max var index *) +fun clean_unify_z sgn ix pat z = + let val (t, (FakeTs, Ts,absterm)) = prep_zipper_match z in + Seq.map (fn insts => (insts, FakeTs, Ts, absterm)) + (clean_unify sgn ix (t, pat)) end; + + +(* FOR DEBUGGING... +type trace_subst_errT = int (* subgoal *) + * thm (* thm with all goals *) + * (Thm.cterm list (* certified free var placeholders for vars *) + * thm) (* trivial thm of goal concl *) + (* possible matches/unifiers *) + * thm (* rule *) + * (((indexname * typ) list (* type instantiations *) + * (indexname * term) list ) (* term instantiations *) + * (string * typ) list (* Type abs env *) + * term) (* outer term *); + +val trace_subst_err = (ref NONE : trace_subst_errT option ref); +val trace_subst_search = ref false; +exception trace_subst_exp of trace_subst_errT; +*) + + +fun bot_left_leaf_of (l $ r) = bot_left_leaf_of l + | bot_left_leaf_of (Abs(s,ty,t)) = bot_left_leaf_of t + | bot_left_leaf_of x = x; + +(* Avoid considering replacing terms which have a var at the head as + they always succeed trivially, and uninterestingly. *) +fun valid_match_start z = + (case bot_left_leaf_of (Z.trm z) of + Var _ => false + | _ => true); + +(* search from top, left to right, then down *) +val search_lr_all = ZipperSearch.all_bl_ur; + +(* search from top, left to right, then down *) +fun search_lr_valid validf = + let + fun sf_valid_td_lr z = + let val here = if validf z then [Z.Here z] else [] in + case Z.trm z + of _ $ _ => [Z.LookIn (Z.move_down_left z)] + @ here + @ [Z.LookIn (Z.move_down_right z)] + | Abs _ => here @ [Z.LookIn (Z.move_down_abs z)] + | _ => here + end; + in Z.lzy_search sf_valid_td_lr end; + +(* search from bottom to top, left to right *) + +fun search_bt_valid validf = + let + fun sf_valid_td_lr z = + let val here = if validf z then [Z.Here z] else [] in + case Z.trm z + of _ $ _ => [Z.LookIn (Z.move_down_left z), + Z.LookIn (Z.move_down_right z)] @ here + | Abs _ => [Z.LookIn (Z.move_down_abs z)] @ here + | _ => here + end; + in Z.lzy_search sf_valid_td_lr end; + +fun searchf_unify_gen f (sgn, maxidx, z) lhs = + Seq.map (clean_unify_z sgn maxidx lhs) + (Z.limit_apply f z); + +(* search all unifications *) +val searchf_lr_unify_all = + searchf_unify_gen search_lr_all; + +(* search only for 'valid' unifiers (non abs subterms and non vars) *) +val searchf_lr_unify_valid = + searchf_unify_gen (search_lr_valid valid_match_start); + +val searchf_bt_unify_valid = + searchf_unify_gen (search_bt_valid valid_match_start); + +(* apply a substitution in the conclusion of the theorem th *) +(* cfvs are certified free var placeholders for goal params *) +(* conclthm is a theorem of for just the conclusion *) +(* m is instantiation/match information *) +(* rule is the equation for substitution *) +fun apply_subst_in_concl i th (cfvs, conclthm) rule m = + (RWInst.rw m rule conclthm) + |> IsaND.unfix_frees cfvs + |> RWInst.beta_eta_contract + |> (fn r => Tactic.rtac r i th); + +(* substitute within the conclusion of goal i of gth, using a meta +equation rule. Note that we assume rule has var indicies zero'd *) +fun prep_concl_subst i gth = + let + val th = Thm.incr_indexes 1 gth; + val tgt_term = Thm.prop_of th; + + val sgn = Thm.theory_of_thm th; + val ctermify = Thm.cterm_of sgn; + val trivify = Thm.trivial o ctermify; + + val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term; + val cfvs = rev (map ctermify fvs); + + val conclterm = Logic.strip_imp_concl fixedbody; + val conclthm = trivify conclterm; + val maxidx = Thm.maxidx_of th; + val ft = ((Z.move_down_right (* ==> *) + o Z.move_down_left (* Trueprop *) + o Z.mktop + o Thm.prop_of) conclthm) + in + ((cfvs, conclthm), (sgn, maxidx, ft)) + end; + +(* substitute using an object or meta level equality *) +fun eqsubst_tac' ctxt searchf instepthm i th = + let + val (cvfsconclthm, searchinfo) = prep_concl_subst i th; + val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm); + fun rewrite_with_thm r = + let val (lhs,_) = Logic.dest_equals (Thm.concl_of r); + in searchf searchinfo lhs + |> Seq.maps (apply_subst_in_concl i th cvfsconclthm r) end; + in stepthms |> Seq.maps rewrite_with_thm end; + + +(* distinct subgoals *) +fun distinct_subgoals th = + the_default th (SINGLE distinct_subgoals_tac th); + +(* General substitution of multiple occurances using one of + the given theorems*) + + +exception eqsubst_occL_exp of + string * (int list) * (thm list) * int * thm; +fun skip_first_occs_search occ srchf sinfo lhs = + case (skipto_skipseq occ (srchf sinfo lhs)) of + SkipMore _ => Seq.empty + | SkipSeq ss => Seq.flat ss; + +(* The occL is a list of integers indicating which occurence +w.r.t. the search order, to rewrite. Backtracking will also find later +occurences, but all earlier ones are skipped. Thus you can use [0] to +just find all rewrites. *) + +fun eqsubst_tac ctxt occL thms i th = + let val nprems = Thm.nprems_of th in + if nprems < i then Seq.empty else + let val thmseq = (Seq.of_list thms) + fun apply_occ occ th = + thmseq |> Seq.maps + (fn r => eqsubst_tac' + ctxt + (skip_first_occs_search + occ searchf_lr_unify_valid) r + (i + ((Thm.nprems_of th) - nprems)) + th); + val sortedoccL = + Library.sort (Library.rev_order o Library.int_ord) occL; + in + Seq.map distinct_subgoals (Seq.EVERY (map apply_occ sortedoccL) th) + end + end + handle THM _ => raise eqsubst_occL_exp ("THM",occL,thms,i,th); + + +(* inthms are the given arguments in Isar, and treated as eqstep with + the first one, then the second etc *) +fun eqsubst_meth ctxt occL inthms = + Method.SIMPLE_METHOD' (eqsubst_tac ctxt occL inthms); + +(* apply a substitution inside assumption j, keeps asm in the same place *) +fun apply_subst_in_asm i th rule ((cfvs, j, ngoalprems, pth),m) = + let + val th2 = Thm.rotate_rule (j - 1) i th; (* put premice first *) + val preelimrule = + (RWInst.rw m rule pth) + |> (Seq.hd o prune_params_tac) + |> Thm.permute_prems 0 ~1 (* put old asm first *) + |> IsaND.unfix_frees cfvs (* unfix any global params *) + |> RWInst.beta_eta_contract; (* normal form *) + (* val elimrule = + preelimrule + |> Tactic.make_elim (* make into elim rule *) + |> Thm.lift_rule (th2, i); (* lift into context *) + *) + in + (* ~j because new asm starts at back, thus we subtract 1 *) + Seq.map (Thm.rotate_rule (~j) ((Thm.nprems_of rule) + i)) + (Tactic.dtac preelimrule i th2) + + (* (Thm.bicompose + false (* use unification *) + (true, (* elim resolution *) + elimrule, (2 + (Thm.nprems_of rule)) - ngoalprems) + i th2) *) + end; + + +(* prepare to substitute within the j'th premise of subgoal i of gth, +using a meta-level equation. Note that we assume rule has var indicies +zero'd. Note that we also assume that premt is the j'th premice of +subgoal i of gth. Note the repetition of work done for each +assumption, i.e. this can be made more efficient for search over +multiple assumptions. *) +fun prep_subst_in_asm i gth j = + let + val th = Thm.incr_indexes 1 gth; + val tgt_term = Thm.prop_of th; + + val sgn = Thm.theory_of_thm th; + val ctermify = Thm.cterm_of sgn; + val trivify = Thm.trivial o ctermify; + + val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term; + val cfvs = rev (map ctermify fvs); + + val asmt = nth (Logic.strip_imp_prems fixedbody) (j - 1); + val asm_nprems = length (Logic.strip_imp_prems asmt); + + val pth = trivify asmt; + val maxidx = Thm.maxidx_of th; + + val ft = ((Z.move_down_right (* trueprop *) + o Z.mktop + o Thm.prop_of) pth) + in ((cfvs, j, asm_nprems, pth), (sgn, maxidx, ft)) end; + +(* prepare subst in every possible assumption *) +fun prep_subst_in_asms i gth = + map (prep_subst_in_asm i gth) + ((fn l => Library.upto (1, length l)) + (Logic.prems_of_goal (Thm.prop_of gth) i)); + + +(* substitute in an assumption using an object or meta level equality *) +fun eqsubst_asm_tac' ctxt searchf skipocc instepthm i th = + let + val asmpreps = prep_subst_in_asms i th; + val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm); + fun rewrite_with_thm r = + let val (lhs,_) = Logic.dest_equals (Thm.concl_of r) + fun occ_search occ [] = Seq.empty + | occ_search occ ((asminfo, searchinfo)::moreasms) = + (case searchf searchinfo occ lhs of + SkipMore i => occ_search i moreasms + | SkipSeq ss => + Seq.append (Seq.map (Library.pair asminfo) (Seq.flat ss)) + (occ_search 1 moreasms)) + (* find later substs also *) + in + occ_search skipocc asmpreps |> Seq.maps (apply_subst_in_asm i th r) + end; + in stepthms |> Seq.maps rewrite_with_thm end; + + +fun skip_first_asm_occs_search searchf sinfo occ lhs = + skipto_skipseq occ (searchf sinfo lhs); + +fun eqsubst_asm_tac ctxt occL thms i th = + let val nprems = Thm.nprems_of th + in + if nprems < i then Seq.empty else + let val thmseq = (Seq.of_list thms) + fun apply_occ occK th = + thmseq |> Seq.maps + (fn r => + eqsubst_asm_tac' ctxt (skip_first_asm_occs_search + searchf_lr_unify_valid) occK r + (i + ((Thm.nprems_of th) - nprems)) + th); + val sortedoccs = + Library.sort (Library.rev_order o Library.int_ord) occL + in + Seq.map distinct_subgoals + (Seq.EVERY (map apply_occ sortedoccs) th) + end + end + handle THM _ => raise eqsubst_occL_exp ("THM",occL,thms,i,th); + +(* inthms are the given arguments in Isar, and treated as eqstep with + the first one, then the second etc *) +fun eqsubst_asm_meth ctxt occL inthms = + Method.SIMPLE_METHOD' (eqsubst_asm_tac ctxt occL inthms); + +(* syntax for options, given "(asm)" will give back true, without + gives back false *) +val options_syntax = + (Args.parens (Args.$$$ "asm") >> (K true)) || + (Scan.succeed false); + +val ith_syntax = + Scan.optional (Args.parens (Scan.repeat OuterParse.nat)) [0]; + +(* combination method that takes a flag (true indicates that subst +should be done to an assumption, false = apply to the conclusion of +the goal) as well as the theorems to use *) +fun subst_meth src = + Method.syntax ((Scan.lift options_syntax) -- (Scan.lift ith_syntax) -- Attrib.thms) src + #> (fn (((asmflag, occL), inthms), ctxt) => + (if asmflag then eqsubst_asm_meth else eqsubst_meth) ctxt occL inthms); + + +val setup = + Method.add_method ("subst", subst_meth, "single-step substitution"); + +end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Tools/intuitionistic.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/intuitionistic.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,100 @@ +(* Title: Tools/intuitionistic.ML + Author: Stefan Berghofer, TU Muenchen + +Simple intuitionistic proof search. +*) + +signature INTUITIONISTIC = +sig + val prover_tac: Proof.context -> int option -> int -> tactic + val method_setup: bstring -> theory -> theory +end; + +structure Intuitionistic: INTUITIONISTIC = +struct + +(* main tactic *) + +local + +val remdups_tac = SUBGOAL (fn (g, i) => + let val prems = Logic.strip_assums_hyp g in + REPEAT_DETERM_N (length prems - length (distinct op aconv prems)) + (Tactic.ematch_tac [Drule.remdups_rl] i THEN Tactic.eq_assume_tac i) + end); + +fun REMDUPS tac = tac THEN_ALL_NEW remdups_tac; + +val bires_tac = Tactic.biresolution_from_nets_tac ContextRules.orderlist; + +fun safe_step_tac ctxt = + ContextRules.Swrap ctxt + (eq_assume_tac ORELSE' + bires_tac true (ContextRules.netpair_bang ctxt)); + +fun unsafe_step_tac ctxt = + ContextRules.wrap ctxt + (assume_tac APPEND' + bires_tac false (ContextRules.netpair_bang ctxt) APPEND' + bires_tac false (ContextRules.netpair ctxt)); + +fun step_tac ctxt i = + REPEAT_DETERM1 (REMDUPS (safe_step_tac ctxt) i) ORELSE + REMDUPS (unsafe_step_tac ctxt) i; + +fun intprover_tac ctxt gs d lim = SUBGOAL (fn (g, i) => if d > lim then no_tac else + let + val ps = Logic.strip_assums_hyp g; + val c = Logic.strip_assums_concl g; + in + if member (fn ((ps1, c1), (ps2, c2)) => + c1 aconv c2 andalso + length ps1 = length ps2 andalso + gen_eq_set (op aconv) (ps1, ps2)) gs (ps, c) then no_tac + else (step_tac ctxt THEN_ALL_NEW intprover_tac ctxt ((ps, c) :: gs) (d + 1) lim) i + end); + +in + +fun prover_tac ctxt opt_lim = + SELECT_GOAL (DEEPEN (2, the_default 20 opt_lim) (intprover_tac ctxt [] 0) 4 1); + +end; + + +(* method setup *) + +local + +val introN = "intro"; +val elimN = "elim"; +val destN = "dest"; +val ruleN = "rule"; + +fun modifier name kind kind' att = + Args.$$$ name |-- (kind >> K NONE || kind' |-- OuterParse.nat --| Args.colon >> SOME) + >> (pair (I: Proof.context -> Proof.context) o att); + +val modifiers = + [modifier destN Args.bang_colon Args.bang ContextRules.dest_bang, + modifier destN Args.colon (Scan.succeed ()) ContextRules.dest, + modifier elimN Args.bang_colon Args.bang ContextRules.elim_bang, + modifier elimN Args.colon (Scan.succeed ()) ContextRules.elim, + modifier introN Args.bang_colon Args.bang ContextRules.intro_bang, + modifier introN Args.colon (Scan.succeed ()) ContextRules.intro, + Args.del -- Args.colon >> K (I, ContextRules.rule_del)]; + +val method = + Method.bang_sectioned_args' modifiers (Scan.lift (Scan.option OuterParse.nat)) + (fn n => fn prems => fn ctxt => Method.METHOD (fn facts => + HEADGOAL (Method.insert_tac (prems @ facts) THEN' + ObjectLogic.atomize_prems_tac THEN' prover_tac ctxt n))); + +in + +fun method_setup name = Method.add_method (name, method, "intuitionistic proof search"); + +end; + +end; + diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/Tools/project_rule.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/project_rule.ML Wed Mar 04 11:05:29 2009 +0100 @@ -0,0 +1,65 @@ +(* Title: Tools/project_rule.ML + Author: Makarius + +Transform mutual rule: + + HH ==> (x1:A1 --> P1 x1) & ... & (xn:An --> Pn xn) + +into projection: + + xi:Ai ==> HH ==> Pi xi +*) + +signature PROJECT_RULE_DATA = +sig + val conjunct1: thm + val conjunct2: thm + val mp: thm +end; + +signature PROJECT_RULE = +sig + val project: Proof.context -> int -> thm -> thm + val projects: Proof.context -> int list -> thm -> thm list + val projections: Proof.context -> thm -> thm list +end; + +functor ProjectRuleFun(Data: PROJECT_RULE_DATA): PROJECT_RULE = +struct + +fun conj1 th = th RS Data.conjunct1; +fun conj2 th = th RS Data.conjunct2; +fun imp th = th RS Data.mp; + +fun projects ctxt is raw_rule = + let + fun proj 1 th = the_default th (try conj1 th) + | proj k th = proj (k - 1) (conj2 th); + fun prems k th = + (case try imp th of + NONE => (k, th) + | SOME th' => prems (k + 1) th'); + val ((_, [rule]), ctxt') = Variable.import_thms true [raw_rule] ctxt; + fun result i = + rule + |> proj i + |> prems 0 |-> (fn k => + Thm.permute_prems 0 (~ k) + #> singleton (Variable.export ctxt' ctxt) + #> Drule.zero_var_indexes + #> RuleCases.save raw_rule + #> RuleCases.add_consumes k); + in map result is end; + +fun project ctxt i th = hd (projects ctxt [i] th); + +fun projections ctxt raw_rule = + let + fun projs k th = + (case try conj2 th of + NONE => k + | SOME th' => projs (k + 1) th'); + val ((_, [rule]), _) = Variable.import_thms true [raw_rule] ctxt; + in projects ctxt (1 upto projs 1 rule) raw_rule end; + +end; diff -r 3a1aef73b2b2 -r aea5d7fa7ef5 src/ZF/Tools/inductive_package.ML