# HG changeset patch # User clasohm # Date 748174838 -7200 # Node ID a5a9c433f639f25e5e02c0d090104cacdf113e6f Initial revision diff -r 000000000000 -r a5a9c433f639 CHANGES-92f.txt --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/CHANGES-92f.txt Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,40 @@ +**** Isabelle-92f : a faster version of Isabelle-92 **** + +Isabelle now runs faster through a combination of improvements: pattern +unification, discrimination nets and removal of assumptions during +simplification. Classical reasoning (e.g. fast_tac) runs up to 30% faster +when large numbers of rules are involved. Rewriting (e.g. SIMP_TAC) runs +up to 3 times faster for large subgoals. + +The new version will not benefit everybody; unless you require greater +speed, it may be best to stay with the existing version. The new changes +have not been documented properly, and there are a few incompatibilities. + +THE SPEEDUPS + +Pattern unification is completely invisible to users. It efficiently +handles a common case of higher-order unification. + +Discrimination nets replace the old stringtrees. They provide fast lookup +in a large set of rules for matching or unification. New "net" tactics +replace the "compat_..." tactics based on stringtrees. Tactics +biresolve_from_nets_tac, bimatch_from_nets_tac, resolve_from_net_tac and +match_from_net_tac take a net, rather than a list of rules, and perform +resolution or matching. Tactics net_biresolve_tac, net_bimatch_tac +net_resolve_tac and net_match_tac take a list of rules, build a net +(internally) and perform resolution or matching. + +The tactical METAHYPS, which allows a subgoal's hypotheses to be taken as a +list of theorems, has been extended to handle unknowns (although not type +unknowns). The simplification tactics now use METAHYPS to economise on +storage consumption, and to avoid problems involving "parameters" bound in +a subgoal. The modified simplifier now requires the auto_tac to take an +extra argument: a list of theorems, which represents the assumptions of the +current subgoal. + +OTHER CHANGES + +Apart from minor improvements in Pure Isabelle, the main other changes are +extensions to object-logics. HOL now contains a treatment of co-induction +and co-recursion, while ZF contains a formalization of equivalence classes, +the integers and binary arithmetic. None of this material is documented. diff -r 000000000000 -r a5a9c433f639 COPYRIGHT --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/COPYRIGHT Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,21 @@ +ISABELLE COPYRIGHT NOTICE, LICENCE AND DISCLAIMER. + +Copyright (C) 1992 by the University of Cambridge, Cambridge, England. + +Permission to use, copy, modify, and distribute this software and its +documentation for any non-commercial purpose and without fee is hereby +granted, provided that the above copyright notice appears in all copies and +that both the copyright notice and this permission notice and warranty +disclaimer appear in supporting documentation, and that the name of the +University of Cambridge not be used in advertising or publicity pertaining +to distribution of the software without specific, written prior permission. + +The University of Cambridge disclaims all warranties with regard to this +software, including all implied warranties of merchantability and fitness. +In no event shall the University of Cambridge be liable for any special, +indirect or consequential damages or any damages whatsoever resulting from +loss of use, data or profits, whether in an action of contract, negligence +or other tortious action, arising out of or in connection with the use or +performance of this software. + + diff -r 000000000000 -r a5a9c433f639 EMAILDIST-README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/EMAILDIST-README Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,37 @@ +ISABELLE -- INSTRUCTIONS FOR UNPACKING THE EMAIL DISTRIBUTION + +The Isabelle email distribution consists of about 8 installments, each +small enough to send by electronic mail. The files are called Isabelle.aa, +Isabelle.ab, .... They have been generated by tar, compress, uuencode, and +split, and are packed for email using shar. To unpack the files, perform +the following steps: + +STEP 1. Create a new directory to hold Isabelle and move to that +directory (the name of the directory does not matter): + + mkdir Isabelle; cd Isabelle + +STEP 2. Put each message into a separate file and pipe it through unshar. +(If you don't have unshar, remove the header lines generated by the mail +system and submit the file to sh.) + +STEP 3. Concatenate the files into one file using the command + + cat Isabelle.?? > 92.tar.Z.uu + +STEP 4. Undo the uuencode operation using the command + + uudecode 92.tar.Z.uu + +STEP 5. You should now have a file 92.tar.Z; uncompress and unpack it using... + + uncompress -c 92.tar.Z | tar xf - + +STEP 6. You should now have a complete Isabelle directory, called 92. You +may now tidy up by executing + + rm Isabelle.?? *.hdr 92.tar.Z.uu 92.tar.Z + +Consult the file 92/README for information on compiling Isabelle. + + Good luck! diff -r 000000000000 -r a5a9c433f639 README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/README Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,113 @@ + ISABELLE-92 DISTRIBUTION DIRECTORY + +------------------------------------------------------------------------------ +ISABELLE-92 IS INCOMPATIBLE WITH EARLIER VERSIONS. PLEASE CONSULT THE +DOCUMENTATION. +------------------------------------------------------------------------------ + +This directory contains the complete Isabelle system. To build and test the +entire system, including all object-logics, use the shell script make-all. +Pure Isabelle and each of the object-logics can be built separately using the +Makefiles in the respective directories; read them for more information. + + THE MAKEFILES + +The Makefiles can use two different Standard ML compilers: Poly/ML version +1.88MSX or later (from Abstract Hardware Ltd) and Standard ML of New Jersey +(Version 75 or later). Poly/ML is a commercial product and costs money, +but it is reliable and its database system is convenient for interactive +work. SML of New Jersey requires lots of memory and disc space, but it is +free and its code sometimes runs faster. Both compilers are perfectly +satisfactory for running Isabelle. + +The Makefiles and make-all use enviroment variables that you should set +according to your site configuration. + +ISABELLEBIN is the directory to hold Poly/ML databases or New Jersey ML +images. When using Poly/ML, ISABELLEBIN must be an absolute pathname (one +starting with "/"). + +ML_DBASE is an absolute pathname to the initial Poly/ML database (not +required for New Jersey ML). + +ISABELLECOMP is the ML compiler, typically "poly -noDisplay" or "sml". If +ISABELLECOMP begins with the letters "poly" then the Makefiles assume that +it is Poly/ML; if it begins with the letters "sml" then they assume +Standard ML of New Jersey. + + + STRUCTURE OF THIS DIRECTORY + +The directory Pure containes pure Isabelle, which has no object-logic. + +Other important files include... + COPYRIGHT Copyright notice and Disclaimer of Warranty + make-rulenames shell script used during Make + make-all shell script for building entire system + expandshort shell script to expand "shortcuts" in files + prove_goal.el Emacs command to change proof format + xlisten shell script for running Isabelle under X + teeinput shell script to run Isabelle, logging inputs to a file + theory-template.ML template file for defining new theories + Pure directory of source files for Pure Isabelle + Provers directory of generic theorem provers + +xlisten sets up a window running Isabelle, with a separate small "listener" +window, which keeps a log of all input lines. This log is a useful record +of a session. If you are not running X windows, teeinput can still be used at +least to record (if not to display) the log. + +The following subdirectories contain object-logics: + FOL Natural deduction logic (intuitionistic and classical) + ZF Zermelo-Fraenkel Set theory + CTT Constructive Type Theory + HOL Classical Higher-Order Logic + LK Classical sequent calculus + Modal The modal logics T, S4, S43 + LCF Logic for Computable Functions (domain theory) + Cube Barendregt's Lambda Cube + +Object-logics include examples files in subdirectory ex or file ex.ML. +These files can be loaded in batch mode. The commands can also be +executed interactively, using the windows on your workstation. This is a +good way to get started. + +Each object-logic is built on top of Pure Isabelle, and possibly on top of +another object logic (like FOL or LK). A database or binary called Pure is +first created, then the object-logic is loaded on top. Poly/ML extends +Pure using its "make_database" operation. Standard ML of New Jersey starts +with the Pure core image and loads the object-logic's ROOT.ML. + + HOW TO GET A STANDARD ML COMPILER + +To obtain Poly/ML, contact Mike Crawley at Abstract +Hardware Ltd, The Howell Building, Brunel University, Uxbridge UB8 3PH, +England. + +To obtain Standard ML of New Jersey, contact David MacQueen + at AT&T Bell Laboratories, 600 Mountain Avenue, +Murray Hill, NJ 07974, USA. This compiler is available by FTP. Connect to +research.att.com; login as anonymous with your userid as password; set +binary mode; transfer files from the directory dist/ml. + +------------------------------------------------------------------------------ + +Please report any problems you encounter. While we will try to be helpful, +we can accept no responsibility for the deficiences of Isabelle amd their +consequences. + +Lawrence C Paulson E-mail: lcp@cl.cam.ac.uk +Computer Laboratory Phone: +44-223-334600 +University of Cambridge Fax: +44-223-334748 +Pembroke Street +Cambridge CB2 3QG +England + +Tobias Nipkow E-mail: nipkow@informatik.tu-muenchen.de +Institut fuer Informatik Phone: +49-89-2105-2690 +T. U. Muenchen Fax: +49-89-2105-8183 +Postfach 20 24 20 +D-8000 Muenchen 2 +Germany + +Last updated 25 August 1992 diff -r 000000000000 -r a5a9c433f639 agrep --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/agrep Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,2 @@ +#! /bin/csh +grep "$*" {Pure/Syntax,Pure/Thy}/*ML */*ML */ex/*ML diff -r 000000000000 -r a5a9c433f639 edits.txt --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/edits.txt Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,1372 @@ +EDITS TO THE ISABELLE SYSTEM FOR 1993 + +11 January + +*/README: Eliminated references to Makefile.NJ, which no longer exists. + +**** New tar file placed on /homes/lcp (464K) **** + +14 January + +Provers/simp/pr_goal_lhs: now distinct from pr_goal_concl so that tracing +prints conditions correctly. + +{CTT/arith,HOL/ex/arith/ZF/arith}/add_mult_distrib: renamed from +add_mult_dist, to agree with the other _distrib rules + +20 January + +Pure/Syntax/type_ext.ML: "I have fixed a few anomalies in the pretty +printing annotations for types. Only the layout has changed." -- Toby + +21 January + +{CTT/arith,HOL/ex/arith/ZF/arith}/add_inverse_diff: renamed to add_diff_inverse + +22 January + +ZF/ex/equiv: new theory of equivalence classes +ZF/ex/integ: new theory of integers +HOL/set.thy: added indentation of 3 to all binding operators + +ZF/bool/boolI0,boolI1: renamed as bool_0I, bool_1I + +25 January + +MAKE-ALL (NJ 0.75) ran perfectly. It took 3:19 hours!? + +ZF/bool/not,and,or,xor: new + +27 January + +ZF/ex/bin: new theory of binary integer arithmetic + +27 January + +MAKE-ALL (Poly/ML) ran perfectly. It took 6:33 hours??? +(ZF took almost 5 hours!) + +**** New tar file placed on /homes/lcp (472K) **** + +HOL/set/UN_cong,INT_cong: new +HOL/subset/mem_rews,set_congs,set_ss: new +HOL/simpdata/o_apply: new; added to HOL_ss + +29 January + +Pure/Thy/syntax/mk_structure: the dummy theory created by type infixes is +now called name^"(type infix)" instead of "", avoid triggering a spurious +error "Attempt to merge different versions of theory: " in +Pure/sign/merge_stamps + +2 February + +MAKE-ALL (Poly/ML) ran perfectly. It took 2:48 hours. Runs in 1992 took +under 2:20 hours, but the new files in ZF/ex take time: nearly 23 minutes +according to make10836.log. + +Pure/Thy/scan/comment: renamed from komt +Pure/Thy/scan/numeric: renamed from zahl + +Pure/Syntax/syntax,lexicon,type_ext,extension,sextension: modified by +Tobias to change ID, TVAR, ... to lower case. + +Cube/cube.thy,HOL/hol.thy,HOL/set.thy,CTT/ctt.thy,LK/lk.thy,ZF/zf.thy: now +with ID, ... in lower case and other tidying + +3 February + +MAKE-ALL (Poly/ML) ran perfectly. It took 2:50 hours. + +4 February + +HOL/nat/nat_ss: now includes the rule Suc_less_eq: (Suc(m) < Suc(n)) = (m P(s)"; + by (res_inst_tac [("t","s")] (Rep_Sum_inverse RS subst) 1); + by (rtac (rewrite_rule [Sum_def] Rep_Sum RS CollectE) 1); + by (REPEAT (eresolve_tac [disjE,exE,ssubst] 1 ORELSE resolve_tac prems 1)); + val sumE = result(); + +8 February + +Changes from Tobias: +Pure/Thy/parse: now list_of admits the empty phrase, while listof_1 does not +Pure/Thy/syntax: uses new list_of, list_of1 + +9 February + +HOL/ex/arith: moved to main HOL directory +HOL/prod: now define the type "unit" and constant "(): unit" + +11 February + +HOL/arith: eliminated redefinitions of nat_ss and arith_ss + +12 February + +MAKE-ALL (Poly/ML) ran perfectly. It took 2:50 hours. + +Pure/Thy/scan/string: now correctly recognizes ML-style strings. + +15 February + +MAKE-ALL (NJ 0.75) ran perfectly. It took 1:37 hours (on albatross) +MAKE-ALL (NJ 0.75) ran perfectly. It took 2:42 hours (on dunlin) +MAKE-ALL (Poly/ML) ran perfectly. It took 2:53 hours (on dunlin) + +**** New tar file placed on /homes/lcp (480K) **** + +18 February + +Pure/Syntax/earley0A/compile_xgram: Tobias deleted the third argument, as +it was unused. + +Pure/Syntax/earley0A: modified accordingly. + +19 February + +MAKE-ALL (NJ 0.75) ran perfectly. It took 3:37 hours +MAKE-ALL (Poly/ML) ran perfectly. It took 2:52 hours + +**** New tar file placed on /homes/lcp (480K) **** + +20 February + +MAKE-ALL (NJ 0.93) ran perfectly. It took 3:30 hours + +10 March + +HOL/fun/image_eqI: fixed bad pattern + +11 March + +MAKE-ALL (Poly/ML) failed in HOL! + +HOL/fun: moved "mono" proofs to HOL/subset, since they rely on subset laws +of Int and Un. + +12 March + +ZF/ex/misc: new example from Bledsoe + +15 March + +ZF/perm: two new theorems inspired by Pastre + +16 March + +Weakened congruence rules for HOL: speeds simplification considerably by +NOT simplifying the body of a conditional or eliminator. + +HOL/simpdata/mk_weak_congs: new, to make weakened congruence rules + +HOL/simpdata/congs: renamed HOL_congs and weakened the "if" rule + +HOL/simpdata/HOL_congs: now contains polymorphic rules for the overloaded +operators < and <= + +HOL/prod: weakened the congruence rule for split +HOL/sum: weakened the congruence rule for case +HOL/nat: weakened the congruence rule for nat_case and nat_rec +HOL/list: weakened the congruence rule for List_rec and list_rec + +HOL & test rebuilt perfectly + +Pure/goals/prepare_proof/mkresult: fixed bug in signature check. Now +compares the FINAL signature with that from the original theory. + +Pure/goals/prepare_proof: ensures that [prove_]goalw checks that the +definitions do not update the proof state. + +17 March + +MAKE-ALL (Poly/ML) ran perfectly. + +18 March + +MAKE-ALL (Poly/ML) failed in HOL/ex/Substitutions + +HOL/ex/Subst/setplus: changed Set.thy to Setplus.thy where +necessary + +ZF/perm: proved some rules about inj and surj + +ZF/ex/misc: did some of Pastre's examples + +Pure/library/gen_ins,gen_union: new + +HOL/ex/Subst/subst: renamed rangeE to srangeE + +18 March + +MAKE-ALL (Poly/ML) failed in HOL/ex/term due to renaming of list_ss in +ex/Subst/alist + +HOL/list/list_congs: new; re-organized simpsets a bit + +Pure/goals/sign_error: new + +Pure/goals/prepare_proof,by_com: now print the list of new theories when +the signature of the proof state changes + +HOL/prod,sexp: renamed fst, snd to fst_conv, snd_conv to avoid over-writing +the library functions fst, snd + +HOL/fun/image_compose: new + +21 March + +MAKE-ALL (NJ 0.93) ran perfectly. It took 3:50 hours +MAKE-ALL (Poly/ML) ran perfectly. It took 3:21 hours +Much slower now (about 30 minutes!) because of HOL/ex/Subst + +**** New tar file placed on /homes/lcp (504K) **** + +ZF/pair,simpdata: renamed fst, snd to fst_conv, snd_conv to avoid over-writing +the library functions fst, snd + +HOL/prod/prod_fun_imageI,E: new + +HOL/ex/Subst/Unify: renamed to Unifier to avoid clobbering structure Unify +of Pure + +24 March + +HOL/trancl/comp_subset_Sigma: new +HOL/wf/wfI: new + +HOL/Subst: moved from HOL/ex/Subst to shorten pathnames +HOL/Makefile: target 'test' now loads Subst/ROOT separately + +*** Installation of gfp, coinduction, ... to HOL *** + +HOL/gfp,llist: new +HOL/univ,sexp,list: replaced with new version + +Sexp is now the set of all well-founded trees, each of type 'a node set. +There is no longer a type 'sexp'. Initial algebras require more explicit +type checking than before. Defining a type 'sexp' would eliminate this, +but would also require a whole new set of primitives, similar to those +defined in univ.thy but restricted to well-founded trees. + +25 March + +Pure/thm: renamed 'bires' to 'eres' in many places (not exported) -- +biresolution now refers to resolution with (flag,rule) pairs. + +Pure/thm/bicompose_aux: SOUNDNESS BUG concerning variable renaming. A Var in +a premise was getting renamed when its occurrence in the flexflex pairs was +not. Martin Coen supplied the following proof of True=False in HOL: + + val [prem] = goal Set.thy "EX a:{c}.p=a ==> p=c"; + br (prem RS bexE) 1; be ssubst 1; be singletonD 1; + val l1 = result(); + + val rls = [refl] RL [bexI] RL [l1]; + + goal Set.thy "True = False"; + brs rls 1; br singletonI 1; + result(); + +Marcus Moore noted that the error only occurred with +Logic.auto_rename:=false. Elements of the fix: + +1. rename_bvs, rename_bvars and bicompose_aux/newAs take tpairs (the +existing flex-flex pairs) as an extra argument. rename_bvs preserves all +Vars in tpairs. + +2. bicompose_aux/tryasms and res now unpack the "cell" and supply its tpairs +to newAs. + +HOL/lfp,gfp,ex/set: renamed Tarski to lfp_Tarski + +HOL/lfp,list,llist,nat,sexp,trancl,Subst/uterm,ex/simult,ex/term: renamed +def_Tarski to def_lfp_Tarski + +26 March + +MAKE-ALL (NJ 0.93) ran perfectly. It took 4:25 hours! +MAKE-ALL (Poly/ML) ran perfectly. It took 3:54 hours! (jobs overlapped) + +Pure/Thy/scan/is_digit,is_letter: deleted. They are already in +Pure/library, and these versions used non-Standard string comparisons! + +Repairing a fault reported by David Aspinall: + show_types := true; read "a"; (* followed by 'prin it' for NJ *) +Raises exception LIST "hd". Also has the side effect of leaving +show_types set at false. + +Pure/goals/read: no longer creates a null TVar +Pure/Syntax/lexicon/string_of_vname: now handles null names +Pure/Syntax/printer/string_of_typ: tidied + +/usr/groups/theory/isabelle/92/Pure/thm: replaced by new version to fix bug +MAKE-ALL on this directory ran perfectly +/usr/groups/theory/ml-aftp/Isabelle92.tar.Z: replaced by new version + +29 March + +MAKE-ALL (NJ 0.93) ran perfectly. It took 4:14 hours! +MAKE-ALL (Poly/ML) ran perfectly. It took 3:43 hours! + +**** New tar file placed on /homes/lcp (518K) **** + +30 March + +ZF/univ/cons_in_Vfrom: deleted "[| a: Vfrom(A,i); b<=Vfrom(A,i) |] ==> +cons(a,b) : Vfrom(A,succ(i))" since it was useless. + +8 April + +MAKE-ALL (Poly/ML) ran perfectly. It took 3:49 hours! + +**** New tar file placed on /homes/lcp (520K) **** + +**** Updates for pattern unification (Tobias Nipkow) **** + +Pure/pattern.ML: new, pattern unification + +Pure/Makefile and ROOT.ML: included pattern.ML + +Pure/library.ML: added predicate downto0 + +Pure/unify.ML: call pattern unification first. Removed call to could_unify. + +FOL/Makefile/FILES: now mentions ifol.ML (previously repeated fol.ML instead) + +**** Installation of Martin Coen's FOLP (FOL + proof objects) **** + +renamed PFOL, PIFOL to FOLP, IFOLP, etc. + +9 April + +MAKE-ALL (NJ 0.93) ran perfectly. It took 4:05 hours! +MAKE-ALL (Poly/ML) ran perfectly. It took 3:31 hours! + +**** New tar file placed on /homes/lcp (576K) **** + +**** Installation of Discrimination Nets **** + +*Affected files (those mentioning Stringtree, compat_thms or rtr_resolve_tac) +Pure/ROOT.ML,goals.ML,stringtree.ML,tactic.ML +Provers/simp.ML +HOL/ex/meson.ML + +*Affected files (those mentioning compat_resolve_tac) +Pure/tactic.ML +Provers/typedsimp.ML +CTT/ctt.ML + +Pure/stringtree: saved on Isabelle/old +Pure/net: new +Pure/Makefile/FILES: now mentions net.ML, not stringtree.ML +Pure/ROOT: now mentions net.ML, not stringtree.ML + +Pure/goals/compat_goal: DELETED + +Pure/tactic/compat_thms,rtr_resolve_tac,compat_resolve_tac,insert_thm, +delete_thm,head_string: DELETED + +Pure/tactic/biresolve_from_nets_tac, bimatch_from_nets_tac, +net_biresolve_tac, net_bimatch_tac, resolve_from_net_tac, match_from_net_tac, +net_resolve_tac, net_match_tac: NEW + +Pure/tactic/filt_resolve_tac: new implementation using nets! + +Provers/simp: replaced by new version + +Provers/typedsimp: changed compat_resolve_tac to filt_resolve_tac and +updated comments + +CTT/ctt.ML: changed compat_resolve_tac to filt_resolve_tac +ZF/simpdata/typechk_step_tac: changed compat_resolve_tac to filt_resolve_tac + +CTT tested + +HOL/ex/meson/ins_term,has_reps: replaced Stringtree by Net + +FOL tested + +Provers/simp/cong_const: new, replaces head_string call in cong_consts +Provers/simp: renamed variables: atomic to at and cong_consts to ccs + +ZF/ex/bin/integ_of_bin_type: proof required reordering of rules -- +typechk_tac now respects this ordering! + +ZF tested + +DOCUMENTATION + +Logics/CTT: Removed mention of compat_resolve_tac +Ref/goals: deleted compat_goal's entry + +Provers/hypsubst/lasthyp_subst_tac: deleted + +FOLP/ROOT/dest_eq: corrected; now hyp_subst_tac works! + +12 April + +MAKE-ALL (NJ 0.93) ran perfectly. It took 4:03 hours! +MAKE-ALL (Poly/ML) ran perfectly. It took 3:28 hours! + +FOLP/{int-prover,classical}/safe_step_tac: uses eq_assume_tac, not assume_tac +FOLP/{int-prover,classical}/inst_step_tac: restored, calls assume and mp_tac +FOLP/{int-prover,classical}/step_tac: calls inst_step_tac + +{FOL,FOLP}/int-prover/safe_brls: removed (asm_rl,true) since assume_tac is +used explicitly!! + +FOLP/ifolp/uniq_assume_tac: new, since eq_assume_tac is almost useless + +FOLP/{int-prover,classical}/uniq_mp_tac: replace eq_mp_tac and call +uniq_assume_tac + +Provers/classical: REPLACED BY 'NET' VERSION! + +13 April + +MAKE-ALL (Poly/ML) failed in ZF and ran out of quota for Cube. + +Unification bug (nothing to do with pattern unification) +Cleaning of flex-flex pairs attempts to remove all occurrences of bound +variables not common to both sides. Arguments containing "banned" bound +variables are deleted -- but this should ONLY be done if the occurrence is +rigid! + +unify/CHANGE_FAIL: new, for flexible occurrence of bound variable +unify/change_bnos: now takes "flex" as argument, indicating path status + +14 April + +MAKE-ALL (Poly/ML) failed in HOL (ASM_SIMP_TAC redefined!) and LK + +LK/ex/hard-quant/37: added "by flexflex_tac" to compensate for flexflex +changes + +Pure/goals/gethyps: now calls METAHYPS directly + +rm-logfiles: no longer mentions directories. WAS + rm log {Pure,FOL,ZF,LCF,CTT,LK,Modal,HOL,Cube}/make*.log + rm {FOL,ZF,LCF,CTT,LK,Modal,HOL,Cube}/test + rm {FOL,ZF,LCF,CTT,LK,Modal,HOL,Cube}/.*.thy.ML + rm {FOL,ZF,HOL}/ex/.*.thy.ML + +MAKE-ALL (Poly/ML) ran perfectly. It took 2:39 hours! (albatross) + +New version of simp on Isabelle/new -- instantiates unknowns provided only +one rule may do so [SINCE REJECTED DUE TO UNPREDICTABLE BEHAVIOR] + +works with FOLP/ex/nat, but in general could fail in the event of +overlapping rewrite rules, since FOLP always instantiates unknowns during +rewriting. + +ZF: tested with new version + +HOL: tested with new version, appeared to loop in llist/Lmap_ident + +**** NEW VERSION OF ASM_SIMP_TAC, WITH METAHYPS **** + +ZF: failed in perm/comp_mem_injD1: the rule anti_refl_rew is too ambiguous! +ZF/wfrec: all uses of wf_ss' require +by (METAHYPS (fn hyps => cut_facts_tac hyps 1 THEN + SIMP_TAC (wf_ss' addrews (hyps)) 1) 1); + +ZF/epsilon/eclose_least: changed ASM_SIMP_TAC to SIMP_TAC; this makes +METAHYPS version work + +ZF/arith/add_not_less_self: adds anti_refl_rew + +ZF/ex/prop-log/hyps_finite: the use of UN_I is very bad -- too undirected. +Swapping the premises of UN_I would probably allow instantiation. + +ZF otherwise seems to work! + +HOL/llist/llistE: loops! due to rewriting by Rep_LList_LCons of Vars + +HOL/ex/prop-log/comp_lemma: failed due to uninstantiated Var in +(CCONTR_rule RS allI) + +*** REJECTED + +15 April + +These overnight runs involve Provers/simp.ML with old treatment of rules +(match_tac) and no METAHYPS; they test the new flexflex pairs and +discrimination nets, to see whether it runs faster. + +MAKE-ALL (NJ 0.93) ran perfectly. It took 3:39 hours (4 mins faster) +MAKE-ALL (Poly/ML) ran perfectly. It took 3:23 hours (5 mins faster) + +ZF/simpdata/ZF_ss: deleted anti_refl_rew; non-linear patterns slow down +discrimination nets (and this rewrite used only ONCE) + +ZF/mem_not_refl: new; replaces obsolete anti_refl_rew + +**Timing experiments** + +fun HYP_SIMP_TAC ss = METAHYPS (fn hyps => HOL_SIMP_TAC (ss addrews hyps) 1); + +On large examples such as ... +HOL/arith/mod_quo_equality +HOL/llist/LListD_implies_ntrunc_equality +ZF/ex/bin/integ_of_bin_succ +... it is 1.5 to 3 times faster than ASM_SIMP_TAC. But cannot replace +ASM_SIMP_TAC since the auto_tac sometimes fails due to lack of assumptions. +If there are few assumptions then HYP_SIMP_TAC is no better. + +Pure/Makefile: now copies $(ML_DBASE) to $(BIN)/Pure instead of calling +make_database, so that users can call make_database for their object-logics. + +Pure/tctical/SELECT_GOAL: now does nothing if i=1 and there is +only one subgoal. + +19 April + +MAKE-ALL (NJ 0.93) failed in HOL due to lack of disc space. +MAKE-ALL (Poly/ML) ran perfectly. It took 3:23 hours + +**** Installation of new simplifier **** + +Provers/simp/EXEC: now calls METAHYPS and passes the hyps as an extra arg +to the auto_tac. + +FOL,HOL/simpdata: auto_tac now handles the hyps argument + +ZF/simpdata/standard_auto_tac: deleted +ZF/simpdata/auto_tac: added hyps argument +ZF/epsilon/eclose_least_lemma: no special auto_tac + +*/ex/ROOT: no longer use 'cd' commands; instead pathnames contain "ex/..." + +20 April + +MAKE-ALL failed in HOL/Subst + +HOL/Subst/setplus/cla_case: renamed imp_excluded_middle and simplified. +Old version caused ambiguity in rewriting: + "[| P ==> P-->Q; ~P ==> ~P-->Q |] ==> Q"; + +**** New tar file placed on /homes/lcp (????) **** + +Pure/Syntax: improvements to the printing of syntaxes +Pure/Syntax/lexicon.ML: added name_of_token +Pure/Syntax/earley0A.ML: updated print_gram + +21 April + +MAKE-ALL (NJ 0.93) ran perfectly. It took 3:44 hours +MAKE-ALL (Poly/ML) failed in HOL due to lack of disc space + +HOL/list,llist: now share NIL, CONS, List_Fun and List_case + +make-all: now compresses the log files, which were taking up 4M; this +reduces their space by more than 1/3 + +rm-logfiles: now deletes compressed log files. + +** Patrick Meche has noted that if the goal is stated with a leading !! +quantifier, then the list of premises is always empty -- this gives the +effect of an initial (cut_facts_tac prems 1). The final theorem is the +same as it would be without the quantifier. + +ZF: used the point above to simplify many proofs +ZF/domrange/cfast_tac: deleted, it simply called cut_facts_tac + +22 April + +MAKE-ALL (NJ 0.93) ran perfectly. It took 3:52 hours?? +MAKE-ALL (Poly/ML) ran perfectly. It took 3:16 hours + +30 April + +HOL: installation of finite set notation: {x1,...,xn} (by Tobias Nipkow) + +HOL/set.thy,set.ML,fun.ML,equalities.ML: addition of rules for "insert", +new derivations for "singleton" + +HOL/llist.thy,llist.ML: changed {x.False} to {} + +**** New tar file placed on /homes/lcp (584K) **** + +4 May + +MAKE-ALL (NJ 0.93) ran out of space in LK. +MAKE-ALL (Poly/ML) ran perfectly. It took 3:14 hours + +Pure/Makefile: inserted "chmod u+w $(BIN)/Pure;" in case $(ML_DBASE) is +write-protected + +5 May + +HOL/list/not_Cons_self: renamed from l_not_Cons_l +HOL/list/not_CONS_self: new + +HOL/llist.thy/Lconst: changed type and def to remove Leaf +HOL/llist.ML: changed Lconst theorems + +6 May + +MAKE-ALL (Poly/ML) ran perfectly. It took 3:18 hours + +** Installation of new HOL from Tobias ** + +HOL/ex/{finite,prop-log} made like the ZF versions +HOL/hol.thy: type classes plus, minus, times; overloaded operators + - * +HOL/set: set enumeration via "insert" + additions to set_cs and set_ss +HOL/set,subset,equalities: various lemmas to do with {}, insert and - +HOL/llist: One of the proofs needs one fewer commands +HOL/arith: many proofs require type constraints due to overloading + +** end Installation ** + +ZF/ex/misc: added new lemmas from Abrial's paper + +7 May + +HOL/llist.ML/LList_corec_subset1: deleted a fast_tac call; the previous +simplification now proves the subgoal. + +**** New tar file placed on /homes/lcp (584K) **** + +** Installation of new simplifier from Tobias ** + +The "case_splits" parameter of SimpFun is moved from the signature to the +simpset. SIMP_CASE_TAC and ASM_SIMP_CASE_TAC are removed. The ordinary +simplification tactics perform case splits if present in the simpset. + +The simplifier finds out for itself what constant is affected. Instead of +supplying the pair (expand_if,"if"), supply just the rule expand_if. + +This change affects all calls to SIMP_CASE_TAC and all applications of SimpFun. + +MAKE-ALL (Poly/ML) ran perfectly. It took 3:18 hours + +Cube/ex: UNTIL1, UNTIL_THM: replaced by standard tactics DEPTH_SOLVE_1 and +DEPTH_SOLVE + +HOL: installation of NORM tag for simplication. How was it forgotten?? + +HOL/hol.thy: declaration of NORM +HOL/simpdata: NORM_def supplied to SimpFun + +10 May + +MAKE-ALL (Poly/ML) ran perfectly. It took 3:33 hours?? + +11 May + +HOL/prod/Prod_eq: renamed Pair_eq +HOL/ex/lex-prod: wf_lex_prod: simplified proof + +HOL/fun/inj_eq: new + +HOL/llist/sumPairE: deleted, thanks to new simplifier's case splits! + +12 May + +MAKE-ALL (NJ 0.93) ran out of space in HOL. +MAKE-ALL (Poly/ML) failed in HOL. +HOL/Subst/utermlemmas/utlemmas_ss: deleted Prod_eq from the congruence rules + +13 May + +Pure/logic/flexpair: moved to term, with "equals" etc. Now pervasive +Pure/logic/mk_flexpair: now exported +Pure/logic/dest_flexpair: new +Pure/goals/print_exn: now prints the error message for TERM and TYPE + +Pure/Syntax/sextension: now =?= has type ['a::{}, 'a] => prop because +flexflex pairs can have any type at all. Thus == must have the same type. + +Pure/thm/flexpair_def: now =?= and == are equated for all 'a::{}. + +Pure/tctical/equal_abs_elim,equal_abs_elim_list: new (for METAHYPS fix) +Pure/tctical/METAHYPS: now works if new proof state has flexflex pairs + +Pure/Syntax/earley0A,syntax,lexicon: Tokens are represented by strings now, +not by integers. (Changed by Tobias) + +*** Installation of more printing functions *** + +Pure/sign/sg: changed from a type abbrev to a datatype +Pure/type/type_sig: changed from a type abbrev to a datatype +These changes needed for abstract type printing in NJ + +Pure/tctical/print_sg,print_theory: new + +Pure/drule: new file containing derived rules and printing functions. +Mostly from tctical.ML, but includes rewriting rules from tactic.ML. + +Pure/ROOT: loads drule before tctical; TacticalFun,TacticFun,GoalsFun now +depend on Drule and have sharing constraints. + +14 May + +Installing new print functions for New Jersey: incompatible with Poly/ML! + +Pure/NJ/install_pp_nj: new (requires datatypes as above) +Pure/POLY/install_pp_nj: a dummy version + +Pure/ROOT: calls install_pp_nj to install printing for NJ + +*/ROOT: added extra install_pp calls (sg, theory, cterm, typ, ctyp) for +Poly/ML [ZF,LCF,Modal do not need them since they inherit them from another +logic -- make_database is not used] + +17 May + +MAKE-ALL (NJ 0.93) ran perfectly. It took 3:57 hours?? + +Pure/Syntax/lexicon: Yet another leaner and faster version ... (from Tobias) + +18 May + +MAKE-ALL (Poly/ML) ran perfectly. It took 3:36 hours + +19 May + +ZF/equalities/Union_singleton,Inter_singleton: now refer to {b} instead of +complex assumptions + +20 May + +HOL/list: Tobias added the [x1,...,xn] notation and the functions hd, tl, +null and list_case. + +1 June + +MAKE-ALL (Poly/ML) ran perfectly. It took 3:39 hours + +**** New tar file 92.tar.z placed on /homes/lcp (376K) **** + +MAKE-ALL (NJ 0.93) ran perfectly. It took 1:49 hours on albatross. + +Pure/tactic/dres_inst_tac,forw_inst_tac: now call the new +make_elim_preserve to preserve Var indexes when creating the elimination +rule. + +ZF/ex/ramsey: modified calls to dres_inst_tac + +2 June + +Pure/Thy/read/read_thy,use_thy: the .thy.ML file is now written to the +current directory, since the pathname may lead to a non-writeable area. + +HOL/arith: renamed / and // to div and mod +ZF/arith: renamed #/ and #// to div and mod + +MAKE-ALL (Poly/ML) ran perfectly. It took 1:48 hours on albatross. + +**** New tar file 92.tar.z placed on /homes/lcp (376K) **** + +Pure/NJ/commit: new dummy function +FOLP/ex/ROOT: inserted commit call to avoid Poly/ML problems + +make-all: now builds FOLP also! + +3 June + +ZF/zf.thy,HOL/list.thy,HOL/set.thy: now constructions involving {_}, [_], +<_,_> are formatted as {(_)}, [(_)], + +MAKE-ALL (Poly/ML) ran perfectly. It took 4:37 hours on muscovy (with FOLP). + +ZF/Makefile: removed obsolete target for .rules.ML + +All object-logic Makefiles: EXAMPLES ARE NO LONGER SAVED. This saves disc +and avoids problems (in New Jersey ML) of writing to the currently +executing image. + +4 June + +Pure/logic/rewritec: now uses nets for greater speed. Functor LogicFun now +takes Net as argument. + +Pure/ROOT: now loads net before logic. + +MAKE-ALL (Poly/ML) failed in ZF and HOL. + +LK/lk.thy: changed constant "not" to "Not" (for consistency with FOL) + +7 June + +Pure/tactic/is_letdig: moved to library +Pure/Syntax/lexicon/is_qld: deleted, was same as is_letdig + +MAKE-ALL (Poly/ML) ran perfectly. It took 2:07 hours on albatross. +MAKE-ALL (NJ 0.93) ran perfectly. It took 4:41 hours on dunlin. + +HOL/set/UN1,INT1: new union/intersection operators. Binders UN x.B(x), +INT x.B(x). + +HOL/univ,llist: now use UN x.B(x) instead of Union(range(B)) + +HOL/subset: added lattice properties for INT, UN (both forms) + +8 June + +MAKE-ALL (NJ 0.93) ran perfectly. It took 4:45 hours on dunlin. + +**** New tar file 92.tar.z placed on /homes/lcp (384K) **** + +14 June + +HOL/list.thy/List_rec_def: changed pred_sexp (a variable!) to pred_Sexp. +Using def_wfrec hides such errors!! + +**** New tar file 92.tar.gz placed on /homes/lcp (384K) **** + +** NEW VERSION FROM MUNICH WITH ==-REWRITING ** + +** The following changes are Toby's ** + +type.ML: + +Renamed mark_free to freeze_vars and thaw_tvars to thaw_vars. +Added both functions to the signature. + +sign.ML: + +Added val subsig: sg * sg -> bool to signature. +Added trueprop :: prop and mark_prop : prop => prop to pure_sg. + +Added + +val freeze_vars: term -> term +val thaw_vars: term -> term +val strip_all_imp: term * int -> term list * term * int + +Moved rewritec_bottom and rewritec_top to thm.ML. +Only bottom-up rewriting supported any longer. + +thm.ML: + +Added + +(* internal form of conditional ==-rewrite rules *) +type meta_simpset +val add_mss: meta_simpset * thm list -> meta_simpset +val empty_mss: meta_simpset +val mk_mss: thm list -> meta_simpset + +val mark_prop_def: thm +val truepropI: thm +val trueprop_def: thm + +(* bottom-up conditional ==-rewriting with local ==>-assumptions *) +val rewrite_cterm: meta_simpset -> (thm -> thm list) + -> (meta_simpset -> thm list -> Sign.cterm -> thm) + -> Sign.cterm -> thm +val trace_simp: bool ref + +Simplified concl_of: call to Logic.skip_flexpairs redundant. + +drule.ML: + +Added + +(* rewriting *) +val asm_rewrite_rule: (thm -> thm list) -> thm list -> thm -> thm +val rewrite_goal_rule: (thm -> thm list) -> thm list -> int -> thm -> thm +val rewrite_goals_rule: (thm -> thm list) -> thm list -> thm -> thm + +(* derived concepts *) +val forall_trueprop_eq: thm +val implies_trueprop_eq: thm +val mk_trueprop_eq: thm -> thm +val reflexive_eq: thm +val reflexive_thm: thm +val trueprop_implies_eq: thm +val thm_implies: thm -> thm -> thm +val thm_equals: thm -> thm -> thm + +(*Moved here from tactic.ML:*) +val asm_rl: thm +val cut_rl: thm +val revcut_rl: thm + +tactic.ML: + +Added + +val asm_rewrite_goal_tac: (thm -> thm list) -> thm list -> int -> tactic +val asm_rewrite_goals_tac: (thm -> thm list) -> thm list -> tactic +val asm_rewrite_tac: (thm -> thm list) -> thm list -> tactic +val fold_goal_tac: thm list -> int -> tactic +val rewrite_goal_tac: thm list -> int -> tactic + +Moved to drule.ML: +val asm_rl: thm +val cut_rl: thm +val revcut_rl: thm + +goals.ML: + +Changed prepare_proof to make sure that rewriting with empty list of +meta-thms is identity. + +** End of Toby's changes ** + +16 June + +Pure/sign/typ_of,read_ctyp: new +Pure/logic/dest_flexpair: now exported + +Pure/drule/flexpair_intr,flexpair_elim: new; fixes a bug in +flexpair_abs_elim_list + +HOL/equalities/image_empty,image_insert: new +HOL/ex/finite/Fin_imageI: new + +Installed Martin Coen's CCL as new object-logic + +17 June + +** More changes from Munich (Markus Wenzel) ** + +Pure/library: added the, is_some, is_none, separate and improved space_implode +Pure/sign: Sign.extend now calls Syntax.extend with list of constants +Pure/symtab: added is_null +Pure/Syntax/sextension: added empty_sext +Pure/Syntax/syntax: changed Syntax.extend for compatibility with future version + +HOL now exceeds poly's default heap size. Hence HOL/Makefile needs to +specify -h 8000. + +HOL/univ/ntrunc_subsetD, etc: deleted the useless j syntax + + Pure/Makefile + SYNTAX_FILES: added Syntax/ast.ML + + Pure/Syntax/pretty.ML + added str_of: T -> string + + Pure/Syntax/ast.ML + added this file + + Pure/Syntax/extension.ML + Pure/Syntax/parse_tree.ML + Pure/Syntax/printer.ML + Pure/Syntax/ROOT.ML + Pure/Syntax/sextension.ML + Pure/Syntax/syntax.ML + Pure/Syntax/type_ext.ML + Pure/Syntax/xgram.ML + These files have been completely rewritten, though the global structure + is similar to the old one. + + +30-Jun-1993 MMW + New versions of HOL and Cube: use translation rules wherever possible; + + HOL/hol.thy + cleaned up + removed alt_tr', mk_bindopt_tr' + alternative binders now implemented via translation rules and mk_alt_ast_tr' + + HOL/set.thy + cleaned up + removed type "finset" + now uses category "args" for finite sets + junked "ML" section + added "translations" section + + HOL/list.thy + cleaned up + removed type "listenum" + now uses category "args" for lists + junked "ML" section + added "translations" section + + Cube/cube.thy + cleaned up + changed indentation of Lam and Pi from 2 to 3 + removed qnt_tr, qnt_tr', no_asms_tr, no_asms_tr' + fixed fun_tr': all but one newly introduced frees will have type dummyT + added "translations" section + + +30-Jun-1993, 05-Jul-1993 MMW + Improved toplevel pretty printers: + - unified interface for POLY and NJ; + - print functions now insert atomic string into the toplevel's pp stream, + rather than writing it to std_out (advantage: output appears at the + correct position, disadvantage: output cannot be broken); + (Is there anybody in this universe who exactly knows how Poly's install_pp + is supposed to work?); + + Pure/NJ.ML + removed dummy install_pp + added make_pp, install_pp + + Pure/POLY.ML + removed dummy install_pp_nj + added make_pp + + Pure/ROOT.ML + removed install_pp_nj stuff + + Pure/drule.ML + added str_of_sg, str_of_theory, str_of_thm + + Pure/install_pp.ML + added this file + + Pure/sign.ML + added str_of_term, str_of_typ, str_of_cterm, str_of_ctyp + + Pure/goals.ML + added str_of_term, str_of_typ + + CTT/ROOT.ML + Cube/ROOT.ML + FOL/ROOT.ML + FOLP/ROOT.ML + HOL/ROOT.ML + LK/ROOT.ML + replaced install_pp stuff by 'use "../Pure/install_pp.ML"' + + +01-Jul-1993 MMW + Misc small fixes + + CCL/ROOT.ML + HOL/ROOT.ML + added ".ML" suffix to some filenames + + HOL/ex/unsolved.ML + replaced HOL_Rule.thy by HOL.thy + + Pure/NJ.ML + quit was incorrectly int -> unit + +END MMW CHANGES + +Pure/Syntax/sextension/eta_contract: now initially false + +Pure/library/cat_lines: no longer calls "distinct" +Pure/sign: replaced to calls of implode (map (apr(op^,"\n") o ... by cat_lines +NB This could cause duplicate error messages from Pure/sign and Pure/type + +Pure/goals/prove_goalw: now prints some of the information from print_exn + +9 July + +MAKE-ALL (Poly/ML) ran perfectly. It took 2:26 hours on albatross. + +**** New tar file 93.tar.gz placed on /homes/lcp (480K) **** + +12 July + +MAKE-ALL (NJ 0.93) ran perfectly. It took 2:13 hours on albatross. +MAKE-ALL (Poly/ML) ran perfectly. It took 2:25 hours on albatross. + +22 July + +ZF/zf.thy: new version from Marcus Wenzel + +ZF: ** installation of inductive definitions ** + +changing the argument order of "split"; affects fst/snd too +sum.thy zf.thy ex/bin.thy ex/integ.thy ex/simult.thy ex/term.thy +pair.ML ex/integ.ML + +changing the argument order of "case" and adding "Part": sum.thy sum.ML + +ZF/zf.ML/rev_subsetD,rev_bspec: new + +ZF/mono: new rules for implication +ZF/mono/Collect_mono: now for use with implication rules + +ZF/zf.ML/ballE': renamed rev_ballE + +ZF/list.thy,list.ML: files renamed list-fn.thy, list-fn.ML +ZF/list.ML: new version simply holds the datatype definition +NB THE LIST CONSTRUCTORS ARE NOW Nil/Cons, not 0/Pair. + +ZF/extend_ind.ML, datatype.ML: new files +ZF/fin.ML: moved from ex/finite.ML + +23 July + +ZF/ex/sexp: deleted this example -- it seems hardly worth the trouble of +porting. + +ZF/ex/bt.thy,bt.ML: files renamed bt-fn.thy, bt-fn.ML +ZF/ex/bt.ML: new version simply holds the datatype definition + +ZF/ex/term.thy,term.ML: files renamed term-fn.thy, term-fn.ML +ZF/ex/term.ML: new version simply holds the datatype definition + +ZF/sum/InlI,InrI: renamed from sum_InlI, sum_InlI + +26 July + +ZF/univ/rank_ss: new, for proving recursion equations + +ZF/domrange/image_iff,image_singleton_iff,vimage_iff,vimage_singleton_iff, +field_of_prod:new + +ZF/domrange/field_subset: modified + +ZF/list/list_cases: no longer proved by induction! +ZF/wf/wf_trancl: simplified proof + +ZF/equalities: new laws for field + +29 July + +ZF/trancl/trancl_induct: new +ZF/trancl/rtrancl_induct,trancl_induct: now with more type information + +** More changes from Munich (Markus Wenzel) ** + +Update of new syntax module (aka macro system): mostly internal cleanup and +polishing; + + Pure/Syntax/* + added Ast.stat_norm + added Syntax.print_gram, Syntax.print_trans, Syntax.print_syntax + cleaned type and Pure syntax: "_CLASSES" -> "classes", "_SORTS" -> "sorts", + "_==>" -> "==>", "_fun" -> "fun", added some space for printing + Printer: partial fix of the "PROP " problem: print "PROP " before + any Var or Free of type propT + Syntax: added ndependent_tr, dependent_tr' + + Pure/sign.ML: removed declaration of "==>" (now in Syntax.pure_sext) + +Changes to object logics: minor cleanups and replacement of most remaining ML +translations by rewrite rules (see also file "Translations"); + + ZF/zf.thy + added "translations" section + removed all parse/print translations except ndependent_tr, dependent_tr' + fixed dependent_tr': all but one newly introduced frees have type dummyT + replaced id by idt in order to make terms rereadable if !show_types + + Cube/cube.thy + removed necontext + replaced fun_tr/tr' by ndependent_tr/dependent_tr' + + CTT/ctt.thy + added translations rules for PROD and SUM + removed dependent_tr + removed definitions of ndependent_tr, dependent_tr' + + HOL/set.thy: replaced id by idt + + CCL/ROOT.ML: Logtic -> Logic + + CCL/set.thy + added "translations" section + removed "ML" section + replaced id by idt + + CCL/types.thy + added "translations" section + removed definitions of ndependent_tr, dependent_tr' + replaced id by idt + +Yet another improvement of toplevel pretty printers: output now breakable; + + Pure/NJ.ML Pure/POLY.ML improved make_pp + + Pure/install_pp.ML: replaced str_of_* by pprint_* + + Pure/drule.ML: replaced str_of_{sg,theory,thm} by pprint_* + + Pure/sign.ML: replaced str_of_{term,typ,cterm,ctyp} by pprint_* + + Pure/goals.ML: fixed and replaced str_of_{term,typ} by pprint_* + + Pure/Syntax/pretty.ML: added pprint, quote + +Minor changes and additions; + + Pure/sign.ML: renamed stamp "PURE" to "Pure" + + Pure/library.ML + added quote: string -> string + added to_lower: string -> bool + + Pure/NJ.ML,POLY.ML: added file_info of Carsten Clasohm + +30 July + +MAKE-ALL (Poly/ML) ran perfectly. + +Pure/goals/print_sign_exn: new, takes most code from print_exn +Pure/goals/prove_goalw: displays exceptions using print_sign_exn + +Pure/drule/print_sg: now calls pretty_sg to agree with pprint_sg + +Pure/library,...: replaced front/nth_tail by take/drop. + +Pure/term/typ_tfrees,typ_tvars,term_tfrees,term_tvars: new +thm/mk_rew_triple, drule/types_sorts, sign/zero_tvar_indices: now use the above + +Pure/logic/add_term_vars,add_term_frees,insert_aterm,atless: +moved to term, joining similar functions for type variables; +Logic.vars and Logic.frees are now term_vars and term_frees + +Pure/term/subst_free: new + +Pure/tactic/is_fact: newly exported + +Provers/simp/mk_congs: uses filter_out is_fact to delete trivial cong rules + +Pure/tactic/rename_last_tac: now uses Syntax.is_identifier instead of +forall is_letdig + +**** New tar file 93.tar.gz placed on /homes/lcp (448K) **** + +2 August + +MAKE-ALL (NJ 0.93) failed in ZF due to Compiler bug: elabDecl:open:FctBodyStr +MAKE-ALL (Poly/ML) failed in ZF/enum. It took 2:33 hours on albatross. + +Pure/drule/triv_forall_equality: new +Pure/tactic/prune_params_tac: new + +Provers/hypsubst/bound_hyp_subst_tac: new, safer than hyp_subst_tac + +3 August + +Pure/tactic/rule_by_tactic: new + +ZF/perm/compEpair: now proved via rule_by_tactic + +ZF/extend_ind/cases,mk_cases: new +ZF/datatype/mk_free: new +ZF/list: now calls List.mk_cases + +4 August + +Provers/slow_tac,slow_best_tac: new + +5 August + +MAKE-ALL (Poly/ML) failed in ZF + +ZF/sum/sumE2: deleted since unused +ZF/sum/sum_iff,sum_subset_iff,sum_equal_iff: new +ZF/univ/Transset_Vfrom: new; used in proof of Transset_Vset + +6 August + +Pure/goals/prepare_proof: after "Additional hypotheses", now actually +prints them! + +ZF/ordinal/Transset_Union_family, Transset_Inter_family: renamed from +Transset_Union, Transset_Inter + +ZF/ordinal/Transset_Union: new +ZF/univ/pair_in_univ: renamed Pair_in_univ + +ZF/mono/product_mono: generalized to Sigma_mono; changed uses in trancl, univ + +ZF/lfp/lfp_Tarski,def_lfp_Tarski: renamed from Tarski,def_Tarski; changed +uses in extend_ind.ML, nat.ML, trancl.ML. + +ZF/ex/misc: Schroeder-Bernstein Theorem moved here from lfp.ML + +ZF/fixedpt.thy,.ML: renamed from lfp.thy,.ML, and gfp added + +9 August + +ZF/zf.thy/ndependent_tr,dependent_tr': deleted, since they are now on +Syntax/sextension. + +11 August + +Pure/library.ML: added functions +assocs: (''a * 'b list) list -> ''a -> 'b list +transitive_closure: (''a * ''a list) list -> (''a * ''a list) list + +Pure/type.ML: deleted (inefficient) transitive_closure diff -r 000000000000 -r a5a9c433f639 expandshort --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/expandshort Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,31 @@ +#! /bin/sh +# +#expandshort - shell script to expand shorthand goal commands +# ALSO contracts uses of resolve_tac, dresolve_tac, eresolve_tac, +# rewrite_goals_tac on 1-element lists +# +# Usage: +# expandshort FILE1 ... FILEn +# +# leaves previous versions as XXX~~ +# +for f in $* +do +echo Expanding shorthands in $f. \ Backup file is $f~~ +mv $f $f~~; sed -e ' +s/^ba \([0-9]*\); *$/by (assume_tac \1);/ +s/^br \(.*\) \([0-9]*\); *$/by (rtac \1 \2);/ +s/^brs \(.*\) \([0-9]*\); *$/by (resolve_tac \1 \2);/ +s/^bd \(.*\) \([0-9]*\); *$/by (dtac \1 \2);/ +s/^bds \(.*\) \([0-9]*\); *$/by (dresolve_tac \1 \2);/ +s/^be \(.*\) \([0-9]*\); *$/by (etac \1 \2);/ +s/^bes \(.*\) \([0-9]*\); *$/by (eresolve_tac \1 \2);/ +s/^bw \(.*\); *$/by (rewtac \1);/ +s/^bws \(.*\); *$/by (rewrite_goals_tac \1);/ +s/dresolve_tac *\[\([a-zA-Z0-9_]*\)\] */dtac \1 /g +s/eresolve_tac *\[\([a-zA-Z0-9_]*\)\] */etac \1 /g +s/resolve_tac *\[\([a-zA-Z0-9_]*\)\] */rtac \1 /g +s/rewrite_goals_tac *\[\([a-zA-Z0-9_]*\)\]\( *\)/rewtac \1\2/g +' $f~~ > $f +done +echo Finished. diff -r 000000000000 -r a5a9c433f639 get-rulenames --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/get-rulenames Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,20 @@ +#!/bin/sh +# Title: get-rulenames (see also make-rulenames) +# Author: Larry Paulson, Cambridge University Computer Laboratory +# Copyright 1990 University of Cambridge +# +#shell script to generate "val" declarations for a theory's axioms +# also generates a comma-separated list of axiom names +# +# usage: make-rulenames +# +#Rule lines begin with a line containing the word "extend_theory" +# and end with a line containing the word "get_axiom" +#Each rule name xyz must appear on a line that begins +# ("xyz" +#Output lines have the form +# val Eq_comp = ax"Eq_comp"; +# +sed -n -e '/ext[end]*_theory/,/get_axiom/ s/^[ []*("\([^"]*\)".*$/val \1 = ax"\1";/p' $1 +echo +echo `sed -n -e '/ext[end]*_theory/,/get_axiom/ s/^[ []*("\([^"]*\)".*$/\1/p' $1 | tr '\012' ','` diff -r 000000000000 -r a5a9c433f639 make-all --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/make-all Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,169 @@ +#! /bin/sh +# +#make-all: make all systems afresh + +# Creates gzipped log files called makeNNNN.log.gz on each subdirectory and +# displays the last few lines of these files -- this indicates whether +# the "make" failed (whether it terminated due to an error) + +# switches are +# -noforce don't delete old databases/images first +# -clean delete databases/images after use (leaving Pure) +# -notest make databases/images w/o running the examples +# -noexec don't execute, just check settings and Makefiles + +#Environment variables required: +# ISABELLEBIN: the directory to hold Poly/ML databases or New Jersey ML images +# ISABELLECOMP: the ML compiler + +# A typical shell script for /bin/sh is... +# ML_DBASE=/usr/groups/theory/poly2.04/`arch`/ML_dbase +# ISABELLEBIN=/homes/`whoami`/bin +# ISABELLECOMP="poly -noDisplay" +# export ML_DBASE ISABELLEBIN ISABELLECOMP +# nohup make-all $* + +set -e #fail immediately upon errors + +# process command line switches +CLEAN="off"; +FORCE="on"; +TEST="test"; +EXEC="on"; +NO=""; +for A in $* +do + case $A in + -clean) CLEAN="on" ;; + -noforce) FORCE="off" ;; + -notest) TEST="" ;; + -noexec) EXEC="off" + NO="-n" ;; + *) echo "Bad flag for make-all: $A" + echo "Usage: make-all [-noforce] [-clean] [-notest] [-noexec]" + exit ;; + esac +done + +echo Started at `date` +echo Source=`pwd` +echo Destination=${ISABELLEBIN?'No destination directory specified'} +echo force=$FORCE ' ' clean=$CLEAN ' ' +echo Compiler=${ISABELLECOMP?'No compiler specified'} +echo Running on `hostname` +echo Log files will be called make$$.log.gz + +case $FORCE.$EXEC in + on.on) (cd $ISABELLEBIN; rm -f Pure FOL ZF CCL LCF CTT LK Modal HOL Cube FOLP) +esac + +echo +echo +echo '*****Pure Isabelle*****' +(cd Pure; make $NO > make$$.log) +tail Pure/make$$.log +gzip Pure/make$$.log + +echo +echo +echo '*****First-Order Logic (FOL)*****' +(cd FOL; make $NO $TEST > make$$.log) +tail FOL/make$$.log +gzip FOL/make$$.log +#cannot delete FOL yet... it is needed for ZF, CCL and LCF! + +echo +echo +echo '*****Set theory (ZF)*****' +(cd ZF; make $NO $TEST > make$$.log) +tail ZF/make$$.log +gzip ZF/make$$.log +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/ZF +esac + +echo +echo +echo '*****Classical Computational Logic (CCL)*****' +(cd CCL; make $NO $TEST > make$$.log) +tail CCL/make$$.log +gzip CCL/make$$.log +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/CCL +esac + +echo +echo +echo '*****Domain Theory (LCF)*****' +(cd LCF; make $NO $TEST > make$$.log) +tail LCF/make$$.log +gzip LCF/make$$.log +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/FOL $ISABELLEBIN/LCF +esac + +echo +echo +echo '*****Constructive Type Theory (CTT)*****' +(cd CTT; make $NO $TEST > make$$.log) +tail CTT/make$$.log +gzip CTT/make$$.log +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/CTT +esac + +echo +echo +echo '*****Classical Sequent Calculus (LK)*****' +(cd LK; make $NO $TEST > make$$.log) +tail LK/make$$.log +gzip LK/make$$.log +#cannot delete LK yet... it is needed for Modal! + +echo +echo +echo '*****Modal logic (Modal)*****' +(cd Modal; make $NO $TEST > make$$.log) +tail Modal/make$$.log +gzip Modal/make$$.log +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/LK $ISABELLEBIN/Modal +esac + +echo +echo +echo '*****Higher-Order Logic (HOL)*****' +(cd HOL; make $NO $TEST > make$$.log) +tail HOL/make$$.log +gzip HOL/make$$.log +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/HOL +esac + +echo +echo +echo '*****The Lambda-Cube (Cube)*****' +(cd Cube; make $NO $TEST > make$$.log) +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/Cube +esac +tail Cube/make$$.log +gzip Cube/make$$.log + +echo +echo +echo '*****First-Order Logic with Proof Terms (FOLP)*****' +(cd FOLP; make $NO $TEST > make$$.log) +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/FOLP +esac +tail FOLP/make$$.log +gzip FOLP/make$$.log + +case $TEST.$EXEC in + test.on) echo + echo '***** Now check the dates on the "test" files *****' + ls -lrt FOL/test ZF/test CCL/test LCF/test CTT/test\ + LK/test Modal/test HOL/test Cube/test FOLP/test +esac +echo Finished at `date` diff -r 000000000000 -r a5a9c433f639 make-dist --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/make-dist Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,21 @@ +#!/bin/sh +#make-dist +#make a distribution directory of Isabelle sources. Example: +# rm -r /usr/groups/theory/isabelle/91 +# make-dist /usr/groups/theory/isabelle/91 + +#BEFORE MAKING A NEW DISTRIBUTION VERSION, CHECK... +# * that make-all works perfectly +# * that README files are up-to-date +# * that the version number has been updated + +#This version copies EVERYTHING!!!!!!!!!!!!!!!! + +set -e #terminate if error + +#Pure Isabelle +mkdir ${1?'No destination directory specified'} +cp -ipr . $1 + +#TO WRITE POLY/ML AND ISABELLE TAPES, USE SHELL SCRIPT write-dist +#TO PACK FOR EMAIL, USE SHELL SCRIPTS make-emaildist, send-emaildist diff -r 000000000000 -r a5a9c433f639 make-rulenames --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/make-rulenames Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,36 @@ +#!/bin/sh +# Title: make-rulenames +# Author: Larry Paulson, Cambridge University Computer Laboratory +# Copyright 1990 University of Cambridge +# +#shell script for adding signature and val declarations to a rules file +# usage: make-rulenames +# +#Input is the file ruleshell.ML, which defines a theory. +#Output is .rules.ML +# +# +#Rule lines begin with a line containing the word "extend_theory" +# and end with a line containing the word "get_axiom" +# +#Each rule name xyz must appear on a line that begins +# ("xyz" +# ENSURE THAT THE FIRST RULE LINE DOES NOT CONTAIN A "[" CHARACTER! +#The file RULESIG gets lines like val Eq_comp: thm +# These are inserted after the line containing the string INSERT-RULESIG +# +#The file RULENAMES gets lines like val Eq_comp = ax"Eq_comp"; +# These are inserted after the line containing the string INSERT-RULENAMES +#The input file should define the function "ax" above this point. +# +set -eu #terminate if error or unset variable +if [ ! '(' -d $1 -a -f $1/ruleshell.ML ')' ]; \ + then echo $1 is not a suitable directory; exit 1; \ + fi +sed -n -e '/extend_theory/,/get_axiom/ s/^ *("\([^"]*\)".*$/ val \1: thm/p' $1/ruleshell.ML > RULESIG +sed -n -e '/extend_theory/,/get_axiom/ s/^ *("\([^"]*\)".*$/val \1 = ax"\1";/p' $1/ruleshell.ML > RULENAMES +sed -e '/INSERT-RULESIG/ r RULESIG +/INSERT-RULENAMES/ r RULENAMES' $1/ruleshell.ML > $1/.rules.ML +#WARNING: there must be no spaces after the filename in the "r" command!! +rm RULESIG RULENAMES + diff -r 000000000000 -r a5a9c433f639 prove_goal.el --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/prove_goal.el Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,125 @@ +;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;;; +;; special function for Isabelle +;; +;; +; goalify.el +; +; Emacs command to change "goal" proofs to "prove_goal" proofs +; and reverse IN A REGION. +; [would be difficult in "sed" since replacements involve multiple lines] +; +;; origin is prove_goalify.el +;; enhanced by Franz Regensburger +;; corrected some errors in regular expressions +;; changed name prove_goalify --> goalify +;; added inverse functions ungoalify +; +; function goalify: +; +; val PAT = goalARGS;$ +; COMMANDS;$ +; val ID = result(); +; +; to +; +; val ID = prove_goalARGS +; (fn PAT=> +; [ +; COMMANDS +; ]); +;; +;; Note: PAT must be an identifier. _ as pattern is not supported. +;; +; function ungoalify: +; +; val ID = prove_goalARGS +; (fn PAT=> +; [ +; COMMANDS +; ]); +; +; +; to +; val PAT = goalARGS;$ +; COMMANDS;$ +; val ID = result(); +; + + +(defun ungoalify (alpha omega) + "Change well-formed prove_goal proofs to goal...result" + (interactive "r" + "*") + ; 0: restrict editing to region + (narrow-to-region alpha omega) + + ; 1: insert delimiter ID + (goto-char (point-min)) + (replace-regexp + "[ \t]*val[ \t]+\\([^ \t\n=]+\\)[ \t\n=]+prove_goal" "\\1") + + ; 2: insertt delimiter ARGS  PAT  and  before first command + (goto-char (point-min)) + (replace-regexp + "[ \n\t]*(fn[ \t]+\\([^=]+\\)=>[^(]*" "\\1\n") + + ; 3: shift  over all commands + ; Note: only one line per command + (goto-char (point-max)) + (while (not (equal (point) (point-min))) + (goto-char (point-min)) + (replace-regexp + "[ \t]*\\(.*\\),[ \t]*\n" "by \\1;\n")) + + ; 4: fix last  + (goto-char (point-min)) + (replace-regexp + "[ \t]*\\(.*\\)[ \t\n]*\][ \t\n]*)[ \t\n]*;" "by \\1;") + + ; 5: arange new val Pat = goal .. + (goto-char (point-min)) + (replace-regexp + "\\([^]*\\)\\([^]*\\)\\([^]*\\)\n\\([^]*\\)" + "val \\3= goal\\2;\n\\4\nval \\1 = result();") + + ; widen again + (widen) +) + + +(defun goalify (alpha omega) + "Change well-formed goal...result proofs to use prove_goal" + (interactive "r" + "*") + + ; 0: restrict editing to region + (narrow-to-region alpha omega) + + ; 1: delimit the identifier in "val ID = result()" using ^Q + (goto-char (point-min)) + (replace-regexp "val[ \t\n]+\\([^ \t\n=]+\\)[ \t\n=]+result();[ \t]*$" + "\\1") + + ; 2: replace terminal \"; by  + (goto-char (point-min)) + (replace-regexp "\";[ \t]*$" "") + + ; 3: replace lines "by ...;" with "...," + (goto-char (point-min)) + (replace-regexp "by[ \n\t]*\\([^;]*\\)[ \t\n]*;" "\t\\1,") + + ; 4: removing the extra commas, those followed by ^Q + (goto-char (point-min)) + (replace-regexp ",[ \n\t]*" "") + + ; 5: transforming goal... to prove_goal... + (goto-char (point-min)) + (replace-regexp + "val[ \t\n]+\\([^ \n\t=]+\\)[ \t\n=]+goal\\([^]*\\) +\\([^]*\\)\\([^]*\\)" + "val \\4 = prove_goal\\2\"\n (fn \\1 =>\n\t[\n\\3\n\t]);") + + ; 6: widen again + (widen) +) + diff -r 000000000000 -r a5a9c433f639 rm-logfiles --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/rm-logfiles Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,7 @@ +#! /bin/sh +#rm-logfiles: remove useless files from subdirectories +rm log */make*.log */make*.log.gz */make*.log.z +rm */test +rm */.*.thy.ML +rm */ex/.*.thy.ML +rm HOL/Subst/.*.thy.ML diff -r 000000000000 -r a5a9c433f639 src/CCL/CCL.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/CCL.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,362 @@ +(* Title: CCL/ccl + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For ccl.thy. +*) + +open CCL; + +val ccl_data_defs = [apply_def,fix_def]; + +(*** Simplifier for pre-order and equality ***) + +structure CCL_SimpData : SIMP_DATA = + struct + val refl_thms = [refl, po_refl, iff_refl] + val trans_thms = [trans, iff_trans, po_trans] + val red1 = iffD1 + val red2 = iffD2 + val mk_rew_rules = mk_rew_rules + val case_splits = [] (*NO IF'S!*) + val norm_thms = norm_thms + val subst_thms = [subst]; + val dest_red = dest_red + end; + +structure CCL_Simp = SimpFun(CCL_SimpData); +open CCL_Simp; + +val auto_ss = empty_ss setauto (fn hyps => ares_tac (TrueI::hyps)); + +val po_refl_iff_T = make_iff_T po_refl; + +val CCL_ss = auto_ss addcongs (FOL_congs @ set_congs) + addrews ([po_refl_iff_T] @ FOL_rews @ mem_rews); + +(*** Congruence Rules ***) + +(*similar to AP_THM in Gordon's HOL*) +val fun_cong = prove_goal CCL.thy "(f::'a=>'b) = g ==> f(x)=g(x)" + (fn [prem] => [rtac (prem RS subst) 1, rtac refl 1]); + +(*similar to AP_TERM in Gordon's HOL and FOL's subst_context*) +val arg_cong = prove_goal CCL.thy "x=y ==> f(x)=f(y)" + (fn [prem] => [rtac (prem RS subst) 1, rtac refl 1]); + +goal CCL.thy "(ALL x. f(x) = g(x)) --> (%x.f(x)) = (%x.g(x))"; +by (SIMP_TAC (CCL_ss addrews [eq_iff]) 1); +by (fast_tac (set_cs addIs [po_abstractn]) 1); +val abstractn = standard (allI RS (result() RS mp)); + +fun type_of_terms (Const("Trueprop",_) $ + (Const("op =",(Type ("fun", [t,_]))) $ _ $ _)) = t; + +fun abs_prems thm = + let fun do_abs n thm (Type ("fun", [_,t])) = do_abs n (abstractn RSN (n,thm)) t + | do_abs n thm _ = thm + fun do_prems n [] thm = thm + | do_prems n (x::xs) thm = do_prems (n+1) xs (do_abs n thm (type_of_terms x)); + in do_prems 1 (prems_of thm) thm + end; + +fun ccl_mk_congs thy cs = map abs_prems (mk_congs thy cs); + +val ccl_congs = ccl_mk_congs CCL.thy + ["op [=","SIM","POgen","EQgen","pair","lambda","case","op `","fix"]; + +val caseBs = [caseBtrue,caseBfalse,caseBpair,caseBlam,caseBbot]; + +(*** Termination and Divergence ***) + +goalw CCL.thy [Trm_def,Dvg_def] "Trm(t) <-> ~ t = bot"; +br iff_refl 1; +val Trm_iff = result(); + +goalw CCL.thy [Trm_def,Dvg_def] "Dvg(t) <-> t = bot"; +br iff_refl 1; +val Dvg_iff = result(); + +(*** Constructors are injective ***) + +val prems = goal CCL.thy + "[| x=a; y=b; x=y |] ==> a=b"; +by (REPEAT (SOMEGOAL (ares_tac (prems@[box_equals])))); +val eq_lemma = result(); + +fun mk_inj_rl thy rews congs s = + let fun mk_inj_lemmas r = ([arg_cong] RL [(r RS (r RS eq_lemma))]); + val inj_lemmas = flat (map mk_inj_lemmas rews); + val tac = REPEAT (ares_tac [iffI,allI,conjI] 1 ORELSE + eresolve_tac inj_lemmas 1 ORELSE + ASM_SIMP_TAC (CCL_ss addrews rews + addcongs congs) 1) + in prove_goal thy s (fn _ => [tac]) + end; + +val ccl_injs = map (mk_inj_rl CCL.thy caseBs ccl_congs) + [" = <-> (a=a' & b=b')", + "(lam x.b(x) = lam x.b'(x)) <-> ((ALL z.b(z)=b'(z)))"]; + +val pair_inject = ((hd ccl_injs) RS iffD1) RS conjE; + +(*** Constructors are distinct ***) + +local + fun pairs_of f x [] = [] + | pairs_of f x (y::ys) = (f x y) :: (f y x) :: (pairs_of f x ys); + + fun mk_combs ff [] = [] + | mk_combs ff (x::xs) = (pairs_of ff x xs) @ mk_combs ff xs; + +(* Doesn't handle binder types correctly *) + fun saturate thy sy name = + let fun arg_str 0 a s = s + | arg_str 1 a s = "(" ^ a ^ "a" ^ s ^ ")" + | arg_str n a s = arg_str (n-1) a ("," ^ a ^ (chr((ord "a")+n-1)) ^ s); + val sg = sign_of thy; + val T = case Sign.Symtab.lookup(#const_tab(Sign.rep_sg sg),sy) of + None => error(sy^" not declared") | Some(T) => T; + val arity = length (fst (strip_type T)); + in sy ^ (arg_str arity name "") end; + + fun mk_thm_str thy a b = "~ " ^ (saturate thy a "a") ^ " = " ^ (saturate thy b "b"); + + val lemma = prove_goal CCL.thy "t=t' --> case(t,b,c,d,e) = case(t',b,c,d,e)" + (fn _ => [SIMP_TAC (CCL_ss addcongs ccl_congs) 1]) RS mp; + fun mk_lemma (ra,rb) = [lemma] RL [ra RS (rb RS eq_lemma)] RL + [distinctness RS notE,sym RS (distinctness RS notE)]; +in + fun mk_lemmas rls = flat (map mk_lemma (mk_combs pair rls)); + fun mk_dstnct_rls thy xs = mk_combs (mk_thm_str thy) xs; +end; + + +val caseB_lemmas = mk_lemmas caseBs; + +val ccl_dstncts = + let fun mk_raw_dstnct_thm rls s = + prove_goal CCL.thy s (fn _=> [rtac notI 1,eresolve_tac rls 1]) + in map (mk_raw_dstnct_thm caseB_lemmas) + (mk_dstnct_rls CCL.thy ["bot","true","false","pair","lambda"]) end; + +fun mk_dstnct_thms thy defs inj_rls xs = + let fun mk_dstnct_thm rls s = prove_goalw thy defs s + (fn _ => [SIMP_TAC (CCL_ss addrews (rls@inj_rls)) 1]) + in map (mk_dstnct_thm ccl_dstncts) (mk_dstnct_rls thy xs) end; + +fun mkall_dstnct_thms thy defs i_rls xss = flat (map (mk_dstnct_thms thy defs i_rls) xss); + +(*** Rewriting and Proving ***) + +fun XH_to_I rl = rl RS iffD2; +fun XH_to_D rl = rl RS iffD1; +val XH_to_E = make_elim o XH_to_D; +val XH_to_Is = map XH_to_I; +val XH_to_Ds = map XH_to_D; +val XH_to_Es = map XH_to_E; + +val ccl_rews = caseBs @ ccl_injs @ ccl_dstncts; +val ccl_ss = CCL_ss addrews ccl_rews addcongs ccl_congs; + +val ccl_cs = set_cs addSEs (pair_inject::(ccl_dstncts RL [notE])) + addSDs (XH_to_Ds ccl_injs); + +(****** Facts from gfp Definition of [= and = ******) + +val major::prems = goal Set.thy "[| A=B; a:B <-> P |] ==> a:A <-> P"; +brs (prems RL [major RS ssubst]) 1; +val XHlemma1 = result(); + +goal CCL.thy "(P(t,t') <-> Q) --> ( : {p.EX t t'.p= & P(t,t')} <-> Q)"; +by (fast_tac ccl_cs 1); +val XHlemma2 = result() RS mp; + +(*** Pre-Order ***) + +goalw CCL.thy [POgen_def,SIM_def] "mono(%X.POgen(X))"; +br monoI 1; +by (safe_tac ccl_cs); +by (REPEAT_SOME (resolve_tac [exI,conjI,refl])); +by (ALLGOALS (SIMP_TAC ccl_ss)); +by (ALLGOALS (fast_tac set_cs)); +val POgen_mono = result(); + +goalw CCL.thy [POgen_def,SIM_def] + " : POgen(R) <-> t= bot | (t=true & t'=true) | (t=false & t'=false) | \ +\ (EX a a' b b'.t= & t'= & : R & : R) | \ +\ (EX f f'.t=lam x.f(x) & t'=lam x.f'(x) & (ALL x. : R))"; +br (iff_refl RS XHlemma2) 1; +val POgenXH = result(); + +goal CCL.thy + "t [= t' <-> t=bot | (t=true & t'=true) | (t=false & t'=false) | \ +\ (EX a a' b b'.t= & t'= & a [= a' & b [= b') | \ +\ (EX f f'.t=lam x.f(x) & t'=lam x.f'(x) & (ALL x.f(x) [= f'(x)))"; +by (SIMP_TAC (ccl_ss addrews [PO_iff]) 1); +br (rewrite_rule [POgen_def,SIM_def] + (POgen_mono RS (PO_def RS def_gfp_Tarski) RS XHlemma1)) 1; +br (iff_refl RS XHlemma2) 1; +val poXH = result(); + +goal CCL.thy "bot [= b"; +br (poXH RS iffD2) 1; +by (SIMP_TAC ccl_ss 1); +val po_bot = result(); + +goal CCL.thy "a [= bot --> a=bot"; +br impI 1; +bd (poXH RS iffD1) 1; +be rev_mp 1; +by (SIMP_TAC ccl_ss 1); +val bot_poleast = result() RS mp; + +goal CCL.thy " [= <-> a [= a' & b [= b'"; +br (poXH RS iff_trans) 1; +by (SIMP_TAC ccl_ss 1); +by (fast_tac ccl_cs 1); +val po_pair = result(); + +goal CCL.thy "lam x.f(x) [= lam x.f'(x) <-> (ALL x. f(x) [= f'(x))"; +br (poXH RS iff_trans) 1; +by (SIMP_TAC ccl_ss 1); +by (REPEAT (ares_tac [iffI,allI] 1 ORELSE eresolve_tac [exE,conjE] 1)); +by (ASM_SIMP_TAC ccl_ss 1); +by (fast_tac ccl_cs 1); +val po_lam = result(); + +val ccl_porews = [po_bot,po_pair,po_lam]; + +val [p1,p2,p3,p4,p5] = goal CCL.thy + "[| t [= t'; a [= a'; b [= b'; !!x y.c(x,y) [= c'(x,y); \ +\ !!u.d(u) [= d'(u) |] ==> case(t,a,b,c,d) [= case(t',a',b',c',d')"; +br (p1 RS po_cong RS po_trans) 1; +br (p2 RS po_cong RS po_trans) 1; +br (p3 RS po_cong RS po_trans) 1; +br (p4 RS po_abstractn RS po_abstractn RS po_cong RS po_trans) 1; +by (res_inst_tac [("f1","%d.case(t',a',b',c',d)")] + (p5 RS po_abstractn RS po_cong RS po_trans) 1); +br po_refl 1; +val case_pocong = result(); + +val [p1,p2] = goalw CCL.thy ccl_data_defs + "[| f [= f'; a [= a' |] ==> f ` a [= f' ` a'"; +by (REPEAT (ares_tac [po_refl,case_pocong,p1,p2 RS po_cong] 1)); +val apply_pocong = result(); + + +val prems = goal CCL.thy "~ lam x.b(x) [= bot"; +br notI 1; +bd bot_poleast 1; +be (distinctness RS notE) 1; +val npo_lam_bot = result(); + +val eq1::eq2::prems = goal CCL.thy + "[| x=a; y=b; x[=y |] ==> a[=b"; +br (eq1 RS subst) 1; +br (eq2 RS subst) 1; +brs prems 1; +val po_lemma = result(); + +goal CCL.thy "~ [= lam x.f(x)"; +br notI 1; +br (npo_lam_bot RS notE) 1; +be (case_pocong RS (caseBlam RS (caseBpair RS po_lemma))) 1; +by (REPEAT (resolve_tac [po_refl,npo_lam_bot] 1)); +val npo_pair_lam = result(); + +goal CCL.thy "~ lam x.f(x) [= "; +br notI 1; +br (npo_lam_bot RS notE) 1; +be (case_pocong RS (caseBpair RS (caseBlam RS po_lemma))) 1; +by (REPEAT (resolve_tac [po_refl,npo_lam_bot] 1)); +val npo_lam_pair = result(); + +fun mk_thm s = prove_goal CCL.thy s (fn _ => + [rtac notI 1,dtac case_pocong 1,etac rev_mp 5, + ALLGOALS (SIMP_TAC ccl_ss), + REPEAT (resolve_tac [po_refl,npo_lam_bot] 1)]); + +val npo_rls = [npo_pair_lam,npo_lam_pair] @ map mk_thm + ["~ true [= false", "~ false [= true", + "~ true [= ", "~ [= true", + "~ true [= lam x.f(x)","~ lam x.f(x) [= true", + "~ false [= ", "~ [= false", + "~ false [= lam x.f(x)","~ lam x.f(x) [= false"]; + +(* Coinduction for [= *) + +val prems = goal CCL.thy "[| : R; R <= POgen(R) |] ==> t [= u"; +br (PO_def RS def_coinduct RS (PO_iff RS iffD2)) 1; +by (REPEAT (ares_tac prems 1)); +val po_coinduct = result(); + +fun po_coinduct_tac s i = res_inst_tac [("R",s)] po_coinduct i; + +(*************** EQUALITY *******************) + +goalw CCL.thy [EQgen_def,SIM_def] "mono(%X.EQgen(X))"; +br monoI 1; +by (safe_tac set_cs); +by (REPEAT_SOME (resolve_tac [exI,conjI,refl])); +by (ALLGOALS (SIMP_TAC ccl_ss)); +by (ALLGOALS (fast_tac set_cs)); +val EQgen_mono = result(); + +goalw CCL.thy [EQgen_def,SIM_def] + " : EQgen(R) <-> (t=bot & t'=bot) | (t=true & t'=true) | \ +\ (t=false & t'=false) | \ +\ (EX a a' b b'.t= & t'= & : R & : R) | \ +\ (EX f f'.t=lam x.f(x) & t'=lam x.f'(x) & (ALL x. : R))"; +br (iff_refl RS XHlemma2) 1; +val EQgenXH = result(); + +goal CCL.thy + "t=t' <-> (t=bot & t'=bot) | (t=true & t'=true) | (t=false & t'=false) | \ +\ (EX a a' b b'.t= & t'= & a=a' & b=b') | \ +\ (EX f f'.t=lam x.f(x) & t'=lam x.f'(x) & (ALL x.f(x)=f'(x)))"; +by (subgoal_tac + " : EQ <-> (t=bot & t'=bot) | (t=true & t'=true) | (t=false & t'=false) | \ +\ (EX a a' b b'.t= & t'= & : EQ & : EQ) | \ +\ (EX f f'.t=lam x.f(x) & t'=lam x.f'(x) & (ALL x. : EQ))" 1); +be rev_mp 1; +by (SIMP_TAC (CCL_ss addrews [EQ_iff RS iff_sym]) 1); +br (rewrite_rule [EQgen_def,SIM_def] + (EQgen_mono RS (EQ_def RS def_gfp_Tarski) RS XHlemma1)) 1; +br (iff_refl RS XHlemma2) 1; +val eqXH = result(); + +val prems = goal CCL.thy "[| : R; R <= EQgen(R) |] ==> t = u"; +br (EQ_def RS def_coinduct RS (EQ_iff RS iffD2)) 1; +by (REPEAT (ares_tac prems 1)); +val eq_coinduct = result(); + +val prems = goal CCL.thy + "[| : R; R <= EQgen(lfp(%x.EQgen(x) Un R Un EQ)) |] ==> t = u"; +br (EQ_def RS def_coinduct3 RS (EQ_iff RS iffD2)) 1; +by (REPEAT (ares_tac (EQgen_mono::prems) 1)); +val eq_coinduct3 = result(); + +fun eq_coinduct_tac s i = res_inst_tac [("R",s)] eq_coinduct i; +fun eq_coinduct3_tac s i = res_inst_tac [("R",s)] eq_coinduct3 i; + +(*** Untyped Case Analysis and Other Facts ***) + +goalw CCL.thy [apply_def] "(EX f.t=lam x.f(x)) --> t = lam x.(t ` x)"; +by (safe_tac ccl_cs); +by (SIMP_TAC ccl_ss 1); +val cond_eta = result() RS mp; + +goal CCL.thy "(t=bot) | (t=true) | (t=false) | (EX a b.t=) | (EX f.t=lam x.f(x))"; +by (cut_facts_tac [refl RS (eqXH RS iffD1)] 1); +by (fast_tac set_cs 1); +val exhaustion = result(); + +val prems = goal CCL.thy + "[| P(bot); P(true); P(false); !!x y.P(); !!b.P(lam x.b(x)) |] ==> P(t)"; +by (cut_facts_tac [exhaustion] 1); +by (REPEAT_SOME (ares_tac prems ORELSE' eresolve_tac [disjE,exE,ssubst])); +val term_case = result(); + +fun term_case_tac a i = res_inst_tac [("t",a)] term_case i; diff -r 000000000000 -r a5a9c433f639 src/CCL/CCL.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/CCL.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,148 @@ +(* Title: CCL/ccl.thy + ID: $Id$ + Author: Martin Coen + Copyright 1993 University of Cambridge + +Classical Computational Logic for Untyped Lambda Calculus with reduction to +weak head-normal form. + +Based on FOL extended with set collection, a primitive higher-order logic. +HOL is too strong - descriptions prevent a type of programs being defined +which contains only executable terms. +*) + +CCL = Gfp + + +classes prog < term + +default prog + +types i 0 + +arities + i :: prog + fun :: (prog,prog)prog + +consts + (*** Evaluation Judgement ***) + "--->" :: "[i,i]=>prop" (infixl 20) + + (*** Bisimulations for pre-order and equality ***) + "[=" :: "['a,'a]=>o" (infixl 50) + SIM :: "[i,i,i set]=>o" + POgen,EQgen :: "i set => i set" + PO,EQ :: "i set" + + (*** Term Formers ***) + true,false :: "i" + pair :: "[i,i]=>i" ("(1<_,/_>)") + lambda :: "(i=>i)=>i" (binder "lam " 55) + case :: "[i,i,i,[i,i]=>i,(i=>i)=>i]=>i" + "`" :: "[i,i]=>i" (infixl 56) + bot :: "i" + fix :: "(i=>i)=>i" + + (*** Defined Predicates ***) + Trm,Dvg :: "i => o" + +rules + + (******* EVALUATION SEMANTICS *******) + + (** This is the evaluation semantics from which the axioms below were derived. **) + (** It is included here just as an evaluator for FUN and has no influence on **) + (** inference in the theory CCL. **) + + trueV "true ---> true" + falseV "false ---> false" + pairV " ---> " + lamV "lam x.b(x) ---> lam x.b(x)" + caseVtrue "[| t ---> true; d ---> c |] ==> case(t,d,e,f,g) ---> c" + caseVfalse "[| t ---> false; e ---> c |] ==> case(t,d,e,f,g) ---> c" + caseVpair "[| t ---> ; f(a,b) ---> c |] ==> case(t,d,e,f,g) ---> c" + caseVlam "[| t ---> lam x.b(x); g(b) ---> c |] ==> case(t,d,e,f,g) ---> c" + + (*** Properties of evaluation: note that "t ---> c" impies that c is canonical ***) + + canonical "[| t ---> c; c==true ==> u--->v; \ +\ c==false ==> u--->v; \ +\ !!a b.c== ==> u--->v; \ +\ !!f.c==lam x.f(x) ==> u--->v |] ==> \ +\ u--->v" + + (* Should be derivable - but probably a bitch! *) + substitute "[| a==a'; t(a)--->c(a) |] ==> t(a')--->c(a')" + + (************** LOGIC ***************) + + (*** Definitions used in the following rules ***) + + apply_def "f ` t == case(f,bot,bot,%x y.bot,%u.u(t))" + bot_def "bot == (lam x.x`x)`(lam x.x`x)" + fix_def "fix(f) == (lam x.f(x`x))`(lam x.f(x`x))" + + (* The pre-order ([=) is defined as a simulation, and behavioural equivalence (=) *) + (* as a bisimulation. They can both be expressed as (bi)simulations up to *) + (* behavioural equivalence (ie the relations PO and EQ defined below). *) + + SIM_def + "SIM(t,t',R) == (t=true & t'=true) | (t=false & t'=false) | \ +\ (EX a a' b b'.t= & t'= & : R & : R) | \ +\ (EX f f'.t=lam x.f(x) & t'=lam x.f'(x) & (ALL x. : R))" + + POgen_def "POgen(R) == {p. EX t t'. p= & (t = bot | SIM(t,t',R))}" + EQgen_def "EQgen(R) == {p. EX t t'. p= & (t = bot & t' = bot | SIM(t,t',R))}" + + PO_def "PO == gfp(POgen)" + EQ_def "EQ == gfp(EQgen)" + + (*** Rules ***) + + (** Partial Order **) + + po_refl "a [= a" + po_trans "[| a [= b; b [= c |] ==> a [= c" + po_cong "a [= b ==> f(a) [= f(b)" + + (* Extend definition of [= to program fragments of higher type *) + po_abstractn "(!!x. f(x) [= g(x)) ==> (%x.f(x)) [= (%x.g(x))" + + (** Equality - equivalence axioms inherited from FOL.thy **) + (** - congruence of "=" is axiomatised implicitly **) + + eq_iff "t = t' <-> t [= t' & t' [= t" + + (** Properties of canonical values given by greatest fixed point definitions **) + + PO_iff "t [= t' <-> : PO" + EQ_iff "t = t' <-> : EQ" + + (** Behaviour of non-canonical terms (ie case) given by the following beta-rules **) + + caseBtrue "case(true,d,e,f,g) = d" + caseBfalse "case(false,d,e,f,g) = e" + caseBpair "case(,d,e,f,g) = f(a,b)" + caseBlam "case(lam x.b(x),d,e,f,g) = g(b)" + caseBbot "case(bot,d,e,f,g) = bot" (* strictness *) + + (** The theory is non-trivial **) + distinctness "~ lam x.b(x) = bot" + + (*** Definitions of Termination and Divergence ***) + + Dvg_def "Dvg(t) == t = bot" + Trm_def "Trm(t) == ~ Dvg(t)" + +end + + +(* +Would be interesting to build a similar theory for a typed programming language: + ie. true :: bool, fix :: ('a=>'a)=>'a etc...... + +This is starting to look like LCF. +What are the advantages of this approach? + - less axiomatic + - wfd induction / coinduction and fixed point induction available + +*) diff -r 000000000000 -r a5a9c433f639 src/CCL/Fix.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Fix.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,202 @@ +(* Title: CCL/fix + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For fix.thy. +*) + +open Fix; + +val prems = goalw Fix.thy [INCL_def] + "[| !!x.P(x) <-> Q(x) |] ==> INCL(%x.P(x)) <-> INCL(%x.Q(x))"; +by (REPEAT (ares_tac ([refl] @ FOL_congs @ set_congs @ prems) 1)); +val INCL_cong = result(); + +val fix_congs = [INCL_cong] @ ccl_mk_congs Fix.thy ["napply"]; + +(*** Fixed Point Induction ***) + +val [base,step,incl] = goalw Fix.thy [INCL_def] + "[| P(bot); !!x.P(x) ==> P(f(x)); INCL(P) |] ==> P(fix(f))"; +br (incl RS spec RS mp) 1; +by (rtac (Nat_ind RS ballI) 1 THEN atac 1); +by (ALLGOALS (SIMP_TAC term_ss)); +by (REPEAT (ares_tac [base,step] 1)); +val fix_ind = result(); + +(*** Inclusive Predicates ***) + +val prems = goalw Fix.thy [INCL_def] + "INCL(P) <-> (ALL f. (ALL n:Nat. P(f ^ n ` bot)) --> P(fix(f)))"; +br iff_refl 1; +val inclXH = result(); + +val prems = goal Fix.thy + "[| !!f.ALL n:Nat.P(f^n`bot) ==> P(fix(f)) |] ==> INCL(%x.P(x))"; +by (fast_tac (term_cs addIs (prems @ [XH_to_I inclXH])) 1); +val inclI = result(); + +val incl::prems = goal Fix.thy + "[| INCL(P); !!n.n:Nat ==> P(f^n`bot) |] ==> P(fix(f))"; +by (fast_tac (term_cs addIs ([ballI RS (incl RS (XH_to_D inclXH) RS spec RS mp)] + @ prems)) 1); +val inclD = result(); + +val incl::prems = goal Fix.thy + "[| INCL(P); (ALL n:Nat.P(f^n`bot))-->P(fix(f)) ==> R |] ==> R"; +by (fast_tac (term_cs addIs ([incl RS inclD] @ prems)) 1); +val inclE = result(); + +val fix_ss = term_ss addcongs fix_congs; + +(*** Lemmas for Inclusive Predicates ***) + +goal Fix.thy "INCL(%x.~ a(x) [= t)"; +br inclI 1; +bd bspec 1; +br zeroT 1; +be contrapos 1; +br po_trans 1; +ba 2; +br (napplyBzero RS ssubst) 1; +by (rtac po_cong 1 THEN rtac po_bot 1); +val npo_INCL = result(); + +val prems = goal Fix.thy "[| INCL(P); INCL(Q) |] ==> INCL(%x.P(x) & Q(x))"; +by (fast_tac (set_cs addSIs ([inclI] @ (prems RL [inclD]))) 1);; +val conj_INCL = result(); + +val prems = goal Fix.thy "[| !!a.INCL(P(a)) |] ==> INCL(%x.ALL a.P(a,x))"; +by (fast_tac (set_cs addSIs ([inclI] @ (prems RL [inclD]))) 1);; +val all_INCL = result(); + +val prems = goal Fix.thy "[| !!a.a:A ==> INCL(P(a)) |] ==> INCL(%x.ALL a:A.P(a,x))"; +by (fast_tac (set_cs addSIs ([inclI] @ (prems RL [inclD]))) 1);; +val ball_INCL = result(); + +goal Fix.thy "INCL(%x.a(x) = b(x)::'a::prog)"; +by (SIMP_TAC (fix_ss addrews [eq_iff]) 1); +by (REPEAT (resolve_tac [conj_INCL,po_INCL] 1)); +val eq_INCL = result(); + +(*** Derivation of Reachability Condition ***) + +(* Fixed points of idgen *) + +goal Fix.thy "idgen(fix(idgen)) = fix(idgen)"; +br (fixB RS sym) 1; +val fix_idgenfp = result(); + +goalw Fix.thy [idgen_def] "idgen(lam x.x) = lam x.x"; +by (SIMP_TAC term_ss 1); +br (term_case RS allI) 1; +by (ALLGOALS (SIMP_TAC term_ss)); +val id_idgenfp = result(); + +(* All fixed points are lam-expressions *) + +val [prem] = goal Fix.thy "idgen(d) = d ==> d = lam x.?f(x)"; +br (prem RS subst) 1; +bw idgen_def; +br refl 1; +val idgenfp_lam = result(); + +(* Lemmas for rewriting fixed points of idgen *) + +val prems = goalw Fix.thy [idgen_def] + "[| a = b; a ` t = u |] ==> b ` t = u"; +by (SIMP_TAC (term_ss addrews (prems RL [sym])) 1); +val l_lemma= result(); + +val idgen_lemmas = + let fun mk_thm s = prove_goalw Fix.thy [idgen_def] s + (fn [prem] => [rtac (prem RS l_lemma) 1,SIMP_TAC term_ss 1]) + in map mk_thm + [ "idgen(d) = d ==> d ` bot = bot", + "idgen(d) = d ==> d ` true = true", + "idgen(d) = d ==> d ` false = false", + "idgen(d) = d ==> d ` = ", + "idgen(d) = d ==> d ` (lam x.f(x)) = lam x.d ` f(x)"] + end; + +(* Proof of Reachability law - show that fix and lam x.x both give LEAST fixed points + of idgen and hence are they same *) + +val [p1,p2,p3] = goal CCL.thy + "[| ALL x.t ` x [= u ` x; EX f.t=lam x.f(x); EX f.u=lam x.f(x) |] ==> t [= u"; +br (p2 RS cond_eta RS ssubst) 1; +br (p3 RS cond_eta RS ssubst) 1; +br (p1 RS (po_lam RS iffD2)) 1; +val po_eta = result(); + +val [prem] = goalw Fix.thy [idgen_def] "idgen(d) = d ==> d = lam x.?f(x)"; +br (prem RS subst) 1; +br refl 1; +val po_eta_lemma = result(); + +val [prem] = goal Fix.thy + "idgen(d) = d ==> \ +\ {p.EX a b.p= & (EX t.a=fix(idgen) ` t & b = d ` t)} <= \ +\ POgen({p.EX a b.p= & (EX t.a=fix(idgen) ` t & b = d ` t)})"; +by (REPEAT (step_tac term_cs 1)); +by (term_case_tac "t" 1); +by (ALLGOALS (SIMP_TAC (term_ss addrews (POgenXH::([prem,fix_idgenfp] RL idgen_lemmas))))); +by (ALLGOALS (fast_tac set_cs)); +val lemma1 = result(); + +val [prem] = goal Fix.thy + "idgen(d) = d ==> fix(idgen) [= d"; +br (allI RS po_eta) 1; +br (lemma1 RSN(2,po_coinduct)) 1; +by (ALLGOALS (fast_tac (term_cs addIs [prem,po_eta_lemma,fix_idgenfp]))); +val fix_least_idgen = result(); + +val [prem] = goal Fix.thy + "idgen(d) = d ==> \ +\ {p.EX a b.p= & b = d ` a} <= POgen({p.EX a b.p= & b = d ` a})"; +by (REPEAT (step_tac term_cs 1)); +by (term_case_tac "a" 1); +by (ALLGOALS (SIMP_TAC (term_ss addrews (POgenXH::([prem] RL idgen_lemmas))))); +by (ALLGOALS (fast_tac set_cs)); +val lemma2 = result(); + +val [prem] = goal Fix.thy + "idgen(d) = d ==> lam x.x [= d"; +br (allI RS po_eta) 1; +br (lemma2 RSN(2,po_coinduct)) 1; +by (SIMP_TAC term_ss 1); +by (ALLGOALS (fast_tac (term_cs addIs [prem,po_eta_lemma,fix_idgenfp]))); +val id_least_idgen = result(); + +goal Fix.thy "fix(idgen) = lam x.x"; +by (fast_tac (term_cs addIs [eq_iff RS iffD2, + id_idgenfp RS fix_least_idgen, + fix_idgenfp RS id_least_idgen]) 1); +val reachability = result(); + +(********) + +val [prem] = goal Fix.thy "f = lam x.x ==> f`t = t"; +br (prem RS sym RS subst) 1; +br applyB 1; +val id_apply = result(); + +val prems = goal Fix.thy + "[| P(bot); P(true); P(false); \ +\ !!x y.[| P(x); P(y) |] ==> P(); \ +\ !!u.(!!x.P(u(x))) ==> P(lam x.u(x)); INCL(P) |] ==> \ +\ P(t)"; +br (reachability RS id_apply RS subst) 1; +by (res_inst_tac [("x","t")] spec 1); +br fix_ind 1; +bw idgen_def; +by (REPEAT_SOME (ares_tac [allI])); +br (applyBbot RS ssubst) 1; +brs prems 1; +br (applyB RS ssubst )1; +by (res_inst_tac [("t","xa")] term_case 1); +by (ALLGOALS (SIMP_TAC term_ss)); +by (ALLGOALS (fast_tac (term_cs addIs ([all_INCL,INCL_subst] @ prems)))); +val term_ind = result(); + diff -r 000000000000 -r a5a9c433f639 src/CCL/Fix.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Fix.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,26 @@ +(* Title: CCL/Lazy/fix.thy + ID: $Id$ + Author: Martin Coen + Copyright 1993 University of Cambridge + +Tentative attempt at including fixed point induction. +Justified by Smith. +*) + +Fix = Type + + +consts + + idgen :: "[i]=>i" + INCL :: "[i=>o]=>o" + +rules + + idgen_def + "idgen(f) == lam t.case(t,true,false,%x y.,%u.lam x.f ` u(x))" + + INCL_def "INCL(%x.P(x)) == (ALL f.(ALL n:Nat.P(f^n`bot)) --> P(fix(f)))" + po_INCL "INCL(%x.a(x) [= b(x))" + INCL_subst "INCL(P) ==> INCL(%x.P((g::i=>i)(x)))" + +end diff -r 000000000000 -r a5a9c433f639 src/CCL/Gfp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Gfp.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,133 @@ +(* Title: CCL/gfp + ID: $Id$ + +Modified version of + Title: HOL/gfp + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For gfp.thy. The Knaster-Tarski Theorem for greatest fixed points. +*) + +open Gfp; + +(*** Proof of Knaster-Tarski Theorem using gfp ***) + +(* gfp(f) is the least upper bound of {u. u <= f(u)} *) + +val prems = goalw Gfp.thy [gfp_def] "[| A <= f(A) |] ==> A <= gfp(f)"; +by (rtac (CollectI RS Union_upper) 1); +by (resolve_tac prems 1); +val gfp_upperbound = result(); + +val prems = goalw Gfp.thy [gfp_def] + "[| !!u. u <= f(u) ==> u<=A |] ==> gfp(f) <= A"; +by (REPEAT (ares_tac ([Union_least]@prems) 1)); +by (etac CollectD 1); +val gfp_least = result(); + +val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) <= f(gfp(f))"; +by (EVERY1 [rtac gfp_least, rtac subset_trans, atac, + rtac (mono RS monoD), rtac gfp_upperbound, atac]); +val gfp_lemma2 = result(); + +val [mono] = goal Gfp.thy "mono(f) ==> f(gfp(f)) <= gfp(f)"; +by (EVERY1 [rtac gfp_upperbound, rtac (mono RS monoD), + rtac gfp_lemma2, rtac mono]); +val gfp_lemma3 = result(); + +val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) = f(gfp(f))"; +by (REPEAT (resolve_tac [equalityI,gfp_lemma2,gfp_lemma3,mono] 1)); +val gfp_Tarski = result(); + +(*** Coinduction rules for greatest fixed points ***) + +(*weak version*) +val prems = goal Gfp.thy + "[| a: A; A <= f(A) |] ==> a : gfp(f)"; +by (rtac (gfp_upperbound RS subsetD) 1); +by (REPEAT (ares_tac prems 1)); +val coinduct = result(); + +val [prem,mono] = goal Gfp.thy + "[| A <= f(A) Un gfp(f); mono(f) |] ==> \ +\ A Un gfp(f) <= f(A Un gfp(f))"; +by (rtac subset_trans 1); +by (rtac (mono RS mono_Un) 2); +by (rtac (mono RS gfp_Tarski RS subst) 1); +by (rtac (prem RS Un_least) 1); +by (rtac Un_upper2 1); +val coinduct2_lemma = result(); + +(*strong version, thanks to Martin Coen*) +val prems = goal Gfp.thy + "[| a: A; A <= f(A) Un gfp(f); mono(f) |] ==> a : gfp(f)"; +by (rtac (coinduct2_lemma RSN (2,coinduct)) 1); +by (REPEAT (resolve_tac (prems@[UnI1]) 1)); +val coinduct2 = result(); + +(*** Even Stronger version of coinduct [by Martin Coen] + - instead of the condition A <= f(A) + consider A <= (f(A) Un f(f(A)) ...) Un gfp(A) ***) + +val [prem] = goal Gfp.thy "mono(f) ==> mono(%x.f(x) Un A Un B)"; +by (REPEAT (ares_tac [subset_refl, monoI, Un_mono, prem RS monoD] 1)); +val coinduct3_mono_lemma= result(); + +val [prem,mono] = goal Gfp.thy + "[| A <= f(lfp(%x.f(x) Un A Un gfp(f))); mono(f) |] ==> \ +\ lfp(%x.f(x) Un A Un gfp(f)) <= f(lfp(%x.f(x) Un A Un gfp(f)))"; +by (rtac subset_trans 1); +br (mono RS coinduct3_mono_lemma RS lfp_lemma3) 1; +by (rtac (Un_least RS Un_least) 1); +br subset_refl 1; +br prem 1; +br (mono RS gfp_Tarski RS equalityD1 RS subset_trans) 1; +by (rtac (mono RS monoD) 1); +by (rtac (mono RS coinduct3_mono_lemma RS lfp_Tarski RS ssubst) 1); +by (rtac Un_upper2 1); +val coinduct3_lemma = result(); + +val prems = goal Gfp.thy + "[| a:A; A <= f(lfp(%x.f(x) Un A Un gfp(f))); mono(f) |] ==> a : gfp(f)"; +by (rtac (coinduct3_lemma RSN (2,coinduct)) 1); +brs (prems RL [coinduct3_mono_lemma RS lfp_Tarski RS ssubst]) 1; +br (UnI2 RS UnI1) 1; +by (REPEAT (resolve_tac prems 1)); +val coinduct3 = result(); + + +(** Definition forms of gfp_Tarski, to control unfolding **) + +val [rew,mono] = goal Gfp.thy "[| h==gfp(f); mono(f) |] ==> h = f(h)"; +by (rewtac rew); +by (rtac (mono RS gfp_Tarski) 1); +val def_gfp_Tarski = result(); + +val rew::prems = goal Gfp.thy + "[| h==gfp(f); a:A; A <= f(A) |] ==> a: h"; +by (rewtac rew); +by (REPEAT (ares_tac (prems @ [coinduct]) 1)); +val def_coinduct = result(); + +val rew::prems = goal Gfp.thy + "[| h==gfp(f); a:A; A <= f(A) Un h; mono(f) |] ==> a: h"; +by (rewtac rew); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct2]) 1)); +val def_coinduct2 = result(); + +val rew::prems = goal Gfp.thy + "[| h==gfp(f); a:A; A <= f(lfp(%x.f(x) Un A Un h)); mono(f) |] ==> a: h"; +by (rewtac rew); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct3]) 1)); +val def_coinduct3 = result(); + +(*Monotonicity of gfp!*) +val prems = goal Gfp.thy + "[| mono(f); !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)"; +by (rtac gfp_upperbound 1); +by (rtac subset_trans 1); +by (rtac gfp_lemma2 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +val gfp_mono = result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/Gfp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Gfp.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,14 @@ +(* Title: HOL/gfp.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Greatest fixed points +*) + +Gfp = Lfp + +consts gfp :: "['a set=>'a set] => 'a set" +rules + (*greatest fixed point*) + gfp_def "gfp(f) == Union({u. u <= f(u)})" +end diff -r 000000000000 -r a5a9c433f639 src/CCL/Hered.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Hered.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,196 @@ +(* Title: CCL/hered + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For hered.thy. +*) + +open Hered; + +fun type_of_terms (Const("Trueprop",_) $ (Const("op =",(Type ("fun", [t,_])))$_$_)) = t; + +val cong_rls = ccl_mk_congs Hered.thy ["HTTgen"]; + +(*** Hereditary Termination ***) + +goalw Hered.thy [HTTgen_def] "mono(%X.HTTgen(X))"; +br monoI 1; +by (fast_tac set_cs 1); +val HTTgen_mono = result(); + +goalw Hered.thy [HTTgen_def] + "t : HTTgen(A) <-> t=true | t=false | (EX a b.t= & a : A & b : A) | \ +\ (EX f.t=lam x.f(x) & (ALL x.f(x) : A))"; +by (fast_tac set_cs 1); +val HTTgenXH = result(); + +goal Hered.thy + "t : HTT <-> t=true | t=false | (EX a b.t= & a : HTT & b : HTT) | \ +\ (EX f.t=lam x.f(x) & (ALL x.f(x) : HTT))"; +br (rewrite_rule [HTTgen_def] + (HTTgen_mono RS (HTT_def RS def_gfp_Tarski) RS XHlemma1)) 1; +by (fast_tac set_cs 1); +val HTTXH = result(); + +(*** Introduction Rules for HTT ***) + +goal Hered.thy "~ bot : HTT"; +by (fast_tac (term_cs addDs [XH_to_D HTTXH]) 1); +val HTT_bot = result(); + +goal Hered.thy "true : HTT"; +by (fast_tac (term_cs addIs [XH_to_I HTTXH]) 1); +val HTT_true = result(); + +goal Hered.thy "false : HTT"; +by (fast_tac (term_cs addIs [XH_to_I HTTXH]) 1); +val HTT_false = result(); + +goal Hered.thy " : HTT <-> a : HTT & b : HTT"; +br (HTTXH RS iff_trans) 1; +by (fast_tac term_cs 1); +val HTT_pair = result(); + +goal Hered.thy "lam x.f(x) : HTT <-> (ALL x. f(x) : HTT)"; +br (HTTXH RS iff_trans) 1; +by (SIMP_TAC term_ss 1); +by (safe_tac term_cs); +by (ASM_SIMP_TAC term_ss 1); +by (fast_tac term_cs 1); +val HTT_lam = result(); + +local + val raw_HTTrews = [HTT_bot,HTT_true,HTT_false,HTT_pair,HTT_lam]; + fun mk_thm s = prove_goalw Hered.thy data_defs s (fn _ => + [SIMP_TAC (term_ss addrews raw_HTTrews) 1]); +in + val HTT_rews = raw_HTTrews @ + map mk_thm ["one : HTT", + "inl(a) : HTT <-> a : HTT", + "inr(b) : HTT <-> b : HTT", + "zero : HTT", + "succ(n) : HTT <-> n : HTT", + "[] : HTT", + "x.xs : HTT <-> x : HTT & xs : HTT"]; +end; + +val HTT_Is = HTT_rews @ (HTT_rews RL [iffD2]); + +(*** Coinduction for HTT ***) + +val prems = goal Hered.thy "[| t : R; R <= HTTgen(R) |] ==> t : HTT"; +br (HTT_def RS def_coinduct) 1; +by (REPEAT (ares_tac prems 1)); +val HTT_coinduct = result(); + +fun HTT_coinduct_tac s i = res_inst_tac [("R",s)] HTT_coinduct i; + +val prems = goal Hered.thy + "[| t : R; R <= HTTgen(lfp(%x. HTTgen(x) Un R Un HTT)) |] ==> t : HTT"; +br (HTTgen_mono RSN(3,HTT_def RS def_coinduct3)) 1; +by (REPEAT (ares_tac prems 1)); +val HTT_coinduct3 = result(); +val HTT_coinduct3_raw = rewrite_rule [HTTgen_def] HTT_coinduct3; + +fun HTT_coinduct3_tac s i = res_inst_tac [("R",s)] HTT_coinduct3 i; + +val HTTgenIs = map (mk_genIs Hered.thy data_defs HTTgenXH HTTgen_mono) + ["true : HTTgen(R)", + "false : HTTgen(R)", + "[| a : R; b : R |] ==> : HTTgen(R)", + "[| !!x. b(x) : R |] ==> lam x.b(x) : HTTgen(R)", + "one : HTTgen(R)", + "a : lfp(%x. HTTgen(x) Un R Un HTT) ==> \ +\ inl(a) : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))", + "b : lfp(%x. HTTgen(x) Un R Un HTT) ==> \ +\ inr(b) : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))", + "zero : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))", + "n : lfp(%x. HTTgen(x) Un R Un HTT) ==> \ +\ succ(n) : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))", + "[] : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))", + "[| h : lfp(%x. HTTgen(x) Un R Un HTT); t : lfp(%x. HTTgen(x) Un R Un HTT) |] ==>\ +\ h.t : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))"]; + +(*** Formation Rules for Types ***) + +goal Hered.thy "Unit <= HTT"; +by (SIMP_TAC (CCL_ss addrews ([subsetXH,UnitXH] @ HTT_rews)) 1); +val UnitF = result(); + +goal Hered.thy "Bool <= HTT"; +by (SIMP_TAC (CCL_ss addrews ([subsetXH,BoolXH] @ HTT_rews)) 1); +by (fast_tac (set_cs addIs HTT_Is @ (prems RL [subsetD])) 1); +val BoolF = result(); + +val prems = goal Hered.thy "[| A <= HTT; B <= HTT |] ==> A + B <= HTT"; +by (SIMP_TAC (CCL_ss addrews ([subsetXH,PlusXH] @ HTT_rews)) 1); +by (fast_tac (set_cs addIs HTT_Is @ (prems RL [subsetD])) 1); +val PlusF = result(); + +val prems = goal Hered.thy + "[| A <= HTT; !!x.x:A ==> B(x) <= HTT |] ==> SUM x:A.B(x) <= HTT"; +by (SIMP_TAC (CCL_ss addrews ([subsetXH,SgXH] @ HTT_rews)) 1); +by (fast_tac (set_cs addIs HTT_Is @ (prems RL [subsetD])) 1); +val SigmaF = result(); + +(*** Formation Rules for Recursive types - using coinduction these only need ***) +(*** exhaution rule for type-former ***) + +(*Proof by induction - needs induction rule for type*) +goal Hered.thy "Nat <= HTT"; +by (SIMP_TAC (term_ss addrews [subsetXH]) 1); +by (safe_tac set_cs); +be Nat_ind 1; +by (ALLGOALS (fast_tac (set_cs addIs HTT_Is @ (prems RL [subsetD])))); +val NatF = result(); + +goal Hered.thy "Nat <= HTT"; +by (safe_tac set_cs); +be HTT_coinduct3 1; +by (fast_tac (set_cs addIs HTTgenIs + addSEs [HTTgen_mono RS ci3_RI] addEs [XH_to_E NatXH]) 1); +val NatF = result(); + +val [prem] = goal Hered.thy "A <= HTT ==> List(A) <= HTT"; +by (safe_tac set_cs); +be HTT_coinduct3 1; +by (fast_tac (set_cs addSIs HTTgenIs + addSEs [HTTgen_mono RS ci3_RI,prem RS subsetD RS (HTTgen_mono RS ci3_AI)] + addEs [XH_to_E ListXH]) 1); +val ListF = result(); + +val [prem] = goal Hered.thy "A <= HTT ==> Lists(A) <= HTT"; +by (safe_tac set_cs); +be HTT_coinduct3 1; +by (fast_tac (set_cs addSIs HTTgenIs + addSEs [HTTgen_mono RS ci3_RI,prem RS subsetD RS (HTTgen_mono RS ci3_AI)] + addEs [XH_to_E ListsXH]) 1); +val ListsF = result(); + +val [prem] = goal Hered.thy "A <= HTT ==> ILists(A) <= HTT"; +by (safe_tac set_cs); +be HTT_coinduct3 1; +by (fast_tac (set_cs addSIs HTTgenIs + addSEs [HTTgen_mono RS ci3_RI,prem RS subsetD RS (HTTgen_mono RS ci3_AI)] + addEs [XH_to_E IListsXH]) 1); +val IListsF = result(); + +(*** A possible use for this predicate is proving equality from pre-order ***) +(*** but it seems as easy (and more general) to do this directly by coinduction ***) +(* +val prems = goal Hered.thy "[| t : HTT; t [= u |] ==> u [= t"; +by (po_coinduct_tac "{p. EX a b.p= & b : HTT & b [= a}" 1); +by (fast_tac (ccl_cs addIs prems) 1); +by (safe_tac ccl_cs); +bd (poXH RS iffD1) 1; +by (safe_tac (set_cs addSEs [HTT_bot RS notE])); +by (REPEAT_SOME (rtac (POgenXH RS iffD2) ORELSE' etac rev_mp)); +by (ALLGOALS (SIMP_TAC (term_ss addrews HTT_rews))); +by (ALLGOALS (fast_tac ccl_cs)); +val HTT_po_op = result(); + +val prems = goal Hered.thy "[| t : HTT; t [= u |] ==> t = u"; +by (REPEAT (ares_tac (prems @ [conjI RS (eq_iff RS iffD2),HTT_po_op]) 1)); +val HTT_po_eq = result(); +*) diff -r 000000000000 -r a5a9c433f639 src/CCL/Hered.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Hered.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,30 @@ +(* Title: CCL/hered.thy + ID: $Id$ + Author: Martin Coen + Copyright 1993 University of Cambridge + +Hereditary Termination - cf. Martin Lo\"f + +Note that this is based on an untyped equality and so lam x.b(x) is only +hereditarily terminating if ALL x.b(x) is. Not so useful for functions! + +*) + +Hered = Type + + +consts + (*** Predicates ***) + HTTgen :: "i set => i set" + HTT :: "i set" + + +rules + + (*** Definitions of Hereditary Termination ***) + + HTTgen_def + "HTTgen(R) == {t. t=true | t=false | (EX a b.t= & a : R & b : R) | \ +\ (EX f. t=lam x.f(x) & (ALL x.f(x) : R))}" + HTT_def "HTT == gfp(HTTgen)" + +end diff -r 000000000000 -r a5a9c433f639 src/CCL/Lfp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Lfp.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,82 @@ +(* Title: CCL/lfp + ID: $Id$ + +Modified version of + Title: HOL/lfp.ML + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For lfp.thy. The Knaster-Tarski Theorem +*) + +open Lfp; + +(*** Proof of Knaster-Tarski Theorem ***) + +(* lfp(f) is the greatest lower bound of {u. f(u) <= u} *) + +val prems = goalw Lfp.thy [lfp_def] "[| f(A) <= A |] ==> lfp(f) <= A"; +by (rtac (CollectI RS Inter_lower) 1); +by (resolve_tac prems 1); +val lfp_lowerbound = result(); + +val prems = goalw Lfp.thy [lfp_def] + "[| !!u. f(u) <= u ==> A<=u |] ==> A <= lfp(f)"; +by (REPEAT (ares_tac ([Inter_greatest]@prems) 1)); +by (etac CollectD 1); +val lfp_greatest = result(); + +val [mono] = goal Lfp.thy "mono(f) ==> f(lfp(f)) <= lfp(f)"; +by (EVERY1 [rtac lfp_greatest, rtac subset_trans, + rtac (mono RS monoD), rtac lfp_lowerbound, atac, atac]); +val lfp_lemma2 = result(); + +val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) <= f(lfp(f))"; +by (EVERY1 [rtac lfp_lowerbound, rtac (mono RS monoD), + rtac lfp_lemma2, rtac mono]); +val lfp_lemma3 = result(); + +val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) = f(lfp(f))"; +by (REPEAT (resolve_tac [equalityI,lfp_lemma2,lfp_lemma3,mono] 1)); +val lfp_Tarski = result(); + + +(*** General induction rule for least fixed points ***) + +val [lfp,mono,indhyp] = goal Lfp.thy + "[| a: lfp(f); mono(f); \ +\ !!x. [| x: f(lfp(f) Int {x.P(x)}) |] ==> P(x) \ +\ |] ==> P(a)"; +by (res_inst_tac [("a","a")] (Int_lower2 RS subsetD RS CollectD) 1); +by (rtac (lfp RSN (2, lfp_lowerbound RS subsetD)) 1); +by (EVERY1 [rtac Int_greatest, rtac subset_trans, + rtac (Int_lower1 RS (mono RS monoD)), + rtac (mono RS lfp_lemma2), + rtac (CollectI RS subsetI), rtac indhyp, atac]); +val induct = result(); + +(** Definition forms of lfp_Tarski and induct, to control unfolding **) + +val [rew,mono] = goal Lfp.thy "[| h==lfp(f); mono(f) |] ==> h = f(h)"; +by (rewtac rew); +by (rtac (mono RS lfp_Tarski) 1); +val def_lfp_Tarski = result(); + +val rew::prems = goal Lfp.thy + "[| A == lfp(f); a:A; mono(f); \ +\ !!x. [| x: f(A Int {x.P(x)}) |] ==> P(x) \ +\ |] ==> P(a)"; +by (EVERY1 [rtac induct, (*backtracking to force correct induction*) + REPEAT1 o (ares_tac (map (rewrite_rule [rew]) prems))]); +val def_induct = result(); + +(*Monotonicity of lfp!*) +val prems = goal Lfp.thy + "[| mono(g); !!Z. f(Z)<=g(Z) |] ==> lfp(f) <= lfp(g)"; +by (rtac lfp_lowerbound 1); +by (rtac subset_trans 1); +by (resolve_tac prems 1); +by (rtac lfp_lemma2 1); +by (resolve_tac prems 1); +val lfp_mono = result(); + diff -r 000000000000 -r a5a9c433f639 src/CCL/Lfp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Lfp.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,14 @@ +(* Title: HOL/lfp.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +The Knaster-Tarski Theorem +*) + +Lfp = Set + +consts lfp :: "['a set=>'a set] => 'a set" +rules + (*least fixed point*) + lfp_def "lfp(f) == Inter({u. f(u) <= u})" +end diff -r 000000000000 -r a5a9c433f639 src/CCL/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Makefile Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,48 @@ +######################################################################### +# # +# Makefile for Isabelle (CCL) # +# # +######################################################################### + +#To make the system, cd to this directory and type +# make -f Makefile +#To make the system and test it on standard examples, type +# make -f Makefile test + +#Environment variable ISABELLECOMP specifies the compiler. +#Environment variable ISABELLEBIN specifies the destination directory. +#For Poly/ML, ISABELLEBIN must begin with a / + +#Makes FOL if this file is ABSENT -- but not +#if it is out of date, since this Makefile does not know its dependencies! + +BIN = $(ISABELLEBIN) +COMP = $(ISABELLECOMP) + +SET_FILES = ROOT.ML set.thy set.ML subset.ML equalities.ML mono.ML \ + gfp.thy gfp.ML lfp.thy lfp.ML + +CCL_FILES = ccl.thy ccl.ML terms.thy terms.ML types.thy types.ML \ + coinduction.ML hered.thy hered.ML trancl.thy trancl.ML\ + wf.thy wf.ML genrec.ML typecheck.ML eval.ML fix.thy fix.ML + +#Uses cp rather than make_database because Poly/ML allows only 3 levels +$(BIN)/CCL: $(BIN)/FOL $(SET_FILES) $(CCL_FILES) + case "$(COMP)" in \ + poly*) cp $(BIN)/FOL $(BIN)/CCL;\ + echo 'open PolyML; use"ROOT";' | $(COMP) $(BIN)/CCL ;;\ + sml*) echo 'use"ROOT.ML"; xML"$(BIN)/CCL" banner;' | $(BIN)/FOL;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +$(BIN)/FOL: + cd ../FOL; $(MAKE) + +test: ex/ROOT.ML $(BIN)/CCL + case "$(COMP)" in \ + poly*) echo 'use"ex/ROOT.ML"; quit();' | $(COMP) $(BIN)/CCL ;;\ + sml*) echo 'use"ex/ROOT.ML";' | $(BIN)/CCL;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +.PRECIOUS: $(BIN)/FOL $(BIN)/CCL diff -r 000000000000 -r a5a9c433f639 src/CCL/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,37 @@ +(* Title: CCL/ROOT + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Adds Classical Computational Logic to a database containing First-Order Logic. +*) + +val banner = "Classical Computational Logic (in FOL)"; + +(* Higher-Order Set Theory Extension to FOL *) +(* used as basis for CCL *) + +use_thy "set"; +use "subset.ML"; +use "equalities.ML"; +use "mono.ML"; +use_thy "lfp"; +use_thy "gfp"; + +(* CCL - a computational logic for an untyped functional language *) +(* with evaluation to weak head-normal form *) + +use_thy "ccl"; +use_thy "terms"; +use_thy "types"; +use "coinduction.ML"; +use_thy "hered"; + +use_thy "trancl"; +use_thy "wf"; +use "genrec.ML"; +use "typecheck.ML"; +use "eval.ML"; +use_thy "fix"; + +val CCL_build_completed = (); (*indicate successful build*) diff -r 000000000000 -r a5a9c433f639 src/CCL/Set.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Set.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,355 @@ +(* Title: set/set + ID: $Id$ + +For set.thy. + +Modified version of + Title: HOL/set + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For set.thy. Set theory for higher-order logic. A set is simply a predicate. +*) + +open Set; + +val [prem] = goal Set.thy "[| P(a) |] ==> a : {x.P(x)}"; +by (rtac (mem_Collect_iff RS iffD2) 1); +by (rtac prem 1); +val CollectI = result(); + +val prems = goal Set.thy "[| a : {x.P(x)} |] ==> P(a)"; +by (resolve_tac (prems RL [mem_Collect_iff RS iffD1]) 1); +val CollectD = result(); + +val [prem] = goal Set.thy "[| !!x. x:A <-> x:B |] ==> A = B"; +by (rtac (set_extension RS iffD2) 1); +by (rtac (prem RS allI) 1); +val set_ext = result(); + +val prems = goal Set.thy "[| !!x. P(x) <-> Q(x) |] ==> {x. P(x)} = {x. Q(x)}"; +by (REPEAT (ares_tac [set_ext,iffI,CollectI] 1 ORELSE + eresolve_tac ([CollectD] RL (prems RL [iffD1,iffD2])) 1)); +val Collect_cong = result(); + +val CollectE = make_elim CollectD; + +(*** Bounded quantifiers ***) + +val prems = goalw Set.thy [Ball_def] + "[| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)"; +by (REPEAT (ares_tac (prems @ [allI,impI]) 1)); +val ballI = result(); + +val [major,minor] = goalw Set.thy [Ball_def] + "[| ALL x:A. P(x); x:A |] ==> P(x)"; +by (rtac (minor RS (major RS spec RS mp)) 1); +val bspec = result(); + +val major::prems = goalw Set.thy [Ball_def] + "[| ALL x:A. P(x); P(x) ==> Q; ~ x:A ==> Q |] ==> Q"; +by (rtac (major RS spec RS impCE) 1); +by (REPEAT (eresolve_tac prems 1)); +val ballE = result(); + +(*Takes assumptions ALL x:A.P(x) and a:A; creates assumption P(a)*) +fun ball_tac i = etac ballE i THEN contr_tac (i+1); + +val prems = goalw Set.thy [Bex_def] + "[| P(x); x:A |] ==> EX x:A. P(x)"; +by (REPEAT (ares_tac (prems @ [exI,conjI]) 1)); +val bexI = result(); + +val bexCI = prove_goal Set.thy + "[| EX x:A. ~P(x) ==> P(a); a:A |] ==> EX x:A.P(x)" + (fn prems=> + [ (rtac classical 1), + (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ]); + +val major::prems = goalw Set.thy [Bex_def] + "[| EX x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q"; +by (rtac (major RS exE) 1); +by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)); +val bexE = result(); + +(*Trival rewrite rule; (! x:A.P)=P holds only if A is nonempty!*) +val prems = goal Set.thy + "(ALL x:A. True) <-> True"; +by (REPEAT (ares_tac [TrueI,ballI,iffI] 1)); +val ball_rew = result(); + +(** Congruence rules **) + +val prems = goal Set.thy + "[| A=A'; !!x. x:A' ==> P(x) <-> P'(x) |] ==> \ +\ (ALL x:A. P(x)) <-> (ALL x:A'. P'(x))"; +by (resolve_tac (prems RL [ssubst,iffD2]) 1); +by (REPEAT (ares_tac [ballI,iffI] 1 + ORELSE eresolve_tac ([make_elim bspec, mp] @ (prems RL [iffE])) 1)); +val ball_cong = result(); + +val prems = goal Set.thy + "[| A=A'; !!x. x:A' ==> P(x) <-> P'(x) |] ==> \ +\ (EX x:A. P(x)) <-> (EX x:A'. P'(x))"; +by (resolve_tac (prems RL [ssubst,iffD2]) 1); +by (REPEAT (etac bexE 1 + ORELSE ares_tac ([bexI,iffI] @ (prems RL [iffD1,iffD2])) 1)); +val bex_cong = result(); + +(*** Rules for subsets ***) + +val prems = goalw Set.thy [subset_def] "(!!x.x:A ==> x:B) ==> A <= B"; +by (REPEAT (ares_tac (prems @ [ballI]) 1)); +val subsetI = result(); + +(*Rule in Modus Ponens style*) +val major::prems = goalw Set.thy [subset_def] "[| A <= B; c:A |] ==> c:B"; +by (rtac (major RS bspec) 1); +by (resolve_tac prems 1); +val subsetD = result(); + +(*Classical elimination rule*) +val major::prems = goalw Set.thy [subset_def] + "[| A <= B; ~(c:A) ==> P; c:B ==> P |] ==> P"; +by (rtac (major RS ballE) 1); +by (REPEAT (eresolve_tac prems 1)); +val subsetCE = result(); + +(*Takes assumptions A<=B; c:A and creates the assumption c:B *) +fun set_mp_tac i = etac subsetCE i THEN mp_tac i; + +val subset_refl = prove_goal Set.thy "A <= A" + (fn _=> [ (REPEAT (ares_tac [subsetI] 1)) ]); + +goal Set.thy "!!A B C. [| A<=B; B<=C |] ==> A<=C"; +br subsetI 1; +by (REPEAT (eresolve_tac [asm_rl, subsetD] 1)); +val subset_trans = result(); + + +(*** Rules for equality ***) + +(*Anti-symmetry of the subset relation*) +val prems = goal Set.thy "[| A <= B; B <= A |] ==> A = B"; +by (rtac (iffI RS set_ext) 1); +by (REPEAT (ares_tac (prems RL [subsetD]) 1)); +val subset_antisym = result(); +val equalityI = subset_antisym; + +(* Equality rules from ZF set theory -- are they appropriate here? *) +val prems = goal Set.thy "A = B ==> A<=B"; +by (resolve_tac (prems RL [subst]) 1); +by (rtac subset_refl 1); +val equalityD1 = result(); + +val prems = goal Set.thy "A = B ==> B<=A"; +by (resolve_tac (prems RL [subst]) 1); +by (rtac subset_refl 1); +val equalityD2 = result(); + +val prems = goal Set.thy + "[| A = B; [| A<=B; B<=A |] ==> P |] ==> P"; +by (resolve_tac prems 1); +by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1)); +val equalityE = result(); + +val major::prems = goal Set.thy + "[| A = B; [| c:A; c:B |] ==> P; [| ~ c:A; ~ c:B |] ==> P |] ==> P"; +by (rtac (major RS equalityE) 1); +by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1)); +val equalityCE = result(); + +(*Lemma for creating induction formulae -- for "pattern matching" on p + To make the induction hypotheses usable, apply "spec" or "bspec" to + put universal quantifiers over the free variables in p. *) +val prems = goal Set.thy + "[| p:A; !!z. z:A ==> p=z --> R |] ==> R"; +by (rtac mp 1); +by (REPEAT (resolve_tac (refl::prems) 1)); +val setup_induction = result(); + +goal Set.thy "{x.x:A} = A"; +by (REPEAT (ares_tac [equalityI,subsetI,CollectI] 1 ORELSE eresolve_tac [CollectD] 1)); +val trivial_set = result(); + +(*** Rules for binary union -- Un ***) + +val prems = goalw Set.thy [Un_def] "c:A ==> c : A Un B"; +by (REPEAT (resolve_tac (prems @ [CollectI,disjI1]) 1)); +val UnI1 = result(); + +val prems = goalw Set.thy [Un_def] "c:B ==> c : A Un B"; +by (REPEAT (resolve_tac (prems @ [CollectI,disjI2]) 1)); +val UnI2 = result(); + +(*Classical introduction rule: no commitment to A vs B*) +val UnCI = prove_goal Set.thy "(~c:B ==> c:A) ==> c : A Un B" + (fn prems=> + [ (rtac classical 1), + (REPEAT (ares_tac (prems@[UnI1,notI]) 1)), + (REPEAT (ares_tac (prems@[UnI2,notE]) 1)) ]); + +val major::prems = goalw Set.thy [Un_def] + "[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P"; +by (rtac (major RS CollectD RS disjE) 1); +by (REPEAT (eresolve_tac prems 1)); +val UnE = result(); + + +(*** Rules for small intersection -- Int ***) + +val prems = goalw Set.thy [Int_def] + "[| c:A; c:B |] ==> c : A Int B"; +by (REPEAT (resolve_tac (prems @ [CollectI,conjI]) 1)); +val IntI = result(); + +val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:A"; +by (rtac (major RS CollectD RS conjunct1) 1); +val IntD1 = result(); + +val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:B"; +by (rtac (major RS CollectD RS conjunct2) 1); +val IntD2 = result(); + +val [major,minor] = goal Set.thy + "[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P"; +by (rtac minor 1); +by (rtac (major RS IntD1) 1); +by (rtac (major RS IntD2) 1); +val IntE = result(); + + +(*** Rules for set complement -- Compl ***) + +val prems = goalw Set.thy [Compl_def] + "[| c:A ==> False |] ==> c : Compl(A)"; +by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1)); +val ComplI = result(); + +(*This form, with negated conclusion, works well with the Classical prover. + Negated assumptions behave like formulae on the right side of the notional + turnstile...*) +val major::prems = goalw Set.thy [Compl_def] + "[| c : Compl(A) |] ==> ~c:A"; +by (rtac (major RS CollectD) 1); +val ComplD = result(); + +val ComplE = make_elim ComplD; + + +(*** Empty sets ***) + +goalw Set.thy [empty_def] "{x.False} = {}"; +br refl 1; +val empty_eq = result(); + +val [prem] = goalw Set.thy [empty_def] "a : {} ==> P"; +by (rtac (prem RS CollectD RS FalseE) 1); +val emptyD = result(); + +val emptyE = make_elim emptyD; + +val [prem] = goal Set.thy "~ A={} ==> (EX x.x:A)"; +br (prem RS swap) 1; +br equalityI 1; +by (ALLGOALS (fast_tac (FOL_cs addSIs [subsetI] addSEs [emptyD]))); +val not_emptyD = result(); + +(*** Singleton sets ***) + +goalw Set.thy [singleton_def] "a : {a}"; +by (rtac CollectI 1); +by (rtac refl 1); +val singletonI = result(); + +val [major] = goalw Set.thy [singleton_def] "b : {a} ==> b=a"; +by (rtac (major RS CollectD) 1); +val singletonD = result(); + +val singletonE = make_elim singletonD; + +(*** Unions of families ***) + +(*The order of the premises presupposes that A is rigid; b may be flexible*) +val prems = goalw Set.thy [UNION_def] + "[| a:A; b: B(a) |] ==> b: (UN x:A. B(x))"; +by (REPEAT (resolve_tac (prems @ [bexI,CollectI]) 1)); +val UN_I = result(); + +val major::prems = goalw Set.thy [UNION_def] + "[| b : (UN x:A. B(x)); !!x.[| x:A; b: B(x) |] ==> R |] ==> R"; +by (rtac (major RS CollectD RS bexE) 1); +by (REPEAT (ares_tac prems 1)); +val UN_E = result(); + +val prems = goal Set.thy + "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ +\ (UN x:A. C(x)) = (UN x:B. D(x))"; +by (REPEAT (etac UN_E 1 + ORELSE ares_tac ([UN_I,equalityI,subsetI] @ + (prems RL [equalityD1,equalityD2] RL [subsetD])) 1)); +val UN_cong = result(); + +(*** Intersections of families -- INTER x:A. B(x) is Inter(B)``A ) *) + +val prems = goalw Set.thy [INTER_def] + "(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))"; +by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1)); +val INT_I = result(); + +val major::prems = goalw Set.thy [INTER_def] + "[| b : (INT x:A. B(x)); a:A |] ==> b: B(a)"; +by (rtac (major RS CollectD RS bspec) 1); +by (resolve_tac prems 1); +val INT_D = result(); + +(*"Classical" elimination rule -- does not require proving X:C *) +val major::prems = goalw Set.thy [INTER_def] + "[| b : (INT x:A. B(x)); b: B(a) ==> R; ~ a:A ==> R |] ==> R"; +by (rtac (major RS CollectD RS ballE) 1); +by (REPEAT (eresolve_tac prems 1)); +val INT_E = result(); + +val prems = goal Set.thy + "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ +\ (INT x:A. C(x)) = (INT x:B. D(x))"; +by (REPEAT_FIRST (resolve_tac [INT_I,equalityI,subsetI])); +by (REPEAT (dtac INT_D 1 + ORELSE ares_tac (prems RL [equalityD1,equalityD2] RL [subsetD]) 1)); +val INT_cong = result(); + +(*** Rules for Unions ***) + +(*The order of the premises presupposes that C is rigid; A may be flexible*) +val prems = goalw Set.thy [Union_def] + "[| X:C; A:X |] ==> A : Union(C)"; +by (REPEAT (resolve_tac (prems @ [UN_I]) 1)); +val UnionI = result(); + +val major::prems = goalw Set.thy [Union_def] + "[| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R"; +by (rtac (major RS UN_E) 1); +by (REPEAT (ares_tac prems 1)); +val UnionE = result(); + +(*** Rules for Inter ***) + +val prems = goalw Set.thy [Inter_def] + "[| !!X. X:C ==> A:X |] ==> A : Inter(C)"; +by (REPEAT (ares_tac ([INT_I] @ prems) 1)); +val InterI = result(); + +(*A "destruct" rule -- every X in C contains A as an element, but + A:X can hold when X:C does not! This rule is analogous to "spec". *) +val major::prems = goalw Set.thy [Inter_def] + "[| A : Inter(C); X:C |] ==> A:X"; +by (rtac (major RS INT_D) 1); +by (resolve_tac prems 1); +val InterD = result(); + +(*"Classical" elimination rule -- does not require proving X:C *) +val major::prems = goalw Set.thy [Inter_def] + "[| A : Inter(C); A:X ==> R; ~ X:C ==> R |] ==> R"; +by (rtac (major RS INT_E) 1); +by (REPEAT (eresolve_tac prems 1)); +val InterE = result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/Set.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Set.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,71 @@ +(* Title: CCL/set.thy + ID: $Id$ + +Modified version of HOL/set.thy that extends FOL + +*) + +Set = FOL + + +types + set 1 + +arities + set :: (term) term + +consts + Collect :: "['a => o] => 'a set" (*comprehension*) + Compl :: "('a set) => 'a set" (*complement*) + Int :: "['a set, 'a set] => 'a set" (infixl 70) + Un :: "['a set, 'a set] => 'a set" (infixl 65) + Union, Inter :: "(('a set)set) => 'a set" (*...of a set*) + UNION, INTER :: "['a set, 'a => 'b set] => 'b set" (*general*) + Ball, Bex :: "['a set, 'a => o] => o" (*bounded quants*) + mono :: "['a set => 'b set] => o" (*monotonicity*) + ":" :: "['a, 'a set] => o" (infixl 50) (*membership*) + "<=" :: "['a set, 'a set] => o" (infixl 50) + singleton :: "'a => 'a set" ("{_}") + empty :: "'a set" ("{}") + "oo" :: "['b => 'c, 'a => 'b, 'a] => 'c" (infixr 50) (*composition*) + + "@Coll" :: "[idt, o] => 'a set" ("(1{_./ _})") (*collection*) + + (* Big Intersection / Union *) + + "@INTER" :: "[idt, 'a set, 'b set] => 'b set" ("(INT _:_./ _)" [0, 0, 0] 10) + "@UNION" :: "[idt, 'a set, 'b set] => 'b set" ("(UN _:_./ _)" [0, 0, 0] 10) + + (* Bounded Quantifiers *) + + "@Ball" :: "[idt, 'a set, o] => o" ("(ALL _:_./ _)" [0, 0, 0] 10) + "@Bex" :: "[idt, 'a set, o] => o" ("(EX _:_./ _)" [0, 0, 0] 10) + + +translations + "{x. P}" == "Collect(%x. P)" + "INT x:A. B" == "INTER(A, %x. B)" + "UN x:A. B" == "UNION(A, %x. B)" + "ALL x:A. P" == "Ball(A, %x. P)" + "EX x:A. P" == "Bex(A, %x. P)" + + +rules + mem_Collect_iff "(a : {x.P(x)}) <-> P(a)" + set_extension "A=B <-> (ALL x.x:A <-> x:B)" + + Ball_def "Ball(A, P) == ALL x. x:A --> P(x)" + Bex_def "Bex(A, P) == EX x. x:A & P(x)" + mono_def "mono(f) == (ALL A B. A <= B --> f(A) <= f(B))" + subset_def "A <= B == ALL x:A. x:B" + singleton_def "{a} == {x.x=a}" + empty_def "{} == {x.False}" + Un_def "A Un B == {x.x:A | x:B}" + Int_def "A Int B == {x.x:A & x:B}" + Compl_def "Compl(A) == {x. ~x:A}" + INTER_def "INTER(A, B) == {y. ALL x:A. y: B(x)}" + UNION_def "UNION(A, B) == {y. EX x:A. y: B(x)}" + Inter_def "Inter(S) == (INT x:S. x)" + Union_def "Union(S) == (UN x:S. x)" + +end + diff -r 000000000000 -r a5a9c433f639 src/CCL/Term.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Term.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,146 @@ +(* Title: CCL/terms + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For terms.thy. +*) + +open Term; + +val simp_can_defs = [one_def,inl_def,inr_def]; +val simp_ncan_defs = [if_def,when_def,split_def,fst_def,snd_def,thd_def]; +val simp_defs = simp_can_defs @ simp_ncan_defs; + +val ind_can_defs = [zero_def,succ_def,nil_def,cons_def]; +val ind_ncan_defs = [ncase_def,nrec_def,lcase_def,lrec_def]; +val ind_defs = ind_can_defs @ ind_ncan_defs; + +val data_defs = simp_defs @ ind_defs @ [napply_def]; +val genrec_defs = [letrec_def,letrec2_def,letrec3_def]; + +val term_congs = ccl_mk_congs Term.thy + ["inl","inr","succ","op .","split","if","when","ncase","nrec","lcase","lrec", + "fst","snd","thd","let","letrec","letrec2","letrec3","napply"]; + +(*** Beta Rules, including strictness ***) + +goalw Term.thy [let_def] "~ t=bot--> let x be t in f(x) = f(t)"; +by (res_inst_tac [("t","t")] term_case 1); +by (ALLGOALS(SIMP_TAC(CCL_ss addrews [caseBtrue,caseBfalse,caseBpair,caseBlam]))); +val letB = result() RS mp; + +goalw Term.thy [let_def] "let x be bot in f(x) = bot"; +br caseBbot 1; +val letBabot = result(); + +goalw Term.thy [let_def] "let x be t in bot = bot"; +brs ([caseBbot] RL [term_case]) 1; +by (ALLGOALS(SIMP_TAC(CCL_ss addrews [caseBtrue,caseBfalse,caseBpair,caseBlam]))); +val letBbbot = result(); + +goalw Term.thy [apply_def] "(lam x.b(x)) ` a = b(a)"; +by (ALLGOALS(SIMP_TAC(CCL_ss addrews [caseBtrue,caseBfalse,caseBpair,caseBlam]))); +val applyB = result(); + +goalw Term.thy [apply_def] "bot ` a = bot"; +br caseBbot 1; +val applyBbot = result(); + +goalw Term.thy [fix_def] "fix(f) = f(fix(f))"; +by (resolve_tac [applyB RS ssubst] 1 THEN resolve_tac [refl] 1); +val fixB = result(); + +goalw Term.thy [letrec_def] + "letrec g x be h(x,g) in g(a) = h(a,%y.letrec g x be h(x,g) in g(y))"; +by (resolve_tac [fixB RS ssubst] 1 THEN + resolve_tac [applyB RS ssubst] 1 THEN resolve_tac [refl] 1); +val letrecB = result(); + +val rawBs = caseBs @ [applyB,applyBbot,letrecB]; + +fun raw_mk_beta_rl defs s = prove_goalw Term.thy defs s + (fn _ => [SIMP_TAC (CCL_ss addrews rawBs addcongs term_congs) 1]); +fun mk_beta_rl s = raw_mk_beta_rl data_defs s; + +val ifBtrue = mk_beta_rl "if true then t else u = t"; +val ifBfalse = mk_beta_rl "if false then t else u = u"; +val ifBbot = mk_beta_rl "if bot then t else u = bot"; + +val whenBinl = mk_beta_rl "when(inl(a),t,u) = t(a)"; +val whenBinr = mk_beta_rl "when(inr(a),t,u) = u(a)"; +val whenBbot = mk_beta_rl "when(bot,t,u) = bot"; + +val splitB = mk_beta_rl "split(,h) = h(a,b)"; +val splitBbot = mk_beta_rl "split(bot,h) = bot"; +val fstB = mk_beta_rl "fst() = a"; +val fstBbot = mk_beta_rl "fst(bot) = bot"; +val sndB = mk_beta_rl "snd() = b"; +val sndBbot = mk_beta_rl "snd(bot) = bot"; +val thdB = mk_beta_rl "thd(>) = c"; +val thdBbot = mk_beta_rl "thd(bot) = bot"; + +val ncaseBzero = mk_beta_rl "ncase(zero,t,u) = t"; +val ncaseBsucc = mk_beta_rl "ncase(succ(n),t,u) = u(n)"; +val ncaseBbot = mk_beta_rl "ncase(bot,t,u) = bot"; +val nrecBzero = mk_beta_rl "nrec(zero,t,u) = t"; +val nrecBsucc = mk_beta_rl "nrec(succ(n),t,u) = u(n,nrec(n,t,u))"; +val nrecBbot = mk_beta_rl "nrec(bot,t,u) = bot"; + +val lcaseBnil = mk_beta_rl "lcase([],t,u) = t"; +val lcaseBcons = mk_beta_rl "lcase(x.xs,t,u) = u(x,xs)"; +val lcaseBbot = mk_beta_rl "lcase(bot,t,u) = bot"; +val lrecBnil = mk_beta_rl "lrec([],t,u) = t"; +val lrecBcons = mk_beta_rl "lrec(x.xs,t,u) = u(x,xs,lrec(xs,t,u))"; +val lrecBbot = mk_beta_rl "lrec(bot,t,u) = bot"; + +val letrec2B = raw_mk_beta_rl (data_defs @ [letrec2_def]) + "letrec g x y be h(x,y,g) in g(p,q) = \ +\ h(p,q,%u v.letrec g x y be h(x,y,g) in g(u,v))"; +val letrec3B = raw_mk_beta_rl (data_defs @ [letrec3_def]) + "letrec g x y z be h(x,y,z,g) in g(p,q,r) = \ +\ h(p,q,r,%u v w.letrec g x y z be h(x,y,z,g) in g(u,v,w))"; + +val napplyBzero = mk_beta_rl "f^zero`a = a"; +val napplyBsucc = mk_beta_rl "f^succ(n)`a = f(f^n`a)"; + +val termBs = [letB,applyB,applyBbot,splitB,splitBbot, + fstB,fstBbot,sndB,sndBbot,thdB,thdBbot, + ifBtrue,ifBfalse,ifBbot,whenBinl,whenBinr,whenBbot, + ncaseBzero,ncaseBsucc,ncaseBbot,nrecBzero,nrecBsucc,nrecBbot, + lcaseBnil,lcaseBcons,lcaseBbot,lrecBnil,lrecBcons,lrecBbot, + napplyBzero,napplyBsucc]; + +(*** Constructors are injective ***) + +val term_injs = map (mk_inj_rl Term.thy + [applyB,splitB,whenBinl,whenBinr,ncaseBsucc,lcaseBcons] + (ccl_congs @ term_congs)) + ["(inl(a) = inl(a')) <-> (a=a')", + "(inr(a) = inr(a')) <-> (a=a')", + "(succ(a) = succ(a')) <-> (a=a')", + "(a.b = a'.b') <-> (a=a' & b=b')"]; + +(*** Constructors are distinct ***) + +val term_dstncts = mkall_dstnct_thms Term.thy data_defs (ccl_injs @ term_injs) + [["bot","inl","inr"],["bot","zero","succ"],["bot","nil","op ."]]; + +(*** Rules for pre-order [= ***) + +local + fun mk_thm s = prove_goalw Term.thy data_defs s (fn _ => + [SIMP_TAC (ccl_ss addrews (ccl_porews)) 1]); +in + val term_porews = map mk_thm ["inl(a) [= inl(a') <-> a [= a'", + "inr(b) [= inr(b') <-> b [= b'", + "succ(n) [= succ(n') <-> n [= n'", + "x.xs [= x'.xs' <-> x [= x' & xs [= xs'"]; +end; + +(*** Rewriting and Proving ***) + +val term_rews = termBs @ term_injs @ term_dstncts @ ccl_porews @ term_porews; +val term_ss = ccl_ss addrews term_rews addcongs term_congs; + +val term_cs = ccl_cs addSEs (term_dstncts RL [notE]) addSDs (XH_to_Ds term_injs); diff -r 000000000000 -r a5a9c433f639 src/CCL/Term.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Term.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,131 @@ +(* Title: CCL/terms.thy + ID: $Id$ + Author: Martin Coen + Copyright 1993 University of Cambridge + +Definitions of usual program constructs in CCL. + +*) + +Term = CCL + + +consts + + one :: "i" + + if :: "[i,i,i]=>i" ("(3if _/ then _/ else _)" [] 60) + + inl,inr :: "i=>i" + when :: "[i,i=>i,i=>i]=>i" + + split :: "[i,[i,i]=>i]=>i" + fst,snd, + thd :: "i=>i" + + zero :: "i" + succ :: "i=>i" + ncase :: "[i,i,i=>i]=>i" + nrec :: "[i,i,[i,i]=>i]=>i" + + nil :: "i" ("([])") + "." :: "[i,i]=>i" (infixr 80) + lcase :: "[i,i,[i,i]=>i]=>i" + lrec :: "[i,i,[i,i,i]=>i]=>i" + + let :: "[i,i=>i]=>i" + letrec :: "[[i,i=>i]=>i,(i=>i)=>i]=>i" + letrec2 :: "[[i,i,i=>i=>i]=>i,(i=>i=>i)=>i]=>i" + letrec3 :: "[[i,i,i,i=>i=>i=>i]=>i,(i=>i=>i=>i)=>i]=>i" + + "@let" :: "[id,i,i]=>i" ("(3let _ be _/ in _)" [] 60) + "@letrec" :: "[id,id,i,i]=>i" ("(3letrec _ _ be _/ in _)" [] 60) + "@letrec2" :: "[id,id,id,i,i]=>i" ("(3letrec _ _ _ be _/ in _)" [] 60) + "@letrec3" :: "[id,id,id,id,i,i]=>i" ("(3letrec _ _ _ _ be _/ in _)" [] 60) + + napply :: "[i=>i,i,i]=>i" ("(_ ^ _ ` _)") + +rules + + one_def "one == true" + if_def "if b then t else u == case(b,t,u,% x y.bot,%v.bot)" + inl_def "inl(a) == " + inr_def "inr(b) == " + when_def "when(t,f,g) == split(t,%b x.if b then f(x) else g(x))" + split_def "split(t,f) == case(t,bot,bot,f,%u.bot)" + fst_def "fst(t) == split(t,%x y.x)" + snd_def "snd(t) == split(t,%x y.y)" + thd_def "thd(t) == split(t,%x p.split(p,%y z.z))" + zero_def "zero == inl(one)" + succ_def "succ(n) == inr(n)" + ncase_def "ncase(n,b,c) == when(n,%x.b,%y.c(y))" + nrec_def " nrec(n,b,c) == letrec g x be ncase(x,b,%y.c(y,g(y))) in g(n)" + nil_def "[] == inl(one)" + cons_def "h.t == inr()" + lcase_def "lcase(l,b,c) == when(l,%x.b,%y.split(y,c))" + lrec_def "lrec(l,b,c) == letrec g x be lcase(x,b,%h t.c(h,t,g(t))) in g(l)" + + let_def "let x be t in f(x) == case(t,f(true),f(false),%x y.f(),%u.f(lam x.u(x)))" + letrec_def + "letrec g x be h(x,g) in b(g) == b(%x.fix(%f.lam x.h(x,%y.f`y))`x)" + + letrec2_def "letrec g x y be h(x,y,g) in f(g)== \ +\ letrec g' p be split(p,%x y.h(x,y,%u v.g'())) \ +\ in f(%x y.g'())" + + letrec3_def "letrec g x y z be h(x,y,z,g) in f(g) == \ +\ letrec g' p be split(p,%x xs.split(xs,%y z.h(x,y,z,%u v w.g'(>)))) \ +\ in f(%x y z.g'(>))" + + napply_def "f ^n` a == nrec(n,a,%x g.f(g))" + +end + +ML + +(** Quantifier translations: variable binding **) + +fun let_tr [Free(id,T),a,b] = Const("let",dummyT) $ a $ absfree(id,T,b); +fun let_tr' [a,Abs(id,T,b)] = + let val (id',b') = variant_abs(id,T,b) + in Const("@let",dummyT) $ Free(id',T) $ a $ b' end; + +fun letrec_tr [Free(f,S),Free(x,T),a,b] = + Const("letrec",dummyT) $ absfree(x,T,absfree(f,S,a)) $ absfree(f,S,b); +fun letrec2_tr [Free(f,S),Free(x,T),Free(y,U),a,b] = + Const("letrec2",dummyT) $ absfree(x,T,absfree(y,U,absfree(f,S,a))) $ absfree(f,S,b); +fun letrec3_tr [Free(f,S),Free(x,T),Free(y,U),Free(z,V),a,b] = + Const("letrec3",dummyT) $ absfree(x,T,absfree(y,U,absfree(z,U,absfree(f,S,a)))) $ absfree(f,S,b); + +fun letrec_tr' [Abs(x,T,Abs(f,S,a)),Abs(ff,SS,b)] = + let val (f',b') = variant_abs(ff,SS,b) + val (_,a'') = variant_abs(f,S,a) + val (x',a') = variant_abs(x,T,a'') + in Const("@letrec",dummyT) $ Free(f',SS) $ Free(x',T) $ a' $ b' end; +fun letrec2_tr' [Abs(x,T,Abs(y,U,Abs(f,S,a))),Abs(ff,SS,b)] = + let val (f',b') = variant_abs(ff,SS,b) + val ( _,a1) = variant_abs(f,S,a) + val (y',a2) = variant_abs(y,U,a1) + val (x',a') = variant_abs(x,T,a2) + in Const("@letrec2",dummyT) $ Free(f',SS) $ Free(x',T) $ Free(y',U) $ a' $ b' + end; +fun letrec3_tr' [Abs(x,T,Abs(y,U,Abs(z,V,Abs(f,S,a)))),Abs(ff,SS,b)] = + let val (f',b') = variant_abs(ff,SS,b) + val ( _,a1) = variant_abs(f,S,a) + val (z',a2) = variant_abs(z,V,a1) + val (y',a3) = variant_abs(y,U,a2) + val (x',a') = variant_abs(x,T,a3) + in Const("@letrec3",dummyT) $ Free(f',SS) $ Free(x',T) $ Free(y',U) $ Free(z',V) $ a' $ b' + end; + +val parse_translation= + [("@let", let_tr), + ("@letrec", letrec_tr), + ("@letrec2", letrec2_tr), + ("@letrec3", letrec3_tr) + ]; +val print_translation= + [("let", let_tr'), + ("letrec", letrec_tr'), + ("letrec2", letrec2_tr'), + ("letrec3", letrec3_tr') + ]; diff -r 000000000000 -r a5a9c433f639 src/CCL/Trancl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Trancl.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,215 @@ +(* Title: CCL/trancl + ID: $Id$ + +For trancl.thy. + +Modified version of + Title: HOL/trancl.ML + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +*) + +open Trancl; + +(** Natural deduction for trans(r) **) + +val prems = goalw Trancl.thy [trans_def] + "(!! x y z. [| :r; :r |] ==> :r) ==> trans(r)"; +by (REPEAT (ares_tac (prems@[allI,impI]) 1)); +val transI = result(); + +val major::prems = goalw Trancl.thy [trans_def] + "[| trans(r); :r; :r |] ==> :r"; +by (cut_facts_tac [major] 1); +by (fast_tac (FOL_cs addIs prems) 1); +val transD = result(); + +(** Identity relation **) + +goalw Trancl.thy [id_def] " : id"; +by (rtac CollectI 1); +by (rtac exI 1); +by (rtac refl 1); +val idI = result(); + +val major::prems = goalw Trancl.thy [id_def] + "[| p: id; !!x.[| p = |] ==> P \ +\ |] ==> P"; +by (rtac (major RS CollectE) 1); +by (etac exE 1); +by (eresolve_tac prems 1); +val idE = result(); + +(** Composition of two relations **) + +val prems = goalw Trancl.thy [comp_def] + "[| :s; :r |] ==> : r O s"; +by (fast_tac (set_cs addIs prems) 1); +val compI = result(); + +(*proof requires higher-level assumptions or a delaying of hyp_subst_tac*) +val prems = goalw Trancl.thy [comp_def] + "[| xz : r O s; \ +\ !!x y z. [| xz = ; :s; :r |] ==> P \ +\ |] ==> P"; +by (cut_facts_tac prems 1); +by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1)); +val compE = result(); + +val prems = goal Trancl.thy + "[| : r O s; \ +\ !!y. [| :s; :r |] ==> P \ +\ |] ==> P"; +by (rtac compE 1); +by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [pair_inject,ssubst] 1)); +val compEpair = result(); + +val comp_cs = set_cs addIs [compI,idI] + addEs [compE,idE] + addSEs [pair_inject]; + +val prems = goal Trancl.thy + "[| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; +by (cut_facts_tac prems 1); +by (fast_tac comp_cs 1); +val comp_mono = result(); + +(** The relation rtrancl **) + +goal Trancl.thy "mono(%s. id Un (r O s))"; +by (rtac monoI 1); +by (REPEAT (ares_tac [monoI, subset_refl, comp_mono, Un_mono] 1)); +val rtrancl_fun_mono = result(); + +val rtrancl_unfold = rtrancl_fun_mono RS (rtrancl_def RS def_lfp_Tarski); + +(*Reflexivity of rtrancl*) +goal Trancl.thy " : r^*"; +br (rtrancl_unfold RS ssubst) 1; +by (fast_tac comp_cs 1); +val rtrancl_refl = result(); + +(*Closure under composition with r*) +val prems = goal Trancl.thy + "[| : r^*; : r |] ==> : r^*"; +br (rtrancl_unfold RS ssubst) 1; +by (fast_tac (comp_cs addIs prems) 1); +val rtrancl_into_rtrancl = result(); + +(*rtrancl of r contains r*) +val [prem] = goal Trancl.thy "[| : r |] ==> : r^*"; +by (rtac (rtrancl_refl RS rtrancl_into_rtrancl) 1); +by (rtac prem 1); +val r_into_rtrancl = result(); + + +(** standard induction rule **) + +val major::prems = goal Trancl.thy + "[| : r^*; \ +\ !!x. P(); \ +\ !!x y z.[| P(); : r^*; : r |] ==> P() |] \ +\ ==> P()"; +by (rtac (major RS (rtrancl_def RS def_induct)) 1); +by (rtac rtrancl_fun_mono 1); +by (fast_tac (comp_cs addIs prems) 1); +val rtrancl_full_induct = result(); + +(*nice induction rule*) +val major::prems = goal Trancl.thy + "[| : r^*; \ +\ P(a); \ +\ !!y z.[| : r^*; : r; P(y) |] ==> P(z) |] \ +\ ==> P(b)"; +(*by induction on this formula*) +by (subgoal_tac "ALL y. = --> P(y)" 1); +(*now solve first subgoal: this formula is sufficient*) +by (fast_tac FOL_cs 1); +(*now do the induction*) +by (resolve_tac [major RS rtrancl_full_induct] 1); +by (fast_tac (comp_cs addIs prems) 1); +by (fast_tac (comp_cs addIs prems) 1); +val rtrancl_induct = result(); + +(*transitivity of transitive closure!! -- by induction.*) +goal Trancl.thy "trans(r^*)"; +by (rtac transI 1); +by (res_inst_tac [("b","z")] rtrancl_induct 1); +by (DEPTH_SOLVE (eresolve_tac [asm_rl, rtrancl_into_rtrancl] 1)); +val trans_rtrancl = result(); + +(*elimination of rtrancl -- by induction on a special formula*) +val major::prems = goal Trancl.thy + "[| : r^*; (a = b) ==> P; \ +\ !!y.[| : r^*; : r |] ==> P |] \ +\ ==> P"; +by (subgoal_tac "a = b | (EX y. : r^* & : r)" 1); +by (rtac (major RS rtrancl_induct) 2); +by (fast_tac (set_cs addIs prems) 2); +by (fast_tac (set_cs addIs prems) 2); +by (REPEAT (eresolve_tac ([asm_rl,exE,disjE,conjE]@prems) 1)); +val rtranclE = result(); + + +(**** The relation trancl ****) + +(** Conversions between trancl and rtrancl **) + +val [major] = goalw Trancl.thy [trancl_def] + "[| : r^+ |] ==> : r^*"; +by (resolve_tac [major RS compEpair] 1); +by (REPEAT (ares_tac [rtrancl_into_rtrancl] 1)); +val trancl_into_rtrancl = result(); + +(*r^+ contains r*) +val [prem] = goalw Trancl.thy [trancl_def] + "[| : r |] ==> : r^+"; +by (REPEAT (ares_tac [prem,compI,rtrancl_refl] 1)); +val r_into_trancl = result(); + +(*intro rule by definition: from rtrancl and r*) +val prems = goalw Trancl.thy [trancl_def] + "[| : r^*; : r |] ==> : r^+"; +by (REPEAT (resolve_tac ([compI]@prems) 1)); +val rtrancl_into_trancl1 = result(); + +(*intro rule from r and rtrancl*) +val prems = goal Trancl.thy + "[| : r; : r^* |] ==> : r^+"; +by (resolve_tac (prems RL [rtranclE]) 1); +by (etac subst 1); +by (resolve_tac (prems RL [r_into_trancl]) 1); +by (rtac (trans_rtrancl RS transD RS rtrancl_into_trancl1) 1); +by (REPEAT (ares_tac (prems@[r_into_rtrancl]) 1)); +val rtrancl_into_trancl2 = result(); + +(*elimination of r^+ -- NOT an induction rule*) +val major::prems = goal Trancl.thy + "[| : r^+; \ +\ : r ==> P; \ +\ !!y.[| : r^+; : r |] ==> P \ +\ |] ==> P"; +by (subgoal_tac " : r | (EX y. : r^+ & : r)" 1); +by (REPEAT (eresolve_tac ([asm_rl,disjE,exE,conjE]@prems) 1)); +by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); +by (etac rtranclE 1); +by (fast_tac comp_cs 1); +by (fast_tac (comp_cs addSIs [rtrancl_into_trancl1]) 1); +val tranclE = result(); + +(*Transitivity of r^+. + Proved by unfolding since it uses transitivity of rtrancl. *) +goalw Trancl.thy [trancl_def] "trans(r^+)"; +by (rtac transI 1); +by (REPEAT (etac compEpair 1)); +by (rtac (rtrancl_into_rtrancl RS (trans_rtrancl RS transD RS compI)) 1); +by (REPEAT (assume_tac 1)); +val trans_trancl = result(); + +val prems = goal Trancl.thy + "[| : r; : r^+ |] ==> : r^+"; +by (rtac (r_into_trancl RS (trans_trancl RS transD)) 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +val trancl_into_trancl2 = result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/Trancl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Trancl.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,28 @@ +(* Title: CCL/trancl.thy + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Transitive closure of a relation +*) + +Trancl = CCL + + +consts + trans :: "i set => o" (*transitivity predicate*) + id :: "i set" + rtrancl :: "i set => i set" ("(_^*)" [100] 100) + trancl :: "i set => i set" ("(_^+)" [100] 100) + O :: "[i set,i set] => i set" (infixr 60) + +rules + +trans_def "trans(r) == (ALL x y z. :r --> :r --> :r)" +comp_def (*composition of relations*) + "r O s == {xz. EX x y z. xz = & :s & :r}" +id_def (*the identity relation*) + "id == {p. EX x. p = }" +rtrancl_def "r^* == lfp(%s. id Un (r O s))" +trancl_def "r^+ == r O rtrancl(r)" + +end diff -r 000000000000 -r a5a9c433f639 src/CCL/Type.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Type.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,308 @@ +(* Title: CCL/types + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For types.thy. +*) + +open Type; + +val simp_type_defs = [Subtype_def,Unit_def,Bool_def,Plus_def,Sigma_def,Pi_def, + Lift_def,Tall_def,Tex_def]; +val ind_type_defs = [Nat_def,List_def]; + +val simp_data_defs = [one_def,inl_def,inr_def]; +val ind_data_defs = [zero_def,succ_def,nil_def,cons_def]; + +goal Set.thy "A <= B <-> (ALL x.x:A --> x:B)"; +by (fast_tac set_cs 1); +val subsetXH = result(); + +(*** Exhaustion Rules ***) + +fun mk_XH_tac thy defs rls s = prove_goalw thy defs s (fn _ => [cfast_tac rls 1]); +val XH_tac = mk_XH_tac Type.thy simp_type_defs []; + +val EmptyXH = XH_tac "a : {} <-> False"; +val SubtypeXH = XH_tac "a : {x:A.P(x)} <-> (a:A & P(a))"; +val UnitXH = XH_tac "a : Unit <-> a=one"; +val BoolXH = XH_tac "a : Bool <-> a=true | a=false"; +val PlusXH = XH_tac "a : A+B <-> (EX x:A.a=inl(x)) | (EX x:B.a=inr(x))"; +val PiXH = XH_tac "a : PROD x:A.B(x) <-> (EX b.a=lam x.b(x) & (ALL x:A.b(x):B(x)))"; +val SgXH = XH_tac "a : SUM x:A.B(x) <-> (EX x:A.EX y:B(x).a=)"; + +val XHs = [EmptyXH,SubtypeXH,UnitXH,BoolXH,PlusXH,PiXH,SgXH]; + +val LiftXH = XH_tac "a : [A] <-> (a=bot | a:A)"; +val TallXH = XH_tac "a : TALL X.B(X) <-> (ALL X. a:B(X))"; +val TexXH = XH_tac "a : TEX X.B(X) <-> (EX X. a:B(X))"; + +val case_rls = XH_to_Es XHs; + +(*** Canonical Type Rules ***) + +fun mk_canT_tac thy xhs s = prove_goal thy s + (fn prems => [fast_tac (set_cs addIs (prems @ (xhs RL [iffD2]))) 1]); +val canT_tac = mk_canT_tac Type.thy XHs; + +val oneT = canT_tac "one : Unit"; +val trueT = canT_tac "true : Bool"; +val falseT = canT_tac "false : Bool"; +val lamT = canT_tac "[| !!x.x:A ==> b(x):B(x) |] ==> lam x.b(x) : Pi(A,B)"; +val pairT = canT_tac "[| a:A; b:B(a) |] ==> :Sigma(A,B)"; +val inlT = canT_tac "a:A ==> inl(a) : A+B"; +val inrT = canT_tac "b:B ==> inr(b) : A+B"; + +val canTs = [oneT,trueT,falseT,pairT,lamT,inlT,inrT]; + +(*** Non-Canonical Type Rules ***) + +local +val lemma = prove_goal Type.thy "[| a:B(u); u=v |] ==> a : B(v)" + (fn prems => [cfast_tac prems 1]); +in +fun mk_ncanT_tac thy defs top_crls crls s = prove_goalw thy defs s + (fn major::prems => [(resolve_tac ([major] RL top_crls) 1), + (REPEAT_SOME (eresolve_tac (crls @ [exE,bexE,conjE,disjE]))), + (ALLGOALS (ASM_SIMP_TAC term_ss)), + (ALLGOALS (ares_tac (prems RL [lemma]) ORELSE' + eresolve_tac [bspec])), + (safe_tac (ccl_cs addSIs prems))]); +end; + +val ncanT_tac = mk_ncanT_tac Type.thy [] case_rls case_rls; + +val ifT = ncanT_tac + "[| b:Bool; b=true ==> t:A(true); b=false ==> u:A(false) |] ==> \ +\ if b then t else u : A(b)"; + +val applyT = ncanT_tac + "[| f : Pi(A,B); a:A |] ==> f ` a : B(a)"; + +val splitT = ncanT_tac + "[| p:Sigma(A,B); !!x y. [| x:A; y:B(x); p= |] ==> c(x,y):C() |] ==> \ +\ split(p,c):C(p)"; + +val whenT = ncanT_tac + "[| p:A+B; !!x.[| x:A; p=inl(x) |] ==> a(x):C(inl(x)); \ +\ !!y.[| y:B; p=inr(y) |] ==> b(y):C(inr(y)) |] ==> \ +\ when(p,a,b) : C(p)"; + +val ncanTs = [ifT,applyT,splitT,whenT]; + +(*** Subtypes ***) + +val SubtypeD1 = standard ((SubtypeXH RS iffD1) RS conjunct1); +val SubtypeD2 = standard ((SubtypeXH RS iffD1) RS conjunct2); + +val prems = goal Type.thy + "[| a:A; P(a) |] ==> a : {x:A. P(x)}"; +by (REPEAT (resolve_tac (prems@[SubtypeXH RS iffD2,conjI]) 1)); +val SubtypeI = result(); + +val prems = goal Type.thy + "[| a : {x:A. P(x)}; [| a:A; P(a) |] ==> Q |] ==> Q"; +by (REPEAT (resolve_tac (prems@[SubtypeD1,SubtypeD2]) 1)); +val SubtypeE = result(); + +(*** Monotonicity ***) + +goal Type.thy "mono (%X.X)"; +by (REPEAT (ares_tac [monoI] 1)); +val idM = result(); + +goal Type.thy "mono(%X.A)"; +by (REPEAT (ares_tac [monoI,subset_refl] 1)); +val constM = result(); + +val major::prems = goal Type.thy + "mono(%X.A(X)) ==> mono(%X.[A(X)])"; +br (subsetI RS monoI) 1; +bd (LiftXH RS iffD1) 1; +be disjE 1; +be (disjI1 RS (LiftXH RS iffD2)) 1; +br (disjI2 RS (LiftXH RS iffD2)) 1; +be (major RS monoD RS subsetD) 1; +ba 1; +val LiftM = result(); + +val prems = goal Type.thy + "[| mono(%X.A(X)); !!x X. x:A(X) ==> mono(%X.B(X,x)) |] ==> \ +\ mono(%X.Sigma(A(X),B(X)))"; +by (REPEAT (ares_tac ([subsetI RS monoI] @ canTs) 1 ORELSE + eresolve_tac ([bspec,exE,conjE,disjE,bexE] @ case_rls) 1 ORELSE + (resolve_tac (prems RL [monoD RS subsetD]) 1 THEN assume_tac 1) ORELSE + hyp_subst_tac 1)); +val SgM = result(); + +val prems = goal Type.thy + "[| !!x. x:A ==> mono(%X.B(X,x)) |] ==> mono(%X.Pi(A,B(X)))"; +by (REPEAT (ares_tac ([subsetI RS monoI] @ canTs) 1 ORELSE + eresolve_tac ([bspec,exE,conjE,disjE,bexE] @ case_rls) 1 ORELSE + (resolve_tac (prems RL [monoD RS subsetD]) 1 THEN assume_tac 1) ORELSE + hyp_subst_tac 1)); +val PiM = result(); + +val prems = goal Type.thy + "[| mono(%X.A(X)); mono(%X.B(X)) |] ==> mono(%X.A(X)+B(X))"; +by (REPEAT (ares_tac ([subsetI RS monoI] @ canTs) 1 ORELSE + eresolve_tac ([bspec,exE,conjE,disjE,bexE] @ case_rls) 1 ORELSE + (resolve_tac (prems RL [monoD RS subsetD]) 1 THEN assume_tac 1) ORELSE + hyp_subst_tac 1)); +val PlusM = result(); + +(**************** RECURSIVE TYPES ******************) + +(*** Conversion Rules for Fixed Points via monotonicity and Tarski ***) + +goal Type.thy "mono(%X.Unit+X)"; +by (REPEAT (ares_tac [PlusM,constM,idM] 1)); +val NatM = result(); +val def_NatB = result() RS (Nat_def RS def_lfp_Tarski); + +goal Type.thy "mono(%X.(Unit+Sigma(A,%y.X)))"; +by (REPEAT (ares_tac [PlusM,SgM,constM,idM] 1)); +val ListM = result(); +val def_ListB = result() RS (List_def RS def_lfp_Tarski); +val def_ListsB = result() RS (Lists_def RS def_gfp_Tarski); + +goal Type.thy "mono(%X.({} + Sigma(A,%y.X)))"; +by (REPEAT (ares_tac [PlusM,SgM,constM,idM] 1)); +val IListsM = result(); +val def_IListsB = result() RS (ILists_def RS def_gfp_Tarski); + +val ind_type_eqs = [def_NatB,def_ListB,def_ListsB,def_IListsB]; + +(*** Exhaustion Rules ***) + +fun mk_iXH_tac teqs ddefs rls s = prove_goalw Type.thy ddefs s + (fn _ => [resolve_tac (teqs RL [XHlemma1]) 1, + fast_tac (set_cs addSIs canTs addSEs case_rls) 1]); + +val iXH_tac = mk_iXH_tac ind_type_eqs ind_data_defs []; + +val NatXH = iXH_tac "a : Nat <-> (a=zero | (EX x:Nat.a=succ(x)))"; +val ListXH = iXH_tac "a : List(A) <-> (a=[] | (EX x:A.EX xs:List(A).a=x.xs))"; +val ListsXH = iXH_tac "a : Lists(A) <-> (a=[] | (EX x:A.EX xs:Lists(A).a=x.xs))"; +val IListsXH = iXH_tac "a : ILists(A) <-> (EX x:A.EX xs:ILists(A).a=x.xs)"; + +val iXHs = [NatXH,ListXH]; +val icase_rls = XH_to_Es iXHs; + +(*** Type Rules ***) + +val icanT_tac = mk_canT_tac Type.thy iXHs; +val incanT_tac = mk_ncanT_tac Type.thy [] icase_rls case_rls; + +val zeroT = icanT_tac "zero : Nat"; +val succT = icanT_tac "n:Nat ==> succ(n) : Nat"; +val nilT = icanT_tac "[] : List(A)"; +val consT = icanT_tac "[| h:A; t:List(A) |] ==> h.t : List(A)"; + +val icanTs = [zeroT,succT,nilT,consT]; + +val ncaseT = incanT_tac + "[| n:Nat; n=zero ==> b:C(zero); \ +\ !!x.[| x:Nat; n=succ(x) |] ==> c(x):C(succ(x)) |] ==> \ +\ ncase(n,b,c) : C(n)"; + +val lcaseT = incanT_tac + "[| l:List(A); l=[] ==> b:C([]); \ +\ !!h t.[| h:A; t:List(A); l=h.t |] ==> c(h,t):C(h.t) |] ==> \ +\ lcase(l,b,c) : C(l)"; + +val incanTs = [ncaseT,lcaseT]; + +(*** Induction Rules ***) + +val ind_Ms = [NatM,ListM]; + +fun mk_ind_tac ddefs tdefs Ms canTs case_rls s = prove_goalw Type.thy ddefs s + (fn major::prems => [resolve_tac (Ms RL ([major] RL (tdefs RL [def_induct]))) 1, + fast_tac (set_cs addSIs (prems @ canTs) addSEs case_rls) 1]); + +val ind_tac = mk_ind_tac ind_data_defs ind_type_defs ind_Ms canTs case_rls; + +val Nat_ind = ind_tac + "[| n:Nat; P(zero); !!x.[| x:Nat; P(x) |] ==> P(succ(x)) |] ==> \ +\ P(n)"; + +val List_ind = ind_tac + "[| l:List(A); P([]); \ +\ !!x xs.[| x:A; xs:List(A); P(xs) |] ==> P(x.xs) |] ==> \ +\ P(l)"; + +val inds = [Nat_ind,List_ind]; + +(*** Primitive Recursive Rules ***) + +fun mk_prec_tac inds s = prove_goal Type.thy s + (fn major::prems => [resolve_tac ([major] RL inds) 1, + ALLGOALS (SIMP_TAC term_ss THEN' + fast_tac (set_cs addSIs prems))]); +val prec_tac = mk_prec_tac inds; + +val nrecT = prec_tac + "[| n:Nat; b:C(zero); \ +\ !!x g.[| x:Nat; g:C(x) |] ==> c(x,g):C(succ(x)) |] ==> \ +\ nrec(n,b,c) : C(n)"; + +val lrecT = prec_tac + "[| l:List(A); b:C([]); \ +\ !!x xs g.[| x:A; xs:List(A); g:C(xs) |] ==> c(x,xs,g):C(x.xs) |] ==> \ +\ lrec(l,b,c) : C(l)"; + +val precTs = [nrecT,lrecT]; + + +(*** Theorem proving ***) + +val [major,minor] = goal Type.thy + "[| : Sigma(A,B); [| a:A; b:B(a) |] ==> P \ +\ |] ==> P"; +br (major RS (XH_to_E SgXH)) 1; +br minor 1; +by (ALLGOALS (fast_tac term_cs)); +val SgE2 = result(); + +(* General theorem proving ignores non-canonical term-formers, *) +(* - intro rules are type rules for canonical terms *) +(* - elim rules are case rules (no non-canonical terms appear) *) + +val type_cs = term_cs addSIs (SubtypeI::(canTs @ icanTs)) + addSEs (SubtypeE::(XH_to_Es XHs)); + + +(*** Infinite Data Types ***) + +val [mono] = goal Type.thy "mono(f) ==> lfp(f) <= gfp(f)"; +br (lfp_lowerbound RS subset_trans) 1; +br (mono RS gfp_lemma3) 1; +br subset_refl 1; +val lfp_subset_gfp = result(); + +val prems = goal Type.thy + "[| a:A; !!x X.[| x:A; ALL y:A.t(y):X |] ==> t(x) : B(X) |] ==> \ +\ t(a) : gfp(B)"; +br coinduct 1; +by (res_inst_tac [("P","%x.EX y:A.x=t(y)")] CollectI 1); +by (ALLGOALS (fast_tac (ccl_cs addSIs prems))); +val gfpI = result(); + +val rew::prem::prems = goal Type.thy + "[| C==gfp(B); a:A; !!x X.[| x:A; ALL y:A.t(y):X |] ==> t(x) : B(X) |] ==> \ +\ t(a) : C"; +by (rewtac rew); +by (REPEAT (ares_tac ((prem RS gfpI)::prems) 1)); +val def_gfpI = result(); + +(* EG *) + +val prems = goal Type.thy + "letrec g x be zero.g(x) in g(bot) : Lists(Nat)"; +by (rtac (refl RS (XH_to_I UnitXH) RS (Lists_def RS def_gfpI)) 1); +br (letrecB RS ssubst) 1; +bw cons_def; +by (fast_tac type_cs 1); +result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/Type.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Type.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,73 @@ +(* Title: CCL/types.thy + ID: $Id$ + Author: Martin Coen + Copyright 1993 University of Cambridge + +Types in CCL are defined as sets of terms. + +*) + +Type = Term + + +consts + + Subtype :: "['a set, 'a => o] => 'a set" + Bool :: "i set" + Unit :: "i set" + "+" :: "[i set, i set] => i set" (infixr 55) + Pi :: "[i set, i => i set] => i set" + Sigma :: "[i set, i => i set] => i set" + Nat :: "i set" + List :: "i set => i set" + Lists :: "i set => i set" + ILists :: "i set => i set" + TAll :: "(i set => i set) => i set" (binder "TALL " 55) + TEx :: "(i set => i set) => i set" (binder "TEX " 55) + Lift :: "i set => i set" ("(3[_])") + + SPLIT :: "[i, [i, i] => i set] => i set" + + "@Pi" :: "[idt, i set, i set] => i set" ("(3PROD _:_./ _)" [] 60) + "@Sigma" :: "[idt, i set, i set] => i set" ("(3SUM _:_./ _)" [] 60) + "@->" :: "[i set, i set] => i set" ("(_ ->/ _)" [54, 53] 53) + "@*" :: "[i set, i set] => i set" ("(_ */ _)" [56, 55] 55) + "@Subtype" :: "[idt, 'a set, o] => 'a set" ("(1{_: _ ./ _})") + +translations + "PROD x:A. B" => "Pi(A, %x. B)" + "SUM x:A. B" => "Sigma(A, %x. B)" + "{x: A. B}" == "Subtype(A, %x. B)" + +rules + + Subtype_def "{x:A.P(x)} == {x.x:A & P(x)}" + Unit_def "Unit == {x.x=one}" + Bool_def "Bool == {x.x=true | x=false}" + Plus_def "A+B == {x. (EX a:A.x=inl(a)) | (EX b:B.x=inr(b))}" + Pi_def "Pi(A,B) == {x.EX b.x=lam x.b(x) & (ALL x:A.b(x):B(x))}" + Sigma_def "Sigma(A,B) == {x.EX a:A.EX b:B(a).x=}" + Nat_def "Nat == lfp(% X.Unit + X)" + List_def "List(A) == lfp(% X.Unit + A*X)" + + Lists_def "Lists(A) == gfp(% X.Unit + A*X)" + ILists_def "ILists(A) == gfp(% X.{} + A*X)" + + Tall_def "TALL X.B(X) == Inter({X.EX Y.X=B(Y)})" + Tex_def "TEX X.B(X) == Union({X.EX Y.X=B(Y)})" + Lift_def "[A] == A Un {bot}" + + SPLIT_def "SPLIT(p,B) == Union({A.EX x y.p= & A=B(x,y)})" + +end + + +ML + +val parse_translation = + [("@->", ndependent_tr "Pi"), + ("@*", ndependent_tr "Sigma")]; + +val print_translation = + [("Pi", dependent_tr' ("@Pi", "@->")), + ("Sigma", dependent_tr' ("@Sigma", "@*"))]; + diff -r 000000000000 -r a5a9c433f639 src/CCL/Wfd.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Wfd.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,208 @@ +(* Title: CCL/wf + ID: $Id$ + +For wf.thy. + +Based on + Titles: ZF/wf.ML and HOL/ex/lex-prod + Authors: Lawrence C Paulson and Tobias Nipkow + Copyright 1992 University of Cambridge + +*) + +open Wfd; + +(***********) + +val wfd_congs = mk_congs Wfd.thy ["Wfd","wf","op **","wmap","ListPR"]; + +(***********) + +val [major,prem] = goalw Wfd.thy [Wfd_def] + "[| Wfd(R); \ +\ !!x.[| ALL y. : R --> P(y) |] ==> P(x) |] ==> \ +\ P(a)"; +by (rtac (major RS spec RS mp RS spec RS CollectD) 1); +by (fast_tac (set_cs addSIs [prem RS CollectI]) 1); +val wfd_induct = result(); + +val [p1,p2,p3] = goal Wfd.thy + "[| !!x y. : R ==> Q(x); \ +\ ALL x. (ALL y. : R --> y : P) --> x : P; \ +\ !!x.Q(x) ==> x:P |] ==> a:P"; +br (p2 RS spec RS mp) 1; +by (fast_tac (set_cs addSIs [p1 RS p3]) 1); +val wfd_strengthen_lemma = result(); + +fun wfd_strengthen_tac s i = res_inst_tac [("Q",s)] wfd_strengthen_lemma i THEN + assume_tac (i+1); + +val wfd::prems = goal Wfd.thy "[| Wfd(r); :r; :r |] ==> P"; +by (subgoal_tac "ALL x. :r --> :r --> P" 1); +by (fast_tac (FOL_cs addIs prems) 1); +br (wfd RS wfd_induct) 1; +by (ALLGOALS (fast_tac (ccl_cs addSIs prems))); +val wf_anti_sym = result(); + +val prems = goal Wfd.thy "[| Wfd(r); : r |] ==> P"; +by (rtac wf_anti_sym 1); +by (REPEAT (resolve_tac prems 1)); +val wf_anti_refl = result(); + +(*** Irreflexive transitive closure ***) + +val [prem] = goal Wfd.thy "Wfd(R) ==> Wfd(R^+)"; +by (rewtac Wfd_def); +by (REPEAT (ares_tac [allI,ballI,impI] 1)); +(*must retain the universal formula for later use!*) +by (rtac allE 1 THEN assume_tac 1); +by (etac mp 1); +br (prem RS wfd_induct) 1; +by (rtac (impI RS allI) 1); +by (etac tranclE 1); +by (fast_tac ccl_cs 1); +be (spec RS mp RS spec RS mp) 1; +by (REPEAT (atac 1)); +val trancl_wf = result(); + +(*** Lexicographic Ordering ***) + +goalw Wfd.thy [lex_def] + "p : ra**rb <-> (EX a a' b b'.p = <,> & ( : ra | a=a' & : rb))"; +by (fast_tac ccl_cs 1); +val lexXH = result(); + +val prems = goal Wfd.thy + " : ra ==> <,> : ra**rb"; +by (fast_tac (ccl_cs addSIs (prems @ [lexXH RS iffD2])) 1); +val lexI1 = result(); + +val prems = goal Wfd.thy + " : rb ==> <,> : ra**rb"; +by (fast_tac (ccl_cs addSIs (prems @ [lexXH RS iffD2])) 1); +val lexI2 = result(); + +val major::prems = goal Wfd.thy + "[| p : ra**rb; \ +\ !!a a' b b'.[| : ra; p=<,> |] ==> R; \ +\ !!a b b'.[| : rb; p = <,> |] ==> R |] ==> \ +\ R"; +br (major RS (lexXH RS iffD1) RS exE) 1; +by (REPEAT_SOME (eresolve_tac ([exE,conjE,disjE]@prems))); +by (ALLGOALS (fast_tac ccl_cs)); +val lexE = result(); + +val [major,minor] = goal Wfd.thy + "[| p : r**s; !!a a' b b'. p = <,> ==> P |] ==>P"; +br (major RS lexE) 1; +by (ALLGOALS (fast_tac (set_cs addSEs [minor]))); +val lex_pair = result(); + +val [wfa,wfb] = goal Wfd.thy + "[| Wfd(R); Wfd(S) |] ==> Wfd(R**S)"; +bw Wfd_def; +by (safe_tac ccl_cs); +by (wfd_strengthen_tac "%x.EX a b.x=" 1); +by (fast_tac (term_cs addSEs [lex_pair]) 1); +by (subgoal_tac "ALL a b.:P" 1); +by (fast_tac ccl_cs 1); +br (wfa RS wfd_induct RS allI) 1; +br (wfb RS wfd_induct RS allI) 1;back(); +by (fast_tac (type_cs addSEs [lexE]) 1); +val lex_wf = result(); + +(*** Mapping ***) + +goalw Wfd.thy [wmap_def] + "p : wmap(f,r) <-> (EX x y. p= & : r)"; +by (fast_tac ccl_cs 1); +val wmapXH = result(); + +val prems = goal Wfd.thy + " : r ==> : wmap(f,r)"; +by (fast_tac (ccl_cs addSIs (prems @ [wmapXH RS iffD2])) 1); +val wmapI = result(); + +val major::prems = goal Wfd.thy + "[| p : wmap(f,r); !!a b.[| : r; p= |] ==> R |] ==> R"; +br (major RS (wmapXH RS iffD1) RS exE) 1; +by (REPEAT_SOME (eresolve_tac ([exE,conjE,disjE]@prems))); +by (ALLGOALS (fast_tac ccl_cs)); +val wmapE = result(); + +val [wf] = goal Wfd.thy + "Wfd(r) ==> Wfd(wmap(f,r))"; +bw Wfd_def; +by (safe_tac ccl_cs); +by (subgoal_tac "ALL b.ALL a.f(a)=b-->a:P" 1); +by (fast_tac ccl_cs 1); +br (wf RS wfd_induct RS allI) 1; +by (safe_tac ccl_cs); +be (spec RS mp) 1; +by (safe_tac (ccl_cs addSEs [wmapE])); +be (spec RS mp RS spec RS mp) 1; +ba 1; +br refl 1; +val wmap_wf = result(); + +(* Projections *) + +val prems = goal Wfd.thy " : r ==> <,> : wmap(fst,r)"; +br wmapI 1; +by (SIMP_TAC (term_ss addrews prems) 1); +val wfstI = result(); + +val prems = goal Wfd.thy " : r ==> <,> : wmap(snd,r)"; +br wmapI 1; +by (SIMP_TAC (term_ss addrews prems) 1); +val wsndI = result(); + +val prems = goal Wfd.thy " : r ==> <>,>> : wmap(thd,r)"; +br wmapI 1; +by (SIMP_TAC (term_ss addrews prems) 1); +val wthdI = result(); + +(*** Ground well-founded relations ***) + +val prems = goalw Wfd.thy [wf_def] + "[| Wfd(r); a : r |] ==> a : wf(r)"; +by (fast_tac (set_cs addSIs prems) 1); +val wfI = result(); + +val prems = goalw Wfd.thy [Wfd_def] "Wfd({})"; +by (fast_tac (set_cs addEs [EmptyXH RS iffD1 RS FalseE]) 1); +val Empty_wf = result(); + +val prems = goalw Wfd.thy [wf_def] "Wfd(wf(R))"; +by (res_inst_tac [("Q","Wfd(R)")] (excluded_middle RS disjE) 1); +by (ALLGOALS (ASM_SIMP_TAC (CCL_ss addcongs wfd_congs))); +br Empty_wf 1; +val wf_wf = result(); + +goalw Wfd.thy [NatPR_def] "p : NatPR <-> (EX x:Nat.p=)"; +by (fast_tac set_cs 1); +val NatPRXH = result(); + +goalw Wfd.thy [ListPR_def] "p : ListPR(A) <-> (EX h:A.EX t:List(A).p=)"; +by (fast_tac set_cs 1); +val ListPRXH = result(); + +val NatPRI = refl RS (bexI RS (NatPRXH RS iffD2)); +val ListPRI = refl RS (bexI RS (bexI RS (ListPRXH RS iffD2))); + +goalw Wfd.thy [Wfd_def] "Wfd(NatPR)"; +by (safe_tac set_cs); +by (wfd_strengthen_tac "%x.x:Nat" 1); +by (fast_tac (type_cs addSEs [XH_to_E NatPRXH]) 1); +be Nat_ind 1; +by (ALLGOALS (fast_tac (type_cs addEs [XH_to_E NatPRXH]))); +val NatPR_wf = result(); + +goalw Wfd.thy [Wfd_def] "Wfd(ListPR(A))"; +by (safe_tac set_cs); +by (wfd_strengthen_tac "%x.x:List(A)" 1); +by (fast_tac (type_cs addSEs [XH_to_E ListPRXH]) 1); +be List_ind 1; +by (ALLGOALS (fast_tac (type_cs addEs [XH_to_E ListPRXH]))); +val ListPR_wf = result(); + diff -r 000000000000 -r a5a9c433f639 src/CCL/Wfd.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/Wfd.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,34 @@ +(* Title: CCL/wf.thy + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Well-founded relations in CCL. +*) + +Wfd = Trancl + Type + + +consts + (*** Predicates ***) + Wfd :: "[i set] => o" + (*** Relations ***) + wf :: "[i set] => i set" + wmap :: "[i=>i,i set] => i set" + "**" :: "[i set,i set] => i set" (infixl 70) + NatPR :: "i set" + ListPR :: "i set => i set" + +rules + + Wfd_def + "Wfd(R) == ALL P.(ALL x.(ALL y. : R --> y:P) --> x:P) --> (ALL a.a:P)" + + wf_def "wf(R) == {x.x:R & Wfd(R)}" + + wmap_def "wmap(f,R) == {p. EX x y. p= & : R}" + lex_def + "ra**rb == {p. EX a a' b b'.p = <,> & ( : ra | (a=a' & : rb))}" + + NatPR_def "NatPR == {p.EX x:Nat. p=}" + ListPR_def "ListPR(A) == {p.EX h:A.EX t:List(A). p=}" +end diff -r 000000000000 -r a5a9c433f639 src/CCL/ccl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ccl.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,362 @@ +(* Title: CCL/ccl + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For ccl.thy. +*) + +open CCL; + +val ccl_data_defs = [apply_def,fix_def]; + +(*** Simplifier for pre-order and equality ***) + +structure CCL_SimpData : SIMP_DATA = + struct + val refl_thms = [refl, po_refl, iff_refl] + val trans_thms = [trans, iff_trans, po_trans] + val red1 = iffD1 + val red2 = iffD2 + val mk_rew_rules = mk_rew_rules + val case_splits = [] (*NO IF'S!*) + val norm_thms = norm_thms + val subst_thms = [subst]; + val dest_red = dest_red + end; + +structure CCL_Simp = SimpFun(CCL_SimpData); +open CCL_Simp; + +val auto_ss = empty_ss setauto (fn hyps => ares_tac (TrueI::hyps)); + +val po_refl_iff_T = make_iff_T po_refl; + +val CCL_ss = auto_ss addcongs (FOL_congs @ set_congs) + addrews ([po_refl_iff_T] @ FOL_rews @ mem_rews); + +(*** Congruence Rules ***) + +(*similar to AP_THM in Gordon's HOL*) +val fun_cong = prove_goal CCL.thy "(f::'a=>'b) = g ==> f(x)=g(x)" + (fn [prem] => [rtac (prem RS subst) 1, rtac refl 1]); + +(*similar to AP_TERM in Gordon's HOL and FOL's subst_context*) +val arg_cong = prove_goal CCL.thy "x=y ==> f(x)=f(y)" + (fn [prem] => [rtac (prem RS subst) 1, rtac refl 1]); + +goal CCL.thy "(ALL x. f(x) = g(x)) --> (%x.f(x)) = (%x.g(x))"; +by (SIMP_TAC (CCL_ss addrews [eq_iff]) 1); +by (fast_tac (set_cs addIs [po_abstractn]) 1); +val abstractn = standard (allI RS (result() RS mp)); + +fun type_of_terms (Const("Trueprop",_) $ + (Const("op =",(Type ("fun", [t,_]))) $ _ $ _)) = t; + +fun abs_prems thm = + let fun do_abs n thm (Type ("fun", [_,t])) = do_abs n (abstractn RSN (n,thm)) t + | do_abs n thm _ = thm + fun do_prems n [] thm = thm + | do_prems n (x::xs) thm = do_prems (n+1) xs (do_abs n thm (type_of_terms x)); + in do_prems 1 (prems_of thm) thm + end; + +fun ccl_mk_congs thy cs = map abs_prems (mk_congs thy cs); + +val ccl_congs = ccl_mk_congs CCL.thy + ["op [=","SIM","POgen","EQgen","pair","lambda","case","op `","fix"]; + +val caseBs = [caseBtrue,caseBfalse,caseBpair,caseBlam,caseBbot]; + +(*** Termination and Divergence ***) + +goalw CCL.thy [Trm_def,Dvg_def] "Trm(t) <-> ~ t = bot"; +br iff_refl 1; +val Trm_iff = result(); + +goalw CCL.thy [Trm_def,Dvg_def] "Dvg(t) <-> t = bot"; +br iff_refl 1; +val Dvg_iff = result(); + +(*** Constructors are injective ***) + +val prems = goal CCL.thy + "[| x=a; y=b; x=y |] ==> a=b"; +by (REPEAT (SOMEGOAL (ares_tac (prems@[box_equals])))); +val eq_lemma = result(); + +fun mk_inj_rl thy rews congs s = + let fun mk_inj_lemmas r = ([arg_cong] RL [(r RS (r RS eq_lemma))]); + val inj_lemmas = flat (map mk_inj_lemmas rews); + val tac = REPEAT (ares_tac [iffI,allI,conjI] 1 ORELSE + eresolve_tac inj_lemmas 1 ORELSE + ASM_SIMP_TAC (CCL_ss addrews rews + addcongs congs) 1) + in prove_goal thy s (fn _ => [tac]) + end; + +val ccl_injs = map (mk_inj_rl CCL.thy caseBs ccl_congs) + [" = <-> (a=a' & b=b')", + "(lam x.b(x) = lam x.b'(x)) <-> ((ALL z.b(z)=b'(z)))"]; + +val pair_inject = ((hd ccl_injs) RS iffD1) RS conjE; + +(*** Constructors are distinct ***) + +local + fun pairs_of f x [] = [] + | pairs_of f x (y::ys) = (f x y) :: (f y x) :: (pairs_of f x ys); + + fun mk_combs ff [] = [] + | mk_combs ff (x::xs) = (pairs_of ff x xs) @ mk_combs ff xs; + +(* Doesn't handle binder types correctly *) + fun saturate thy sy name = + let fun arg_str 0 a s = s + | arg_str 1 a s = "(" ^ a ^ "a" ^ s ^ ")" + | arg_str n a s = arg_str (n-1) a ("," ^ a ^ (chr((ord "a")+n-1)) ^ s); + val sg = sign_of thy; + val T = case Sign.Symtab.lookup(#const_tab(Sign.rep_sg sg),sy) of + None => error(sy^" not declared") | Some(T) => T; + val arity = length (fst (strip_type T)); + in sy ^ (arg_str arity name "") end; + + fun mk_thm_str thy a b = "~ " ^ (saturate thy a "a") ^ " = " ^ (saturate thy b "b"); + + val lemma = prove_goal CCL.thy "t=t' --> case(t,b,c,d,e) = case(t',b,c,d,e)" + (fn _ => [SIMP_TAC (CCL_ss addcongs ccl_congs) 1]) RS mp; + fun mk_lemma (ra,rb) = [lemma] RL [ra RS (rb RS eq_lemma)] RL + [distinctness RS notE,sym RS (distinctness RS notE)]; +in + fun mk_lemmas rls = flat (map mk_lemma (mk_combs pair rls)); + fun mk_dstnct_rls thy xs = mk_combs (mk_thm_str thy) xs; +end; + + +val caseB_lemmas = mk_lemmas caseBs; + +val ccl_dstncts = + let fun mk_raw_dstnct_thm rls s = + prove_goal CCL.thy s (fn _=> [rtac notI 1,eresolve_tac rls 1]) + in map (mk_raw_dstnct_thm caseB_lemmas) + (mk_dstnct_rls CCL.thy ["bot","true","false","pair","lambda"]) end; + +fun mk_dstnct_thms thy defs inj_rls xs = + let fun mk_dstnct_thm rls s = prove_goalw thy defs s + (fn _ => [SIMP_TAC (CCL_ss addrews (rls@inj_rls)) 1]) + in map (mk_dstnct_thm ccl_dstncts) (mk_dstnct_rls thy xs) end; + +fun mkall_dstnct_thms thy defs i_rls xss = flat (map (mk_dstnct_thms thy defs i_rls) xss); + +(*** Rewriting and Proving ***) + +fun XH_to_I rl = rl RS iffD2; +fun XH_to_D rl = rl RS iffD1; +val XH_to_E = make_elim o XH_to_D; +val XH_to_Is = map XH_to_I; +val XH_to_Ds = map XH_to_D; +val XH_to_Es = map XH_to_E; + +val ccl_rews = caseBs @ ccl_injs @ ccl_dstncts; +val ccl_ss = CCL_ss addrews ccl_rews addcongs ccl_congs; + +val ccl_cs = set_cs addSEs (pair_inject::(ccl_dstncts RL [notE])) + addSDs (XH_to_Ds ccl_injs); + +(****** Facts from gfp Definition of [= and = ******) + +val major::prems = goal Set.thy "[| A=B; a:B <-> P |] ==> a:A <-> P"; +brs (prems RL [major RS ssubst]) 1; +val XHlemma1 = result(); + +goal CCL.thy "(P(t,t') <-> Q) --> ( : {p.EX t t'.p= & P(t,t')} <-> Q)"; +by (fast_tac ccl_cs 1); +val XHlemma2 = result() RS mp; + +(*** Pre-Order ***) + +goalw CCL.thy [POgen_def,SIM_def] "mono(%X.POgen(X))"; +br monoI 1; +by (safe_tac ccl_cs); +by (REPEAT_SOME (resolve_tac [exI,conjI,refl])); +by (ALLGOALS (SIMP_TAC ccl_ss)); +by (ALLGOALS (fast_tac set_cs)); +val POgen_mono = result(); + +goalw CCL.thy [POgen_def,SIM_def] + " : POgen(R) <-> t= bot | (t=true & t'=true) | (t=false & t'=false) | \ +\ (EX a a' b b'.t= & t'= & : R & : R) | \ +\ (EX f f'.t=lam x.f(x) & t'=lam x.f'(x) & (ALL x. : R))"; +br (iff_refl RS XHlemma2) 1; +val POgenXH = result(); + +goal CCL.thy + "t [= t' <-> t=bot | (t=true & t'=true) | (t=false & t'=false) | \ +\ (EX a a' b b'.t= & t'= & a [= a' & b [= b') | \ +\ (EX f f'.t=lam x.f(x) & t'=lam x.f'(x) & (ALL x.f(x) [= f'(x)))"; +by (SIMP_TAC (ccl_ss addrews [PO_iff]) 1); +br (rewrite_rule [POgen_def,SIM_def] + (POgen_mono RS (PO_def RS def_gfp_Tarski) RS XHlemma1)) 1; +br (iff_refl RS XHlemma2) 1; +val poXH = result(); + +goal CCL.thy "bot [= b"; +br (poXH RS iffD2) 1; +by (SIMP_TAC ccl_ss 1); +val po_bot = result(); + +goal CCL.thy "a [= bot --> a=bot"; +br impI 1; +bd (poXH RS iffD1) 1; +be rev_mp 1; +by (SIMP_TAC ccl_ss 1); +val bot_poleast = result() RS mp; + +goal CCL.thy " [= <-> a [= a' & b [= b'"; +br (poXH RS iff_trans) 1; +by (SIMP_TAC ccl_ss 1); +by (fast_tac ccl_cs 1); +val po_pair = result(); + +goal CCL.thy "lam x.f(x) [= lam x.f'(x) <-> (ALL x. f(x) [= f'(x))"; +br (poXH RS iff_trans) 1; +by (SIMP_TAC ccl_ss 1); +by (REPEAT (ares_tac [iffI,allI] 1 ORELSE eresolve_tac [exE,conjE] 1)); +by (ASM_SIMP_TAC ccl_ss 1); +by (fast_tac ccl_cs 1); +val po_lam = result(); + +val ccl_porews = [po_bot,po_pair,po_lam]; + +val [p1,p2,p3,p4,p5] = goal CCL.thy + "[| t [= t'; a [= a'; b [= b'; !!x y.c(x,y) [= c'(x,y); \ +\ !!u.d(u) [= d'(u) |] ==> case(t,a,b,c,d) [= case(t',a',b',c',d')"; +br (p1 RS po_cong RS po_trans) 1; +br (p2 RS po_cong RS po_trans) 1; +br (p3 RS po_cong RS po_trans) 1; +br (p4 RS po_abstractn RS po_abstractn RS po_cong RS po_trans) 1; +by (res_inst_tac [("f1","%d.case(t',a',b',c',d)")] + (p5 RS po_abstractn RS po_cong RS po_trans) 1); +br po_refl 1; +val case_pocong = result(); + +val [p1,p2] = goalw CCL.thy ccl_data_defs + "[| f [= f'; a [= a' |] ==> f ` a [= f' ` a'"; +by (REPEAT (ares_tac [po_refl,case_pocong,p1,p2 RS po_cong] 1)); +val apply_pocong = result(); + + +val prems = goal CCL.thy "~ lam x.b(x) [= bot"; +br notI 1; +bd bot_poleast 1; +be (distinctness RS notE) 1; +val npo_lam_bot = result(); + +val eq1::eq2::prems = goal CCL.thy + "[| x=a; y=b; x[=y |] ==> a[=b"; +br (eq1 RS subst) 1; +br (eq2 RS subst) 1; +brs prems 1; +val po_lemma = result(); + +goal CCL.thy "~ [= lam x.f(x)"; +br notI 1; +br (npo_lam_bot RS notE) 1; +be (case_pocong RS (caseBlam RS (caseBpair RS po_lemma))) 1; +by (REPEAT (resolve_tac [po_refl,npo_lam_bot] 1)); +val npo_pair_lam = result(); + +goal CCL.thy "~ lam x.f(x) [= "; +br notI 1; +br (npo_lam_bot RS notE) 1; +be (case_pocong RS (caseBpair RS (caseBlam RS po_lemma))) 1; +by (REPEAT (resolve_tac [po_refl,npo_lam_bot] 1)); +val npo_lam_pair = result(); + +fun mk_thm s = prove_goal CCL.thy s (fn _ => + [rtac notI 1,dtac case_pocong 1,etac rev_mp 5, + ALLGOALS (SIMP_TAC ccl_ss), + REPEAT (resolve_tac [po_refl,npo_lam_bot] 1)]); + +val npo_rls = [npo_pair_lam,npo_lam_pair] @ map mk_thm + ["~ true [= false", "~ false [= true", + "~ true [= ", "~ [= true", + "~ true [= lam x.f(x)","~ lam x.f(x) [= true", + "~ false [= ", "~ [= false", + "~ false [= lam x.f(x)","~ lam x.f(x) [= false"]; + +(* Coinduction for [= *) + +val prems = goal CCL.thy "[| : R; R <= POgen(R) |] ==> t [= u"; +br (PO_def RS def_coinduct RS (PO_iff RS iffD2)) 1; +by (REPEAT (ares_tac prems 1)); +val po_coinduct = result(); + +fun po_coinduct_tac s i = res_inst_tac [("R",s)] po_coinduct i; + +(*************** EQUALITY *******************) + +goalw CCL.thy [EQgen_def,SIM_def] "mono(%X.EQgen(X))"; +br monoI 1; +by (safe_tac set_cs); +by (REPEAT_SOME (resolve_tac [exI,conjI,refl])); +by (ALLGOALS (SIMP_TAC ccl_ss)); +by (ALLGOALS (fast_tac set_cs)); +val EQgen_mono = result(); + +goalw CCL.thy [EQgen_def,SIM_def] + " : EQgen(R) <-> (t=bot & t'=bot) | (t=true & t'=true) | \ +\ (t=false & t'=false) | \ +\ (EX a a' b b'.t= & t'= & : R & : R) | \ +\ (EX f f'.t=lam x.f(x) & t'=lam x.f'(x) & (ALL x. : R))"; +br (iff_refl RS XHlemma2) 1; +val EQgenXH = result(); + +goal CCL.thy + "t=t' <-> (t=bot & t'=bot) | (t=true & t'=true) | (t=false & t'=false) | \ +\ (EX a a' b b'.t= & t'= & a=a' & b=b') | \ +\ (EX f f'.t=lam x.f(x) & t'=lam x.f'(x) & (ALL x.f(x)=f'(x)))"; +by (subgoal_tac + " : EQ <-> (t=bot & t'=bot) | (t=true & t'=true) | (t=false & t'=false) | \ +\ (EX a a' b b'.t= & t'= & : EQ & : EQ) | \ +\ (EX f f'.t=lam x.f(x) & t'=lam x.f'(x) & (ALL x. : EQ))" 1); +be rev_mp 1; +by (SIMP_TAC (CCL_ss addrews [EQ_iff RS iff_sym]) 1); +br (rewrite_rule [EQgen_def,SIM_def] + (EQgen_mono RS (EQ_def RS def_gfp_Tarski) RS XHlemma1)) 1; +br (iff_refl RS XHlemma2) 1; +val eqXH = result(); + +val prems = goal CCL.thy "[| : R; R <= EQgen(R) |] ==> t = u"; +br (EQ_def RS def_coinduct RS (EQ_iff RS iffD2)) 1; +by (REPEAT (ares_tac prems 1)); +val eq_coinduct = result(); + +val prems = goal CCL.thy + "[| : R; R <= EQgen(lfp(%x.EQgen(x) Un R Un EQ)) |] ==> t = u"; +br (EQ_def RS def_coinduct3 RS (EQ_iff RS iffD2)) 1; +by (REPEAT (ares_tac (EQgen_mono::prems) 1)); +val eq_coinduct3 = result(); + +fun eq_coinduct_tac s i = res_inst_tac [("R",s)] eq_coinduct i; +fun eq_coinduct3_tac s i = res_inst_tac [("R",s)] eq_coinduct3 i; + +(*** Untyped Case Analysis and Other Facts ***) + +goalw CCL.thy [apply_def] "(EX f.t=lam x.f(x)) --> t = lam x.(t ` x)"; +by (safe_tac ccl_cs); +by (SIMP_TAC ccl_ss 1); +val cond_eta = result() RS mp; + +goal CCL.thy "(t=bot) | (t=true) | (t=false) | (EX a b.t=) | (EX f.t=lam x.f(x))"; +by (cut_facts_tac [refl RS (eqXH RS iffD1)] 1); +by (fast_tac set_cs 1); +val exhaustion = result(); + +val prems = goal CCL.thy + "[| P(bot); P(true); P(false); !!x y.P(); !!b.P(lam x.b(x)) |] ==> P(t)"; +by (cut_facts_tac [exhaustion] 1); +by (REPEAT_SOME (ares_tac prems ORELSE' eresolve_tac [disjE,exE,ssubst])); +val term_case = result(); + +fun term_case_tac a i = res_inst_tac [("t",a)] term_case i; diff -r 000000000000 -r a5a9c433f639 src/CCL/ccl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ccl.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,148 @@ +(* Title: CCL/ccl.thy + ID: $Id$ + Author: Martin Coen + Copyright 1993 University of Cambridge + +Classical Computational Logic for Untyped Lambda Calculus with reduction to +weak head-normal form. + +Based on FOL extended with set collection, a primitive higher-order logic. +HOL is too strong - descriptions prevent a type of programs being defined +which contains only executable terms. +*) + +CCL = Gfp + + +classes prog < term + +default prog + +types i 0 + +arities + i :: prog + fun :: (prog,prog)prog + +consts + (*** Evaluation Judgement ***) + "--->" :: "[i,i]=>prop" (infixl 20) + + (*** Bisimulations for pre-order and equality ***) + "[=" :: "['a,'a]=>o" (infixl 50) + SIM :: "[i,i,i set]=>o" + POgen,EQgen :: "i set => i set" + PO,EQ :: "i set" + + (*** Term Formers ***) + true,false :: "i" + pair :: "[i,i]=>i" ("(1<_,/_>)") + lambda :: "(i=>i)=>i" (binder "lam " 55) + case :: "[i,i,i,[i,i]=>i,(i=>i)=>i]=>i" + "`" :: "[i,i]=>i" (infixl 56) + bot :: "i" + fix :: "(i=>i)=>i" + + (*** Defined Predicates ***) + Trm,Dvg :: "i => o" + +rules + + (******* EVALUATION SEMANTICS *******) + + (** This is the evaluation semantics from which the axioms below were derived. **) + (** It is included here just as an evaluator for FUN and has no influence on **) + (** inference in the theory CCL. **) + + trueV "true ---> true" + falseV "false ---> false" + pairV " ---> " + lamV "lam x.b(x) ---> lam x.b(x)" + caseVtrue "[| t ---> true; d ---> c |] ==> case(t,d,e,f,g) ---> c" + caseVfalse "[| t ---> false; e ---> c |] ==> case(t,d,e,f,g) ---> c" + caseVpair "[| t ---> ; f(a,b) ---> c |] ==> case(t,d,e,f,g) ---> c" + caseVlam "[| t ---> lam x.b(x); g(b) ---> c |] ==> case(t,d,e,f,g) ---> c" + + (*** Properties of evaluation: note that "t ---> c" impies that c is canonical ***) + + canonical "[| t ---> c; c==true ==> u--->v; \ +\ c==false ==> u--->v; \ +\ !!a b.c== ==> u--->v; \ +\ !!f.c==lam x.f(x) ==> u--->v |] ==> \ +\ u--->v" + + (* Should be derivable - but probably a bitch! *) + substitute "[| a==a'; t(a)--->c(a) |] ==> t(a')--->c(a')" + + (************** LOGIC ***************) + + (*** Definitions used in the following rules ***) + + apply_def "f ` t == case(f,bot,bot,%x y.bot,%u.u(t))" + bot_def "bot == (lam x.x`x)`(lam x.x`x)" + fix_def "fix(f) == (lam x.f(x`x))`(lam x.f(x`x))" + + (* The pre-order ([=) is defined as a simulation, and behavioural equivalence (=) *) + (* as a bisimulation. They can both be expressed as (bi)simulations up to *) + (* behavioural equivalence (ie the relations PO and EQ defined below). *) + + SIM_def + "SIM(t,t',R) == (t=true & t'=true) | (t=false & t'=false) | \ +\ (EX a a' b b'.t= & t'= & : R & : R) | \ +\ (EX f f'.t=lam x.f(x) & t'=lam x.f'(x) & (ALL x. : R))" + + POgen_def "POgen(R) == {p. EX t t'. p= & (t = bot | SIM(t,t',R))}" + EQgen_def "EQgen(R) == {p. EX t t'. p= & (t = bot & t' = bot | SIM(t,t',R))}" + + PO_def "PO == gfp(POgen)" + EQ_def "EQ == gfp(EQgen)" + + (*** Rules ***) + + (** Partial Order **) + + po_refl "a [= a" + po_trans "[| a [= b; b [= c |] ==> a [= c" + po_cong "a [= b ==> f(a) [= f(b)" + + (* Extend definition of [= to program fragments of higher type *) + po_abstractn "(!!x. f(x) [= g(x)) ==> (%x.f(x)) [= (%x.g(x))" + + (** Equality - equivalence axioms inherited from FOL.thy **) + (** - congruence of "=" is axiomatised implicitly **) + + eq_iff "t = t' <-> t [= t' & t' [= t" + + (** Properties of canonical values given by greatest fixed point definitions **) + + PO_iff "t [= t' <-> : PO" + EQ_iff "t = t' <-> : EQ" + + (** Behaviour of non-canonical terms (ie case) given by the following beta-rules **) + + caseBtrue "case(true,d,e,f,g) = d" + caseBfalse "case(false,d,e,f,g) = e" + caseBpair "case(,d,e,f,g) = f(a,b)" + caseBlam "case(lam x.b(x),d,e,f,g) = g(b)" + caseBbot "case(bot,d,e,f,g) = bot" (* strictness *) + + (** The theory is non-trivial **) + distinctness "~ lam x.b(x) = bot" + + (*** Definitions of Termination and Divergence ***) + + Dvg_def "Dvg(t) == t = bot" + Trm_def "Trm(t) == ~ Dvg(t)" + +end + + +(* +Would be interesting to build a similar theory for a typed programming language: + ie. true :: bool, fix :: ('a=>'a)=>'a etc...... + +This is starting to look like LCF. +What are the advantages of this approach? + - less axiomatic + - wfd induction / coinduction and fixed point induction available + +*) diff -r 000000000000 -r a5a9c433f639 src/CCL/coinduction.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/coinduction.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,107 @@ +(* Title: 92/CCL/coinduction + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Lemmas and tactics for using the rule coinduct3 on [= and =. +*) + +val [mono,prem] = goal Lfp.thy "[| mono(f); a : f(lfp(f)) |] ==> a : lfp(f)"; +br ((mono RS lfp_Tarski) RS ssubst) 1; +br prem 1; +val lfpI = result(); + +val prems = goal CCL.thy "[| a=a'; a' : A |] ==> a : A"; +by (SIMP_TAC (term_ss addrews prems) 1); +val ssubst_single = result(); + +val prems = goal CCL.thy "[| a=a'; b=b'; : A |] ==> : A"; +by (SIMP_TAC (term_ss addrews prems) 1); +val ssubst_pair = result(); + +(***) + +local +fun mk_thm s = prove_goal Term.thy s (fn mono::prems => + [fast_tac (term_cs addIs ((mono RS coinduct3_mono_lemma RS lfpI)::prems)) 1]); +in +val ci3_RI = mk_thm "[| mono(Agen); a : R |] ==> a : lfp(%x. Agen(x) Un R Un A)"; +val ci3_AgenI = mk_thm "[| mono(Agen); a : Agen(lfp(%x. Agen(x) Un R Un A)) |] ==> \ +\ a : lfp(%x. Agen(x) Un R Un A)"; +val ci3_AI = mk_thm "[| mono(Agen); a : A |] ==> a : lfp(%x. Agen(x) Un R Un A)"; +end; + +fun mk_genIs thy defs genXH gen_mono s = prove_goalw thy defs s + (fn prems => [rtac (genXH RS iffD2) 1, + (SIMP_TAC term_ss 1), + TRY (fast_tac (term_cs addIs + ([genXH RS iffD2,gen_mono RS coinduct3_mono_lemma RS lfpI] + @ prems)) 1)]); + +(** POgen **) + +goal Term.thy " : PO"; +br (po_refl RS (XH_to_D PO_iff)) 1; +val PO_refl = result(); + +val POgenIs = map (mk_genIs Term.thy data_defs POgenXH POgen_mono) + [" : POgen(R)", + " : POgen(R)", + "[| : R; : R |] ==> <,> : POgen(R)", + "[|!!x. : R |] ==> : POgen(R)", + " : POgen(R)", + " : lfp(%x. POgen(x) Un R Un PO) ==> \ +\ : POgen(lfp(%x. POgen(x) Un R Un PO))", + " : lfp(%x. POgen(x) Un R Un PO) ==> \ +\ : POgen(lfp(%x. POgen(x) Un R Un PO))", + " : POgen(lfp(%x. POgen(x) Un R Un PO))", + " : lfp(%x. POgen(x) Un R Un PO) ==> \ +\ : POgen(lfp(%x. POgen(x) Un R Un PO))", + "<[],[]> : POgen(lfp(%x. POgen(x) Un R Un PO))", + "[| : lfp(%x. POgen(x) Un R Un PO); \ +\ : lfp(%x. POgen(x) Un R Un PO) |] ==> \ +\ : POgen(lfp(%x. POgen(x) Un R Un PO))"]; + +fun POgen_tac (rla,rlb) i = + SELECT_GOAL (safe_tac ccl_cs) i THEN + rtac (rlb RS (rla RS ssubst_pair)) i THEN + (REPEAT (resolve_tac (POgenIs @ [PO_refl RS (POgen_mono RS ci3_AI)] @ + (POgenIs RL [POgen_mono RS ci3_AgenI]) @ [POgen_mono RS ci3_RI]) i)); + +(** EQgen **) + +goal Term.thy " : EQ"; +br (refl RS (EQ_iff RS iffD1)) 1; +val EQ_refl = result(); + +val EQgenIs = map (mk_genIs Term.thy data_defs EQgenXH EQgen_mono) + [" : EQgen(R)", + " : EQgen(R)", + "[| : R; : R |] ==> <,> : EQgen(R)", + "[|!!x. : R |] ==> : EQgen(R)", + " : EQgen(R)", + " : lfp(%x. EQgen(x) Un R Un EQ) ==> \ +\ : EQgen(lfp(%x. EQgen(x) Un R Un EQ))", + " : lfp(%x. EQgen(x) Un R Un EQ) ==> \ +\ : EQgen(lfp(%x. EQgen(x) Un R Un EQ))", + " : EQgen(lfp(%x. EQgen(x) Un R Un EQ))", + " : lfp(%x. EQgen(x) Un R Un EQ) ==> \ +\ : EQgen(lfp(%x. EQgen(x) Un R Un EQ))", + "<[],[]> : EQgen(lfp(%x. EQgen(x) Un R Un EQ))", + "[| : lfp(%x. EQgen(x) Un R Un EQ); \ +\ : lfp(%x. EQgen(x) Un R Un EQ) |] ==> \ +\ : EQgen(lfp(%x. EQgen(x) Un R Un EQ))"]; + +fun EQgen_raw_tac i = + (REPEAT (resolve_tac (EQgenIs @ [EQ_refl RS (EQgen_mono RS ci3_AI)] @ + (EQgenIs RL [EQgen_mono RS ci3_AgenI]) @ [EQgen_mono RS ci3_RI]) i)); + +(* Goals of the form R <= EQgen(R) - rewrite elements : EQgen(R) using rews and *) +(* then reduce this to a goal : R (hopefully?) *) +(* rews are rewrite rules that would cause looping in the simpifier *) + +fun EQgen_tac simp_set rews i = + SELECT_GOAL (TRY (safe_tac ccl_cs) THEN + resolve_tac ((rews@[refl]) RL ((rews@[refl]) RL [ssubst_pair])) i THEN + ALLGOALS (SIMP_TAC simp_set) THEN + ALLGOALS EQgen_raw_tac) i; diff -r 000000000000 -r a5a9c433f639 src/CCL/equalities.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/equalities.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,134 @@ +(* Title: CCL/equalities + ID: $Id$ + +Modified version of + Title: HOL/equalities + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Equalities involving union, intersection, inclusion, etc. +*) + +writeln"File HOL/equalities"; + +val eq_cs = set_cs addSIs [equalityI]; + +(** Binary Intersection **) + +goal Set.thy "A Int A = A"; +by (fast_tac eq_cs 1); +val Int_absorb = result(); + +goal Set.thy "A Int B = B Int A"; +by (fast_tac eq_cs 1); +val Int_commute = result(); + +goal Set.thy "(A Int B) Int C = A Int (B Int C)"; +by (fast_tac eq_cs 1); +val Int_assoc = result(); + +goal Set.thy "(A Un B) Int C = (A Int C) Un (B Int C)"; +by (fast_tac eq_cs 1); +val Int_Un_distrib = result(); + +goal Set.thy "(A<=B) <-> (A Int B = A)"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val subset_Int_eq = result(); + +(** Binary Union **) + +goal Set.thy "A Un A = A"; +by (fast_tac eq_cs 1); +val Un_absorb = result(); + +goal Set.thy "A Un B = B Un A"; +by (fast_tac eq_cs 1); +val Un_commute = result(); + +goal Set.thy "(A Un B) Un C = A Un (B Un C)"; +by (fast_tac eq_cs 1); +val Un_assoc = result(); + +goal Set.thy "(A Int B) Un C = (A Un C) Int (B Un C)"; +by (fast_tac eq_cs 1); +val Un_Int_distrib = result(); + +goal Set.thy + "(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)"; +by (fast_tac eq_cs 1); +val Un_Int_crazy = result(); + +goal Set.thy "(A<=B) <-> (A Un B = B)"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val subset_Un_eq = result(); + +(** Simple properties of Compl -- complement of a set **) + +goal Set.thy "A Int Compl(A) = {x.False}"; +by (fast_tac eq_cs 1); +val Compl_disjoint = result(); + +goal Set.thy "A Un Compl(A) = {x.True}"; +by (fast_tac eq_cs 1); +val Compl_partition = result(); + +goal Set.thy "Compl(Compl(A)) = A"; +by (fast_tac eq_cs 1); +val double_complement = result(); + +goal Set.thy "Compl(A Un B) = Compl(A) Int Compl(B)"; +by (fast_tac eq_cs 1); +val Compl_Un = result(); + +goal Set.thy "Compl(A Int B) = Compl(A) Un Compl(B)"; +by (fast_tac eq_cs 1); +val Compl_Int = result(); + +goal Set.thy "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))"; +by (fast_tac eq_cs 1); +val Compl_UN = result(); + +goal Set.thy "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))"; +by (fast_tac eq_cs 1); +val Compl_INT = result(); + +(*Halmos, Naive Set Theory, page 16.*) + +goal Set.thy "((A Int B) Un C = A Int (B Un C)) <-> (C<=A)"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val Un_Int_assoc_eq = result(); + + +(** Big Union and Intersection **) + +goal Set.thy "Union(A Un B) = Union(A) Un Union(B)"; +by (fast_tac eq_cs 1); +val Union_Un_distrib = result(); + +val prems = goal Set.thy + "(Union(C) Int A = {x.False}) <-> (ALL B:C. B Int A = {x.False})"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val Union_disjoint = result(); + +goal Set.thy "Inter(A Un B) = Inter(A) Int Inter(B)"; +by (best_tac eq_cs 1); +val Inter_Un_distrib = result(); + +(** Unions and Intersections of Families **) + +goal Set.thy "(UN x:A. B(x)) = Union({Y. EX x:A. Y=B(x)})"; +by (fast_tac eq_cs 1); +val UN_eq = result(); + +(*Look: it has an EXISTENTIAL quantifier*) +goal Set.thy "(INT x:A. B(x)) = Inter({Y. EX x:A. Y=B(x)})"; +by (fast_tac eq_cs 1); +val INT_eq = result(); + +goal Set.thy "A Int Union(B) = (UN C:B. A Int C)"; +by (fast_tac eq_cs 1); +val Int_Union_image = result(); + +goal Set.thy "A Un Inter(B) = (INT C:B. A Un C)"; +by (fast_tac eq_cs 1); +val Un_Inter_image = result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/eval.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/eval.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,104 @@ +(* Title: 92/CCL/eval + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +*) + + + +(*** Evaluation ***) + +val EVal_rls = ref [trueV,falseV,pairV,lamV,caseVtrue,caseVfalse,caseVpair,caseVlam]; +val eval_tac = DEPTH_SOLVE_1 (resolve_tac (!EVal_rls) 1); +fun ceval_tac rls = DEPTH_SOLVE_1 (resolve_tac (!EVal_rls@rls) 1); + +val prems = goalw thy [apply_def] + "[| f ---> lam x.b(x); b(a) ---> c |] ==> f ` a ---> c"; +by (ceval_tac prems); +val applyV = result(); + +EVal_rls := !EVal_rls @ [applyV]; + +val major::prems = goalw thy [let_def] + "[| t ---> a; f(a) ---> c |] ==> let x be t in f(x) ---> c"; +br (major RS canonical) 1; +by (REPEAT (DEPTH_SOLVE_1 (resolve_tac ([major]@prems@(!EVal_rls)) 1 ORELSE + eresolve_tac [substitute] 1))); +val letV = result(); + +val prems = goalw thy [fix_def] + "f(fix(f)) ---> c ==> fix(f) ---> c"; +br applyV 1; +br lamV 1; +brs prems 1; +val fixV = result(); + +val prems = goalw thy [letrec_def] + "h(t,%y.letrec g x be h(x,g) in g(y)) ---> c ==> \ +\ letrec g x be h(x,g) in g(t) ---> c"; +by (REPEAT (resolve_tac (prems @ [fixV,applyV,lamV]) 1)); +val letrecV = result(); + +EVal_rls := !EVal_rls @ [letV,letrecV,fixV]; + +fun mk_V_rl s = prove_goalw thy data_defs s (fn prems => [ceval_tac prems]); + +val V_rls = map mk_V_rl + ["true ---> true", + "false ---> false", + "[| b--->true; t--->c |] ==> if b then t else u ---> c", + "[| b--->false; u--->c |] ==> if b then t else u ---> c", + " ---> ", + "[| t ---> ; h(a,b) ---> c |] ==> split(t,h) ---> c", + "zero ---> zero", + "succ(n) ---> succ(n)", + "[| n ---> zero; t ---> c |] ==> ncase(n,t,u) ---> c", + "[| n ---> succ(x); u(x) ---> c |] ==> ncase(n,t,u) ---> c", + "[| n ---> zero; t ---> c |] ==> nrec(n,t,u) ---> c", + "[| n--->succ(x); u(x,nrec(x,t,u))--->c |] ==> nrec(n,t,u)--->c", + "[] ---> []", + "h.t ---> h.t", + "[| l ---> []; t ---> c |] ==> lcase(l,t,u) ---> c", + "[| l ---> x.xs; u(x,xs) ---> c |] ==> lcase(l,t,u) ---> c", + "[| l ---> []; t ---> c |] ==> lrec(l,t,u) ---> c", + "[| l--->x.xs; u(x,xs,lrec(xs,t,u))--->c |] ==> lrec(l,t,u)--->c"]; + +EVal_rls := !EVal_rls @ V_rls; + +(* Factorial *) + +val prems = goal thy + "letrec f n be ncase(n,succ(zero),%x.nrec(n,zero,%y g.nrec(f(x),g,%z h.succ(h)))) \ +\ in f(succ(succ(zero))) ---> ?a"; +by (ceval_tac []); + +val prems = goal thy + "letrec f n be ncase(n,succ(zero),%x.nrec(n,zero,%y g.nrec(f(x),g,%z h.succ(h)))) \ +\ in f(succ(succ(succ(zero)))) ---> ?a"; +by (ceval_tac []); + +(* Less Than Or Equal *) + +fun isle x y = prove_goal thy + ("letrec f p be split(p,%m n.ncase(m,true,%x.ncase(n,false,%y.f()))) \ +\ in f(<"^x^","^y^">) ---> ?a") + (fn prems => [ceval_tac []]); + +isle "succ(zero)" "succ(zero)"; +isle "succ(zero)" "succ(succ(succ(succ(zero))))"; +isle "succ(succ(succ(succ(succ(zero)))))" "succ(succ(succ(succ(zero))))"; + + +(* Reverse *) + +val prems = goal thy + "letrec id l be lcase(l,[],%x xs.x.id(xs)) \ +\ in id(zero.succ(zero).[]) ---> ?a"; +by (ceval_tac []); + +val prems = goal thy + "letrec rev l be lcase(l,[],%x xs.lrec(rev(xs),x.[],%y ys g.y.g)) \ +\ in rev(zero.succ(zero).(succ((lam x.x)`succ(zero))).([])) ---> ?a"; +by (ceval_tac []); + diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/Flag.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/Flag.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,46 @@ +(* Title: CCL/ex/flag + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For flag.thy. +*) + +open Flag; + +(******) + +val flag_defs = [Colour_def,red_def,white_def,blue_def,ccase_def]; + +(******) + +val ColourXH = mk_XH_tac Flag.thy (simp_type_defs @flag_defs) [] + "a : Colour <-> (a=red | a=white | a=blue)"; + +val Colour_case = XH_to_E ColourXH; + +val redT = mk_canT_tac Flag.thy [ColourXH] "red : Colour"; +val whiteT = mk_canT_tac Flag.thy [ColourXH] "white : Colour"; +val blueT = mk_canT_tac Flag.thy [ColourXH] "blue : Colour"; + + +val ccaseT = mk_ncanT_tac Flag.thy flag_defs case_rls case_rls + "[| c:Colour; \ +\ c=red ==> r : C(red); c=white ==> w : C(white); c=blue ==> b : C(blue) |] ==> \ +\ ccase(c,r,w,b) : C(c)"; + +(***) + +val prems = goalw Flag.thy [flag_def] + "flag : List(Colour)->List(Colour)*List(Colour)*List(Colour)"; +by (typechk_tac [redT,whiteT,blueT,ccaseT] 1); +by clean_ccs_tac; +be (ListPRI RS (ListPR_wf RS wfI)) 1; +ba 1; +result(); + + +val prems = goalw Flag.thy [flag_def] + "flag : PROD l:List(Colour).{x:List(Colour)*List(Colour)*List(Colour).FLAG(x,l)}"; +by (gen_ccs_tac [redT,whiteT,blueT,ccaseT] 1); +by (REPEAT_SOME (ares_tac [ListPRI RS (ListPR_wf RS wfI)])); diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/Flag.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/Flag.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,48 @@ +(* Title: CCL/ex/flag.thy + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Dutch national flag program - except that the point of Dijkstra's example was to use +arrays and this uses lists. + +*) + +Flag = List + + +consts + + Colour :: "i set" + red, white, blue :: "i" + ccase :: "[i,i,i,i]=>i" + flag :: "i" + +rules + + Colour_def "Colour == Unit + Unit + Unit" + red_def "red == inl(one)" + white_def "white == inr(inl(one))" + blue_def "blue == inr(inr(one))" + + ccase_def "ccase(c,r,w,b) == when(c,%x.r,%wb.when(wb,%x.w,%x.b))" + + flag_def "flag == lam l.letrec \ +\ flagx l be lcase(l,<[],<[],[]>>, \ +\ %h t. split(flagx(t),%lr p.split(p,%lw lb. \ +\ ccase(h, >, \ +\ >, \ +\ >)))) \ +\ in flagx(l)" + + Flag_def + "Flag(l,x) == ALL lr:List(Colour).ALL lw:List(Colour).ALL lb:List(Colour). \ +\ x = > --> \ +\ (ALL c:Colour.(c mem lr = true --> c=red) & \ +\ (c mem lw = true --> c=white) & \ +\ (c mem lb = true --> c=blue)) & \ +\ Perm(l,lr @ lw @ lb)" + +end + + + diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/List.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/List.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,108 @@ +(* Title: CCL/ex/list + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For list.thy. +*) + +open List; + +val list_defs = [map_def,comp_def,append_def,filter_def,flat_def, + insert_def,isort_def,partition_def,qsort_def]; + +(****) + +val listBs = map (fn s=>prove_goalw List.thy list_defs s (fn _ => [SIMP_TAC term_ss 1])) + ["(f o g) = (%a.f(g(a)))", + "(f o g)(a) = f(g(a))", + "map(f,[]) = []", + "map(f,x.xs) = f(x).map(f,xs)", + "[] @ m = m", + "x.xs @ m = x.(xs @ m)", + "filter(f,[]) = []", + "filter(f,x.xs) = if f`x then x.filter(f,xs) else filter(f,xs)", + "flat([]) = []", + "flat(x.xs) = x @ flat(xs)", + "insert(f,a,[]) = a.[]", + "insert(f,a,x.xs) = if f`a`x then a.x.xs else x.insert(f,a,xs)"]; + +val list_congs = ccl_mk_congs List.thy ["map","op @","filter","flat","insert","napply"]; + +val list_ss = nat_ss addrews listBs addcongs list_congs; + +(****) + +val [prem] = goal List.thy "n:Nat ==> map(f) ^ n ` [] = []"; +br (prem RS Nat_ind) 1; +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val nmapBnil = result(); + +val [prem] = goal List.thy "n:Nat ==> map(f)^n`(x.xs) = f^n`x.map(f)^n`xs"; +br (prem RS Nat_ind) 1; +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val nmapBcons = result(); + +(***) + +val prems = goalw List.thy [map_def] + "[| !!x.x:A==>f(x):B; l : List(A) |] ==> map(f,l) : List(B)"; +by (typechk_tac prems 1); +val mapT = result(); + +val prems = goalw List.thy [append_def] + "[| l : List(A); m : List(A) |] ==> l @ m : List(A)"; +by (typechk_tac prems 1); +val appendT = result(); + +val prems = goal List.thy + "[| l : {l:List(A). m : {m:List(A).P(l @ m)}} |] ==> l @ m : {x:List(A). P(x)}"; +by (cut_facts_tac prems 1); +by (fast_tac (set_cs addSIs [SubtypeI,appendT] addSEs [SubtypeE]) 1); +val appendTS = result(); + +val prems = goalw List.thy [filter_def] + "[| f:A->Bool; l : List(A) |] ==> filter(f,l) : List(A)"; +by (typechk_tac prems 1); +val filterT = result(); + +val prems = goalw List.thy [flat_def] + "l : List(List(A)) ==> flat(l) : List(A)"; +by (typechk_tac (appendT::prems) 1); +val flatT = result(); + +val prems = goalw List.thy [insert_def] + "[| f : A->A->Bool; a:A; l : List(A) |] ==> insert(f,a,l) : List(A)"; +by (typechk_tac prems 1); +val insertT = result(); + +val prems = goal List.thy + "[| f : {f:A->A->Bool. a : {a:A. l : {l:List(A).P(insert(f,a,l))}}} |] ==> \ +\ insert(f,a,l) : {x:List(A). P(x)}"; +by (cut_facts_tac prems 1); +by (fast_tac (set_cs addSIs [SubtypeI,insertT] addSEs [SubtypeE]) 1); +val insertTS = result(); + +val prems = goalw List.thy [partition_def] + "[| f:A->Bool; l : List(A) |] ==> partition(f,l) : List(A)*List(A)"; +by (typechk_tac prems 1); +by clean_ccs_tac; +br (ListPRI RS wfstI RS (ListPR_wf RS wmap_wf RS wfI)) 2; +br (ListPRI RS wfstI RS (ListPR_wf RS wmap_wf RS wfI)) 1; +by (REPEAT (atac 1)); +val partitionT = result(); + +(*** Correctness Conditions for Insertion Sort ***) + + +val prems = goalw List.thy [isort_def] + "f:A->A->Bool ==> isort(f) : PROD l:List(A).{x: List(A). Ord(f,x) & Perm(x,l)}"; +by (gen_ccs_tac ([insertTS,insertT]@prems) 1); + + +(*** Correctness Conditions for Quick Sort ***) + +val prems = goalw List.thy [qsort_def] + "f:A->A->Bool ==> qsort(f) : PROD l:List(A).{x: List(A). Ord(f,x) & Perm(x,l)}"; +by (gen_ccs_tac ([partitionT,appendTS,appendT]@prems) 1); + diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/List.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/List.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,44 @@ +(* Title: CCL/ex/list.thy + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Programs defined over lists. +*) + +List = Nat + + +consts + map :: "[i=>i,i]=>i" + "o" :: "[i=>i,i=>i]=>i=>i" (infixr 55) + "@" :: "[i,i]=>i" (infixr 55) + mem :: "[i,i]=>i" (infixr 55) + filter :: "[i,i]=>i" + flat :: "i=>i" + partition :: "[i,i]=>i" + insert :: "[i,i,i]=>i" + isort :: "i=>i" + qsort :: "i=>i" + +rules + + map_def "map(f,l) == lrec(l,[],%x xs g.f(x).g)" + comp_def "f o g == (%x.f(g(x)))" + append_def "l @ m == lrec(l,m,%x xs g.x.g)" + mem_def "a mem l == lrec(l,false,%h t g.if eq(a,h) then true else g)" + filter_def "filter(f,l) == lrec(l,[],%x xs g.if f`x then x.g else g)" + flat_def "flat(l) == lrec(l,[],%h t g.h @ g)" + + insert_def "insert(f,a,l) == lrec(l,a.[],%h t g.if f`a`h then a.h.t else h.g)" + isort_def "isort(f) == lam l.lrec(l,[],%h t g.insert(f,h,g))" + + partition_def + "partition(f,l) == letrec part l a b be lcase(l,,%x xs.\ +\ if f`x then part(xs,x.a,b) else part(xs,a,x.b)) \ +\ in part(l,[],[])" + qsort_def "qsort(f) == lam l. letrec qsortx l be lcase(l,[],%h t. \ +\ let p be partition(f`h,t) \ +\ in split(p,%x y.qsortx(x) @ h.qsortx(y))) \ +\ in qsortx(l)" + +end diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/Nat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/Nat.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,75 @@ +(* Title: CCL/ex/nat + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For nat.thy. +*) + +open Nat; + +val nat_defs = [not_def,add_def,mult_def,sub_def,le_def,lt_def,ack_def,napply_def]; + +val natBs = map (fn s=>prove_goalw Nat.thy nat_defs s (fn _ => [SIMP_TAC term_ss 1])) + ["not(true) = false", + "not(false) = true", + "zero #+ n = n", + "succ(n) #+ m = succ(n #+ m)", + "zero #* n = zero", + "succ(n) #* m = m #+ (n #* m)", + "f^zero`a = a", + "f^succ(n)`a = f(f^n`a)"]; + +val nat_congs = ccl_mk_congs Nat.thy ["not","op #+","op #*","op #-","op ##", + "op #<","op #<=","ackermann","napply"]; + +val nat_ss = term_ss addrews natBs addcongs nat_congs; + +(*** Lemma for napply ***) + +val [prem] = goal Nat.thy "n:Nat ==> f^n`f(a) = f^succ(n)`a"; +br (prem RS Nat_ind) 1; +by (ALLGOALS (ASM_SIMP_TAC (nat_ss addcongs [read_instantiate [("f","f")] arg_cong]))); +val napply_f = result(); + +(****) + +val prems = goalw Nat.thy [add_def] "[| a:Nat; b:Nat |] ==> a #+ b : Nat"; +by (typechk_tac prems 1); +val addT = result(); + +val prems = goalw Nat.thy [mult_def] "[| a:Nat; b:Nat |] ==> a #* b : Nat"; +by (typechk_tac (addT::prems) 1); +val multT = result(); + +(* Defined to return zero if a a #- b : Nat"; +by (typechk_tac (prems) 1); +by clean_ccs_tac; +be (NatPRI RS wfstI RS (NatPR_wf RS wmap_wf RS wfI)) 1; +val subT = result(); + +val prems = goalw Nat.thy [le_def] "[| a:Nat; b:Nat |] ==> a #<= b : Bool"; +by (typechk_tac (prems) 1); +by clean_ccs_tac; +be (NatPRI RS wfstI RS (NatPR_wf RS wmap_wf RS wfI)) 1; +val leT = result(); + +val prems = goalw Nat.thy [not_def,lt_def] "[| a:Nat; b:Nat |] ==> a #< b : Bool"; +by (typechk_tac (prems@[leT]) 1); +val ltT = result(); + +(* Correctness conditions for subtractive division **) + +val prems = goalw Nat.thy [div_def] + "[| a:Nat; b:{x:Nat.~x=zero} |] ==> a ## b : {x:Nat. DIV(a,b,x)}"; +by (gen_ccs_tac (prems@[ltT,subT]) 1); + +(* Termination Conditions for Ackermann's Function *) + +val prems = goalw Nat.thy [ack_def] + "[| a:Nat; b:Nat |] ==> ackermann(a,b) : Nat"; +by (gen_ccs_tac prems 1); +val relI = NatPR_wf RS (NatPR_wf RS lex_wf RS wfI); +by (REPEAT (eresolve_tac [NatPRI RS (lexI1 RS relI),NatPRI RS (lexI2 RS relI)] 1)); +result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/Nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/Nat.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,38 @@ +(* Title: CCL/ex/nat.thy + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Programs defined over the natural numbers +*) + +Nat = Wfd + + +consts + + not :: "i=>i" + "#+","#*","#-", + "##","#<","#<=" :: "[i,i]=>i" (infixr 60) + ackermann :: "[i,i]=>i" + +rules + + not_def "not(b) == if b then false else true" + + add_def "a #+ b == nrec(a,b,%x g.succ(g))" + mult_def "a #* b == nrec(a,zero,%x g.b #+ g)" + sub_def "a #- b == letrec sub x y be ncase(y,x,%yy.ncase(x,zero,%xx.sub(xx,yy))) \ +\ in sub(a,b)" + le_def "a #<= b == letrec le x y be ncase(x,true,%xx.ncase(y,false,%yy.le(xx,yy))) \ +\ in le(a,b)" + lt_def "a #< b == not(b #<= a)" + + div_def "a ## b == letrec div x y be if x #< y then zero else succ(div(x#-y,y)) \ +\ in div(a,b)" + ack_def + "ackermann(a,b) == letrec ack n m be ncase(n,succ(m),%x. \ +\ ncase(m,ack(x,succ(zero)),%y.ack(x,ack(succ(x),y))))\ +\ in ack(a,b)" + +end + diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,17 @@ +(* Title: CCL/ex/ROOT + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Executes all examples for Classical Computational Logic +*) + +CCL_build_completed; (*Cause examples to fail if CCL did*) + +writeln"Root file for CCL examples"; +proof_timing := true; +time_use_thy "ex/nat"; +time_use_thy "ex/list"; +time_use_thy "ex/stream"; +time_use_thy "ex/flag"; +maketest"END: Root file for CCL examples"; diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/Stream.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/Stream.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,112 @@ +(* Title: CCL/ex/stream + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For stream.thy. + +Proving properties about infinite lists using coinduction: + Lists(A) is the set of all finite and infinite lists of elements of A. + ILists(A) is the set of infinite lists of elements of A. +*) + +open Stream; + +(*** Map of composition is composition of maps ***) + +val prems = goal Stream.thy "l:Lists(A) ==> map(f o g,l) = map(f,map(g,l))"; +by (eq_coinduct3_tac + "{p. EX x y.p= & (EX l:Lists(A).x=map(f o g,l) & y=map(f,map(g,l)))}" 1); +by (fast_tac (ccl_cs addSIs prems) 1); +by (safe_tac type_cs); +be (XH_to_E ListsXH) 1; +by (EQgen_tac list_ss [] 1); +by (SIMP_TAC list_ss 1); +by (fast_tac ccl_cs 1); +val map_comp = result(); + +(*** Mapping the identity function leaves a list unchanged ***) + +val prems = goal Stream.thy "l:Lists(A) ==> map(%x.x,l) = l"; +by (eq_coinduct3_tac + "{p. EX x y.p= & (EX l:Lists(A).x=map(%x.x,l) & y=l)}" 1); +by (fast_tac (ccl_cs addSIs prems) 1); +by (safe_tac type_cs); +be (XH_to_E ListsXH) 1; +by (EQgen_tac list_ss [] 1); +by (fast_tac ccl_cs 1); +val map_id = result(); + +(*** Mapping distributes over append ***) + +val prems = goal Stream.thy + "[| l:Lists(A); m:Lists(A) |] ==> map(f,l@m) = map(f,l) @ map(f,m)"; +by (eq_coinduct3_tac "{p. EX x y.p= & (EX l:Lists(A).EX m:Lists(A). \ +\ x=map(f,l@m) & y=map(f,l) @ map(f,m))}" 1); +by (fast_tac (ccl_cs addSIs prems) 1); +by (safe_tac type_cs); +be (XH_to_E ListsXH) 1; +by (EQgen_tac list_ss [] 1); +be (XH_to_E ListsXH) 1; +by (EQgen_tac list_ss [] 1); +by (fast_tac ccl_cs 1); +val map_append = result(); + +(*** Append is associative ***) + +val prems = goal Stream.thy + "[| k:Lists(A); l:Lists(A); m:Lists(A) |] ==> k @ l @ m = (k @ l) @ m"; +by (eq_coinduct3_tac "{p. EX x y.p= & (EX k:Lists(A).EX l:Lists(A).EX m:Lists(A). \ +\ x=k @ l @ m & y=(k @ l) @ m)}" 1); +by (fast_tac (ccl_cs addSIs prems) 1); +by (safe_tac type_cs); +be (XH_to_E ListsXH) 1; +by (EQgen_tac list_ss [] 1); +be (XH_to_E ListsXH) 1;back(); +by (EQgen_tac list_ss [] 1); +be (XH_to_E ListsXH) 1; +by (EQgen_tac list_ss [] 1); +by (fast_tac ccl_cs 1); +val append_assoc = result(); + +(*** Appending anything to an infinite list doesn't alter it ****) + +val prems = goal Stream.thy "l:ILists(A) ==> l @ m = l"; +by (eq_coinduct3_tac "{p. EX x y.p= & (EX l:ILists(A).EX m.x=l@m & y=l)}" 1); +by (fast_tac (ccl_cs addSIs prems) 1); +by (safe_tac set_cs); +be (XH_to_E IListsXH) 1; +by (EQgen_tac list_ss [] 1); +by (fast_tac ccl_cs 1); +val ilist_append = result(); + +(*** The equivalance of two versions of an iteration function ***) +(* *) +(* fun iter1(f,a) = a.iter1(f,f(a)) *) +(* fun iter2(f,a) = a.map(f,iter2(f,a)) *) + +goalw Stream.thy [iter1_def] "iter1(f,a) = a.iter1(f,f(a))"; +br (letrecB RS trans) 1; +by (SIMP_TAC term_ss 1); +val iter1B = result(); + +goalw Stream.thy [iter2_def] "iter2(f,a) = a . map(f,iter2(f,a))"; +br (letrecB RS trans) 1; +br refl 1; +val iter2B = result(); + +val [prem] =goal Stream.thy + "n:Nat ==> map(f) ^ n ` iter2(f,a) = f ^ n ` a . map(f) ^ n ` map(f,iter2(f,a))"; +br (iter2B RS ssubst) 1;back();back(); +by (SIMP_TAC (list_ss addrews [prem RS nmapBcons]) 1); +val iter2Blemma = result(); + +goal Stream.thy "iter1(f,a) = iter2(f,a)"; +by (eq_coinduct3_tac + "{p. EX x y.p= & (EX n:Nat.x=iter1(f,f^n`a) & y=map(f)^n`iter2(f,a))}" 1); +by (fast_tac (type_cs addSIs [napplyBzero RS sym,napplyBzero RS sym RS arg_cong]) 1); +by (EQgen_tac list_ss [iter1B,iter2Blemma] 1); +by (rtac (napply_f RS ssubst) 1 THEN atac 1); +by (res_inst_tac [("f1","f")] (napplyBsucc RS subst) 1); +by (fast_tac type_cs 1); +val iter1_iter2_eq = result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/Stream.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/Stream.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,20 @@ +(* Title: CCL/ex/stream.thy + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Programs defined over streams. +*) + +Stream = List + + +consts + + iter1,iter2 :: "[i=>i,i]=>i" + +rules + + iter1_def "iter1(f,a) == letrec iter x be x.iter(f(x)) in iter(a)" + iter2_def "iter2(f,a) == letrec iter x be x.map(f,iter(x)) in iter(a)" + +end diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/flag.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/flag.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,46 @@ +(* Title: CCL/ex/flag + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For flag.thy. +*) + +open Flag; + +(******) + +val flag_defs = [Colour_def,red_def,white_def,blue_def,ccase_def]; + +(******) + +val ColourXH = mk_XH_tac Flag.thy (simp_type_defs @flag_defs) [] + "a : Colour <-> (a=red | a=white | a=blue)"; + +val Colour_case = XH_to_E ColourXH; + +val redT = mk_canT_tac Flag.thy [ColourXH] "red : Colour"; +val whiteT = mk_canT_tac Flag.thy [ColourXH] "white : Colour"; +val blueT = mk_canT_tac Flag.thy [ColourXH] "blue : Colour"; + + +val ccaseT = mk_ncanT_tac Flag.thy flag_defs case_rls case_rls + "[| c:Colour; \ +\ c=red ==> r : C(red); c=white ==> w : C(white); c=blue ==> b : C(blue) |] ==> \ +\ ccase(c,r,w,b) : C(c)"; + +(***) + +val prems = goalw Flag.thy [flag_def] + "flag : List(Colour)->List(Colour)*List(Colour)*List(Colour)"; +by (typechk_tac [redT,whiteT,blueT,ccaseT] 1); +by clean_ccs_tac; +be (ListPRI RS (ListPR_wf RS wfI)) 1; +ba 1; +result(); + + +val prems = goalw Flag.thy [flag_def] + "flag : PROD l:List(Colour).{x:List(Colour)*List(Colour)*List(Colour).FLAG(x,l)}"; +by (gen_ccs_tac [redT,whiteT,blueT,ccaseT] 1); +by (REPEAT_SOME (ares_tac [ListPRI RS (ListPR_wf RS wfI)])); diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/flag.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/flag.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,48 @@ +(* Title: CCL/ex/flag.thy + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Dutch national flag program - except that the point of Dijkstra's example was to use +arrays and this uses lists. + +*) + +Flag = List + + +consts + + Colour :: "i set" + red, white, blue :: "i" + ccase :: "[i,i,i,i]=>i" + flag :: "i" + +rules + + Colour_def "Colour == Unit + Unit + Unit" + red_def "red == inl(one)" + white_def "white == inr(inl(one))" + blue_def "blue == inr(inr(one))" + + ccase_def "ccase(c,r,w,b) == when(c,%x.r,%wb.when(wb,%x.w,%x.b))" + + flag_def "flag == lam l.letrec \ +\ flagx l be lcase(l,<[],<[],[]>>, \ +\ %h t. split(flagx(t),%lr p.split(p,%lw lb. \ +\ ccase(h, >, \ +\ >, \ +\ >)))) \ +\ in flagx(l)" + + Flag_def + "Flag(l,x) == ALL lr:List(Colour).ALL lw:List(Colour).ALL lb:List(Colour). \ +\ x = > --> \ +\ (ALL c:Colour.(c mem lr = true --> c=red) & \ +\ (c mem lw = true --> c=white) & \ +\ (c mem lb = true --> c=blue)) & \ +\ Perm(l,lr @ lw @ lb)" + +end + + + diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/list.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/list.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,108 @@ +(* Title: CCL/ex/list + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For list.thy. +*) + +open List; + +val list_defs = [map_def,comp_def,append_def,filter_def,flat_def, + insert_def,isort_def,partition_def,qsort_def]; + +(****) + +val listBs = map (fn s=>prove_goalw List.thy list_defs s (fn _ => [SIMP_TAC term_ss 1])) + ["(f o g) = (%a.f(g(a)))", + "(f o g)(a) = f(g(a))", + "map(f,[]) = []", + "map(f,x.xs) = f(x).map(f,xs)", + "[] @ m = m", + "x.xs @ m = x.(xs @ m)", + "filter(f,[]) = []", + "filter(f,x.xs) = if f`x then x.filter(f,xs) else filter(f,xs)", + "flat([]) = []", + "flat(x.xs) = x @ flat(xs)", + "insert(f,a,[]) = a.[]", + "insert(f,a,x.xs) = if f`a`x then a.x.xs else x.insert(f,a,xs)"]; + +val list_congs = ccl_mk_congs List.thy ["map","op @","filter","flat","insert","napply"]; + +val list_ss = nat_ss addrews listBs addcongs list_congs; + +(****) + +val [prem] = goal List.thy "n:Nat ==> map(f) ^ n ` [] = []"; +br (prem RS Nat_ind) 1; +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val nmapBnil = result(); + +val [prem] = goal List.thy "n:Nat ==> map(f)^n`(x.xs) = f^n`x.map(f)^n`xs"; +br (prem RS Nat_ind) 1; +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val nmapBcons = result(); + +(***) + +val prems = goalw List.thy [map_def] + "[| !!x.x:A==>f(x):B; l : List(A) |] ==> map(f,l) : List(B)"; +by (typechk_tac prems 1); +val mapT = result(); + +val prems = goalw List.thy [append_def] + "[| l : List(A); m : List(A) |] ==> l @ m : List(A)"; +by (typechk_tac prems 1); +val appendT = result(); + +val prems = goal List.thy + "[| l : {l:List(A). m : {m:List(A).P(l @ m)}} |] ==> l @ m : {x:List(A). P(x)}"; +by (cut_facts_tac prems 1); +by (fast_tac (set_cs addSIs [SubtypeI,appendT] addSEs [SubtypeE]) 1); +val appendTS = result(); + +val prems = goalw List.thy [filter_def] + "[| f:A->Bool; l : List(A) |] ==> filter(f,l) : List(A)"; +by (typechk_tac prems 1); +val filterT = result(); + +val prems = goalw List.thy [flat_def] + "l : List(List(A)) ==> flat(l) : List(A)"; +by (typechk_tac (appendT::prems) 1); +val flatT = result(); + +val prems = goalw List.thy [insert_def] + "[| f : A->A->Bool; a:A; l : List(A) |] ==> insert(f,a,l) : List(A)"; +by (typechk_tac prems 1); +val insertT = result(); + +val prems = goal List.thy + "[| f : {f:A->A->Bool. a : {a:A. l : {l:List(A).P(insert(f,a,l))}}} |] ==> \ +\ insert(f,a,l) : {x:List(A). P(x)}"; +by (cut_facts_tac prems 1); +by (fast_tac (set_cs addSIs [SubtypeI,insertT] addSEs [SubtypeE]) 1); +val insertTS = result(); + +val prems = goalw List.thy [partition_def] + "[| f:A->Bool; l : List(A) |] ==> partition(f,l) : List(A)*List(A)"; +by (typechk_tac prems 1); +by clean_ccs_tac; +br (ListPRI RS wfstI RS (ListPR_wf RS wmap_wf RS wfI)) 2; +br (ListPRI RS wfstI RS (ListPR_wf RS wmap_wf RS wfI)) 1; +by (REPEAT (atac 1)); +val partitionT = result(); + +(*** Correctness Conditions for Insertion Sort ***) + + +val prems = goalw List.thy [isort_def] + "f:A->A->Bool ==> isort(f) : PROD l:List(A).{x: List(A). Ord(f,x) & Perm(x,l)}"; +by (gen_ccs_tac ([insertTS,insertT]@prems) 1); + + +(*** Correctness Conditions for Quick Sort ***) + +val prems = goalw List.thy [qsort_def] + "f:A->A->Bool ==> qsort(f) : PROD l:List(A).{x: List(A). Ord(f,x) & Perm(x,l)}"; +by (gen_ccs_tac ([partitionT,appendTS,appendT]@prems) 1); + diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/list.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/list.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,44 @@ +(* Title: CCL/ex/list.thy + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Programs defined over lists. +*) + +List = Nat + + +consts + map :: "[i=>i,i]=>i" + "o" :: "[i=>i,i=>i]=>i=>i" (infixr 55) + "@" :: "[i,i]=>i" (infixr 55) + mem :: "[i,i]=>i" (infixr 55) + filter :: "[i,i]=>i" + flat :: "i=>i" + partition :: "[i,i]=>i" + insert :: "[i,i,i]=>i" + isort :: "i=>i" + qsort :: "i=>i" + +rules + + map_def "map(f,l) == lrec(l,[],%x xs g.f(x).g)" + comp_def "f o g == (%x.f(g(x)))" + append_def "l @ m == lrec(l,m,%x xs g.x.g)" + mem_def "a mem l == lrec(l,false,%h t g.if eq(a,h) then true else g)" + filter_def "filter(f,l) == lrec(l,[],%x xs g.if f`x then x.g else g)" + flat_def "flat(l) == lrec(l,[],%h t g.h @ g)" + + insert_def "insert(f,a,l) == lrec(l,a.[],%h t g.if f`a`h then a.h.t else h.g)" + isort_def "isort(f) == lam l.lrec(l,[],%h t g.insert(f,h,g))" + + partition_def + "partition(f,l) == letrec part l a b be lcase(l,,%x xs.\ +\ if f`x then part(xs,x.a,b) else part(xs,a,x.b)) \ +\ in part(l,[],[])" + qsort_def "qsort(f) == lam l. letrec qsortx l be lcase(l,[],%h t. \ +\ let p be partition(f`h,t) \ +\ in split(p,%x y.qsortx(x) @ h.qsortx(y))) \ +\ in qsortx(l)" + +end diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/nat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/nat.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,75 @@ +(* Title: CCL/ex/nat + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For nat.thy. +*) + +open Nat; + +val nat_defs = [not_def,add_def,mult_def,sub_def,le_def,lt_def,ack_def,napply_def]; + +val natBs = map (fn s=>prove_goalw Nat.thy nat_defs s (fn _ => [SIMP_TAC term_ss 1])) + ["not(true) = false", + "not(false) = true", + "zero #+ n = n", + "succ(n) #+ m = succ(n #+ m)", + "zero #* n = zero", + "succ(n) #* m = m #+ (n #* m)", + "f^zero`a = a", + "f^succ(n)`a = f(f^n`a)"]; + +val nat_congs = ccl_mk_congs Nat.thy ["not","op #+","op #*","op #-","op ##", + "op #<","op #<=","ackermann","napply"]; + +val nat_ss = term_ss addrews natBs addcongs nat_congs; + +(*** Lemma for napply ***) + +val [prem] = goal Nat.thy "n:Nat ==> f^n`f(a) = f^succ(n)`a"; +br (prem RS Nat_ind) 1; +by (ALLGOALS (ASM_SIMP_TAC (nat_ss addcongs [read_instantiate [("f","f")] arg_cong]))); +val napply_f = result(); + +(****) + +val prems = goalw Nat.thy [add_def] "[| a:Nat; b:Nat |] ==> a #+ b : Nat"; +by (typechk_tac prems 1); +val addT = result(); + +val prems = goalw Nat.thy [mult_def] "[| a:Nat; b:Nat |] ==> a #* b : Nat"; +by (typechk_tac (addT::prems) 1); +val multT = result(); + +(* Defined to return zero if a a #- b : Nat"; +by (typechk_tac (prems) 1); +by clean_ccs_tac; +be (NatPRI RS wfstI RS (NatPR_wf RS wmap_wf RS wfI)) 1; +val subT = result(); + +val prems = goalw Nat.thy [le_def] "[| a:Nat; b:Nat |] ==> a #<= b : Bool"; +by (typechk_tac (prems) 1); +by clean_ccs_tac; +be (NatPRI RS wfstI RS (NatPR_wf RS wmap_wf RS wfI)) 1; +val leT = result(); + +val prems = goalw Nat.thy [not_def,lt_def] "[| a:Nat; b:Nat |] ==> a #< b : Bool"; +by (typechk_tac (prems@[leT]) 1); +val ltT = result(); + +(* Correctness conditions for subtractive division **) + +val prems = goalw Nat.thy [div_def] + "[| a:Nat; b:{x:Nat.~x=zero} |] ==> a ## b : {x:Nat. DIV(a,b,x)}"; +by (gen_ccs_tac (prems@[ltT,subT]) 1); + +(* Termination Conditions for Ackermann's Function *) + +val prems = goalw Nat.thy [ack_def] + "[| a:Nat; b:Nat |] ==> ackermann(a,b) : Nat"; +by (gen_ccs_tac prems 1); +val relI = NatPR_wf RS (NatPR_wf RS lex_wf RS wfI); +by (REPEAT (eresolve_tac [NatPRI RS (lexI1 RS relI),NatPRI RS (lexI2 RS relI)] 1)); +result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/nat.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,38 @@ +(* Title: CCL/ex/nat.thy + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Programs defined over the natural numbers +*) + +Nat = Wfd + + +consts + + not :: "i=>i" + "#+","#*","#-", + "##","#<","#<=" :: "[i,i]=>i" (infixr 60) + ackermann :: "[i,i]=>i" + +rules + + not_def "not(b) == if b then false else true" + + add_def "a #+ b == nrec(a,b,%x g.succ(g))" + mult_def "a #* b == nrec(a,zero,%x g.b #+ g)" + sub_def "a #- b == letrec sub x y be ncase(y,x,%yy.ncase(x,zero,%xx.sub(xx,yy))) \ +\ in sub(a,b)" + le_def "a #<= b == letrec le x y be ncase(x,true,%xx.ncase(y,false,%yy.le(xx,yy))) \ +\ in le(a,b)" + lt_def "a #< b == not(b #<= a)" + + div_def "a ## b == letrec div x y be if x #< y then zero else succ(div(x#-y,y)) \ +\ in div(a,b)" + ack_def + "ackermann(a,b) == letrec ack n m be ncase(n,succ(m),%x. \ +\ ncase(m,ack(x,succ(zero)),%y.ack(x,ack(succ(x),y))))\ +\ in ack(a,b)" + +end + diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/stream.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/stream.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,112 @@ +(* Title: CCL/ex/stream + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For stream.thy. + +Proving properties about infinite lists using coinduction: + Lists(A) is the set of all finite and infinite lists of elements of A. + ILists(A) is the set of infinite lists of elements of A. +*) + +open Stream; + +(*** Map of composition is composition of maps ***) + +val prems = goal Stream.thy "l:Lists(A) ==> map(f o g,l) = map(f,map(g,l))"; +by (eq_coinduct3_tac + "{p. EX x y.p= & (EX l:Lists(A).x=map(f o g,l) & y=map(f,map(g,l)))}" 1); +by (fast_tac (ccl_cs addSIs prems) 1); +by (safe_tac type_cs); +be (XH_to_E ListsXH) 1; +by (EQgen_tac list_ss [] 1); +by (SIMP_TAC list_ss 1); +by (fast_tac ccl_cs 1); +val map_comp = result(); + +(*** Mapping the identity function leaves a list unchanged ***) + +val prems = goal Stream.thy "l:Lists(A) ==> map(%x.x,l) = l"; +by (eq_coinduct3_tac + "{p. EX x y.p= & (EX l:Lists(A).x=map(%x.x,l) & y=l)}" 1); +by (fast_tac (ccl_cs addSIs prems) 1); +by (safe_tac type_cs); +be (XH_to_E ListsXH) 1; +by (EQgen_tac list_ss [] 1); +by (fast_tac ccl_cs 1); +val map_id = result(); + +(*** Mapping distributes over append ***) + +val prems = goal Stream.thy + "[| l:Lists(A); m:Lists(A) |] ==> map(f,l@m) = map(f,l) @ map(f,m)"; +by (eq_coinduct3_tac "{p. EX x y.p= & (EX l:Lists(A).EX m:Lists(A). \ +\ x=map(f,l@m) & y=map(f,l) @ map(f,m))}" 1); +by (fast_tac (ccl_cs addSIs prems) 1); +by (safe_tac type_cs); +be (XH_to_E ListsXH) 1; +by (EQgen_tac list_ss [] 1); +be (XH_to_E ListsXH) 1; +by (EQgen_tac list_ss [] 1); +by (fast_tac ccl_cs 1); +val map_append = result(); + +(*** Append is associative ***) + +val prems = goal Stream.thy + "[| k:Lists(A); l:Lists(A); m:Lists(A) |] ==> k @ l @ m = (k @ l) @ m"; +by (eq_coinduct3_tac "{p. EX x y.p= & (EX k:Lists(A).EX l:Lists(A).EX m:Lists(A). \ +\ x=k @ l @ m & y=(k @ l) @ m)}" 1); +by (fast_tac (ccl_cs addSIs prems) 1); +by (safe_tac type_cs); +be (XH_to_E ListsXH) 1; +by (EQgen_tac list_ss [] 1); +be (XH_to_E ListsXH) 1;back(); +by (EQgen_tac list_ss [] 1); +be (XH_to_E ListsXH) 1; +by (EQgen_tac list_ss [] 1); +by (fast_tac ccl_cs 1); +val append_assoc = result(); + +(*** Appending anything to an infinite list doesn't alter it ****) + +val prems = goal Stream.thy "l:ILists(A) ==> l @ m = l"; +by (eq_coinduct3_tac "{p. EX x y.p= & (EX l:ILists(A).EX m.x=l@m & y=l)}" 1); +by (fast_tac (ccl_cs addSIs prems) 1); +by (safe_tac set_cs); +be (XH_to_E IListsXH) 1; +by (EQgen_tac list_ss [] 1); +by (fast_tac ccl_cs 1); +val ilist_append = result(); + +(*** The equivalance of two versions of an iteration function ***) +(* *) +(* fun iter1(f,a) = a.iter1(f,f(a)) *) +(* fun iter2(f,a) = a.map(f,iter2(f,a)) *) + +goalw Stream.thy [iter1_def] "iter1(f,a) = a.iter1(f,f(a))"; +br (letrecB RS trans) 1; +by (SIMP_TAC term_ss 1); +val iter1B = result(); + +goalw Stream.thy [iter2_def] "iter2(f,a) = a . map(f,iter2(f,a))"; +br (letrecB RS trans) 1; +br refl 1; +val iter2B = result(); + +val [prem] =goal Stream.thy + "n:Nat ==> map(f) ^ n ` iter2(f,a) = f ^ n ` a . map(f) ^ n ` map(f,iter2(f,a))"; +br (iter2B RS ssubst) 1;back();back(); +by (SIMP_TAC (list_ss addrews [prem RS nmapBcons]) 1); +val iter2Blemma = result(); + +goal Stream.thy "iter1(f,a) = iter2(f,a)"; +by (eq_coinduct3_tac + "{p. EX x y.p= & (EX n:Nat.x=iter1(f,f^n`a) & y=map(f)^n`iter2(f,a))}" 1); +by (fast_tac (type_cs addSIs [napplyBzero RS sym,napplyBzero RS sym RS arg_cong]) 1); +by (EQgen_tac list_ss [iter1B,iter2Blemma] 1); +by (rtac (napply_f RS ssubst) 1 THEN atac 1); +by (res_inst_tac [("f1","f")] (napplyBsucc RS subst) 1); +by (fast_tac type_cs 1); +val iter1_iter2_eq = result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/ex/stream.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/ex/stream.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,20 @@ +(* Title: CCL/ex/stream.thy + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Programs defined over streams. +*) + +Stream = List + + +consts + + iter1,iter2 :: "[i=>i,i]=>i" + +rules + + iter1_def "iter1(f,a) == letrec iter x be x.iter(f(x)) in iter(a)" + iter2_def "iter2(f,a) == letrec iter x be x.map(f,iter(x)) in iter(a)" + +end diff -r 000000000000 -r a5a9c433f639 src/CCL/fix.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/fix.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,202 @@ +(* Title: CCL/fix + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For fix.thy. +*) + +open Fix; + +val prems = goalw Fix.thy [INCL_def] + "[| !!x.P(x) <-> Q(x) |] ==> INCL(%x.P(x)) <-> INCL(%x.Q(x))"; +by (REPEAT (ares_tac ([refl] @ FOL_congs @ set_congs @ prems) 1)); +val INCL_cong = result(); + +val fix_congs = [INCL_cong] @ ccl_mk_congs Fix.thy ["napply"]; + +(*** Fixed Point Induction ***) + +val [base,step,incl] = goalw Fix.thy [INCL_def] + "[| P(bot); !!x.P(x) ==> P(f(x)); INCL(P) |] ==> P(fix(f))"; +br (incl RS spec RS mp) 1; +by (rtac (Nat_ind RS ballI) 1 THEN atac 1); +by (ALLGOALS (SIMP_TAC term_ss)); +by (REPEAT (ares_tac [base,step] 1)); +val fix_ind = result(); + +(*** Inclusive Predicates ***) + +val prems = goalw Fix.thy [INCL_def] + "INCL(P) <-> (ALL f. (ALL n:Nat. P(f ^ n ` bot)) --> P(fix(f)))"; +br iff_refl 1; +val inclXH = result(); + +val prems = goal Fix.thy + "[| !!f.ALL n:Nat.P(f^n`bot) ==> P(fix(f)) |] ==> INCL(%x.P(x))"; +by (fast_tac (term_cs addIs (prems @ [XH_to_I inclXH])) 1); +val inclI = result(); + +val incl::prems = goal Fix.thy + "[| INCL(P); !!n.n:Nat ==> P(f^n`bot) |] ==> P(fix(f))"; +by (fast_tac (term_cs addIs ([ballI RS (incl RS (XH_to_D inclXH) RS spec RS mp)] + @ prems)) 1); +val inclD = result(); + +val incl::prems = goal Fix.thy + "[| INCL(P); (ALL n:Nat.P(f^n`bot))-->P(fix(f)) ==> R |] ==> R"; +by (fast_tac (term_cs addIs ([incl RS inclD] @ prems)) 1); +val inclE = result(); + +val fix_ss = term_ss addcongs fix_congs; + +(*** Lemmas for Inclusive Predicates ***) + +goal Fix.thy "INCL(%x.~ a(x) [= t)"; +br inclI 1; +bd bspec 1; +br zeroT 1; +be contrapos 1; +br po_trans 1; +ba 2; +br (napplyBzero RS ssubst) 1; +by (rtac po_cong 1 THEN rtac po_bot 1); +val npo_INCL = result(); + +val prems = goal Fix.thy "[| INCL(P); INCL(Q) |] ==> INCL(%x.P(x) & Q(x))"; +by (fast_tac (set_cs addSIs ([inclI] @ (prems RL [inclD]))) 1);; +val conj_INCL = result(); + +val prems = goal Fix.thy "[| !!a.INCL(P(a)) |] ==> INCL(%x.ALL a.P(a,x))"; +by (fast_tac (set_cs addSIs ([inclI] @ (prems RL [inclD]))) 1);; +val all_INCL = result(); + +val prems = goal Fix.thy "[| !!a.a:A ==> INCL(P(a)) |] ==> INCL(%x.ALL a:A.P(a,x))"; +by (fast_tac (set_cs addSIs ([inclI] @ (prems RL [inclD]))) 1);; +val ball_INCL = result(); + +goal Fix.thy "INCL(%x.a(x) = b(x)::'a::prog)"; +by (SIMP_TAC (fix_ss addrews [eq_iff]) 1); +by (REPEAT (resolve_tac [conj_INCL,po_INCL] 1)); +val eq_INCL = result(); + +(*** Derivation of Reachability Condition ***) + +(* Fixed points of idgen *) + +goal Fix.thy "idgen(fix(idgen)) = fix(idgen)"; +br (fixB RS sym) 1; +val fix_idgenfp = result(); + +goalw Fix.thy [idgen_def] "idgen(lam x.x) = lam x.x"; +by (SIMP_TAC term_ss 1); +br (term_case RS allI) 1; +by (ALLGOALS (SIMP_TAC term_ss)); +val id_idgenfp = result(); + +(* All fixed points are lam-expressions *) + +val [prem] = goal Fix.thy "idgen(d) = d ==> d = lam x.?f(x)"; +br (prem RS subst) 1; +bw idgen_def; +br refl 1; +val idgenfp_lam = result(); + +(* Lemmas for rewriting fixed points of idgen *) + +val prems = goalw Fix.thy [idgen_def] + "[| a = b; a ` t = u |] ==> b ` t = u"; +by (SIMP_TAC (term_ss addrews (prems RL [sym])) 1); +val l_lemma= result(); + +val idgen_lemmas = + let fun mk_thm s = prove_goalw Fix.thy [idgen_def] s + (fn [prem] => [rtac (prem RS l_lemma) 1,SIMP_TAC term_ss 1]) + in map mk_thm + [ "idgen(d) = d ==> d ` bot = bot", + "idgen(d) = d ==> d ` true = true", + "idgen(d) = d ==> d ` false = false", + "idgen(d) = d ==> d ` = ", + "idgen(d) = d ==> d ` (lam x.f(x)) = lam x.d ` f(x)"] + end; + +(* Proof of Reachability law - show that fix and lam x.x both give LEAST fixed points + of idgen and hence are they same *) + +val [p1,p2,p3] = goal CCL.thy + "[| ALL x.t ` x [= u ` x; EX f.t=lam x.f(x); EX f.u=lam x.f(x) |] ==> t [= u"; +br (p2 RS cond_eta RS ssubst) 1; +br (p3 RS cond_eta RS ssubst) 1; +br (p1 RS (po_lam RS iffD2)) 1; +val po_eta = result(); + +val [prem] = goalw Fix.thy [idgen_def] "idgen(d) = d ==> d = lam x.?f(x)"; +br (prem RS subst) 1; +br refl 1; +val po_eta_lemma = result(); + +val [prem] = goal Fix.thy + "idgen(d) = d ==> \ +\ {p.EX a b.p= & (EX t.a=fix(idgen) ` t & b = d ` t)} <= \ +\ POgen({p.EX a b.p= & (EX t.a=fix(idgen) ` t & b = d ` t)})"; +by (REPEAT (step_tac term_cs 1)); +by (term_case_tac "t" 1); +by (ALLGOALS (SIMP_TAC (term_ss addrews (POgenXH::([prem,fix_idgenfp] RL idgen_lemmas))))); +by (ALLGOALS (fast_tac set_cs)); +val lemma1 = result(); + +val [prem] = goal Fix.thy + "idgen(d) = d ==> fix(idgen) [= d"; +br (allI RS po_eta) 1; +br (lemma1 RSN(2,po_coinduct)) 1; +by (ALLGOALS (fast_tac (term_cs addIs [prem,po_eta_lemma,fix_idgenfp]))); +val fix_least_idgen = result(); + +val [prem] = goal Fix.thy + "idgen(d) = d ==> \ +\ {p.EX a b.p= & b = d ` a} <= POgen({p.EX a b.p= & b = d ` a})"; +by (REPEAT (step_tac term_cs 1)); +by (term_case_tac "a" 1); +by (ALLGOALS (SIMP_TAC (term_ss addrews (POgenXH::([prem] RL idgen_lemmas))))); +by (ALLGOALS (fast_tac set_cs)); +val lemma2 = result(); + +val [prem] = goal Fix.thy + "idgen(d) = d ==> lam x.x [= d"; +br (allI RS po_eta) 1; +br (lemma2 RSN(2,po_coinduct)) 1; +by (SIMP_TAC term_ss 1); +by (ALLGOALS (fast_tac (term_cs addIs [prem,po_eta_lemma,fix_idgenfp]))); +val id_least_idgen = result(); + +goal Fix.thy "fix(idgen) = lam x.x"; +by (fast_tac (term_cs addIs [eq_iff RS iffD2, + id_idgenfp RS fix_least_idgen, + fix_idgenfp RS id_least_idgen]) 1); +val reachability = result(); + +(********) + +val [prem] = goal Fix.thy "f = lam x.x ==> f`t = t"; +br (prem RS sym RS subst) 1; +br applyB 1; +val id_apply = result(); + +val prems = goal Fix.thy + "[| P(bot); P(true); P(false); \ +\ !!x y.[| P(x); P(y) |] ==> P(); \ +\ !!u.(!!x.P(u(x))) ==> P(lam x.u(x)); INCL(P) |] ==> \ +\ P(t)"; +br (reachability RS id_apply RS subst) 1; +by (res_inst_tac [("x","t")] spec 1); +br fix_ind 1; +bw idgen_def; +by (REPEAT_SOME (ares_tac [allI])); +br (applyBbot RS ssubst) 1; +brs prems 1; +br (applyB RS ssubst )1; +by (res_inst_tac [("t","xa")] term_case 1); +by (ALLGOALS (SIMP_TAC term_ss)); +by (ALLGOALS (fast_tac (term_cs addIs ([all_INCL,INCL_subst] @ prems)))); +val term_ind = result(); + diff -r 000000000000 -r a5a9c433f639 src/CCL/fix.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/fix.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,26 @@ +(* Title: CCL/Lazy/fix.thy + ID: $Id$ + Author: Martin Coen + Copyright 1993 University of Cambridge + +Tentative attempt at including fixed point induction. +Justified by Smith. +*) + +Fix = Type + + +consts + + idgen :: "[i]=>i" + INCL :: "[i=>o]=>o" + +rules + + idgen_def + "idgen(f) == lam t.case(t,true,false,%x y.,%u.lam x.f ` u(x))" + + INCL_def "INCL(%x.P(x)) == (ALL f.(ALL n:Nat.P(f^n`bot)) --> P(fix(f)))" + po_INCL "INCL(%x.a(x) [= b(x))" + INCL_subst "INCL(P) ==> INCL(%x.P((g::i=>i)(x)))" + +end diff -r 000000000000 -r a5a9c433f639 src/CCL/genrec.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/genrec.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,165 @@ +(* Title: 92/CCL/genrec + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +*) + +(*** General Recursive Functions ***) + +val major::prems = goal Wfd.thy + "[| a : A; \ +\ !!p g.[| p:A; ALL x:{x: A. :wf(R)}. g(x) : D(x) |] ==>\ +\ h(p,g) : D(p) |] ==> \ +\ letrec g x be h(x,g) in g(a) : D(a)"; +br (major RS rev_mp) 1; +br (wf_wf RS wfd_induct) 1; +br (letrecB RS ssubst) 1; +br impI 1; +bes prems 1; +br ballI 1; +be (spec RS mp RS mp) 1; +by (REPEAT (eresolve_tac [SubtypeD1,SubtypeD2] 1)); +val letrecT = result(); + +goalw Wfd.thy [SPLIT_def] "SPLIT(,B) = B(a,b)"; +br set_ext 1; +by (fast_tac ccl_cs 1); +val SPLITB = result(); + +val prems = goalw Wfd.thy [letrec2_def] + "[| a : A; b : B; \ +\ !!p q g.[| p:A; q:B; \ +\ ALL x:A.ALL y:{y: B. <,>:wf(R)}. g(x,y) : D(x,y) |] ==>\ +\ h(p,q,g) : D(p,q) |] ==> \ +\ letrec g x y be h(x,y,g) in g(a,b) : D(a,b)"; +br (SPLITB RS subst) 1; +by (REPEAT (ares_tac ([letrecT,pairT,splitT]@prems) 1)); +br (SPLITB RS ssubst) 1; +by (REPEAT (ares_tac ([ballI,SubtypeI]@prems) 1)); +br (SPLITB RS subst) 1; +by (REPEAT (ares_tac ([letrecT,SubtypeI,pairT,splitT]@prems) 1 ORELSE + eresolve_tac [bspec,SubtypeE,sym RS subst] 1)); +val letrec2T = result(); + +goal Wfd.thy "SPLIT(>,%x xs.SPLIT(xs,%y z.B(x,y,z))) = B(a,b,c)"; +by (SIMP_TAC (ccl_ss addrews [SPLITB]) 1); +val lemma = result(); + +val prems = goalw Wfd.thy [letrec3_def] + "[| a : A; b : B; c : C; \ +\ !!p q r g.[| p:A; q:B; r:C; \ +\ ALL x:A.ALL y:B.ALL z:{z:C. <>,>> : wf(R)}. \ +\ g(x,y,z) : D(x,y,z) |] ==>\ +\ h(p,q,r,g) : D(p,q,r) |] ==> \ +\ letrec g x y z be h(x,y,z,g) in g(a,b,c) : D(a,b,c)"; +br (lemma RS subst) 1; +by (REPEAT (ares_tac ([letrecT,pairT,splitT]@prems) 1)); +by (SIMP_TAC (ccl_ss addrews [SPLITB]) 1); +by (REPEAT (ares_tac ([ballI,SubtypeI]@prems) 1)); +br (lemma RS subst) 1; +by (REPEAT (ares_tac ([letrecT,SubtypeI,pairT,splitT]@prems) 1 ORELSE + eresolve_tac [bspec,SubtypeE,sym RS subst] 1)); +val letrec3T = result(); + +val letrecTs = [letrecT,letrec2T,letrec3T]; + + +(*** Type Checking for Recursive Calls ***) + +val major::prems = goal Wfd.thy + "[| ALL x:{x:A.:wf(R)}.g(x):D(x); \ +\ g(a) : D(a) ==> g(a) : E; a:A; :wf(R) |] ==> \ +\ g(a) : E"; +by (REPEAT (ares_tac ([SubtypeI,major RS bspec,major]@prems) 1)); +val rcallT = result(); + +val major::prems = goal Wfd.thy + "[| ALL x:A.ALL y:{y:B.<,>:wf(R)}.g(x,y):D(x,y); \ +\ g(a,b) : D(a,b) ==> g(a,b) : E; a:A; b:B; <,>:wf(R) |] ==> \ +\ g(a,b) : E"; +by (REPEAT (ares_tac ([SubtypeI,major RS bspec RS bspec,major]@prems) 1)); +val rcall2T = result(); + +val major::prems = goal Wfd.thy + "[| ALL x:A.ALL y:B.ALL z:{z:C.<>,>>:wf(R)}. g(x,y,z):D(x,y,z); \ +\ g(a,b,c) : D(a,b,c) ==> g(a,b,c) : E; \ +\ a:A; b:B; c:C; <>,>> : wf(R) |] ==> \ +\ g(a,b,c) : E"; +by (REPEAT (ares_tac ([SubtypeI,major RS bspec RS bspec RS bspec,major]@prems) 1)); +val rcall3T = result(); + +val rcallTs = [rcallT,rcall2T,rcall3T]; + +(*** Instantiating an induction hypothesis with an equality assumption ***) + +val prems = goal Wfd.thy + "[| g(a) = b; ALL x:{x:A.:wf(R)}.g(x):D(x); \ +\ [| ALL x:{x:A.:wf(R)}.g(x):D(x); b=g(a); g(a) : D(a) |] ==> P; \ +\ ALL x:{x:A.:wf(R)}.g(x):D(x) ==> a:A; \ +\ ALL x:{x:A.:wf(R)}.g(x):D(x) ==> :wf(R) |] ==> \ +\ P"; +brs (prems RL prems) 1; +brs (prems RL [sym]) 1; +br rcallT 1; +by (REPEAT (ares_tac prems 1)); +val hyprcallT = result(); + +val prems = goal Wfd.thy + "[| g(a) = b; ALL x:{x:A.:wf(R)}.g(x):D(x);\ +\ [| b=g(a); g(a) : D(a) |] ==> P; a:A; :wf(R) |] ==> \ +\ P"; +brs (prems) 1; +brs (prems RL [sym]) 1; +br rcallT 1; +by (REPEAT (ares_tac prems 1)); +val hyprcallT = result(); + +val prems = goal Wfd.thy + "[| g(a,b) = c; ALL x:A.ALL y:{y:B.<,>:wf(R)}.g(x,y):D(x,y); \ +\ [| c=g(a,b); g(a,b) : D(a,b) |] ==> P; \ +\ a:A; b:B; <,>:wf(R) |] ==> \ +\ P"; +brs (prems) 1; +brs (prems RL [sym]) 1; +br rcall2T 1; +by (REPEAT (ares_tac prems 1)); +val hyprcall2T = result(); + +val prems = goal Wfd.thy + "[| g(a,b,c) = d; \ +\ ALL x:A.ALL y:B.ALL z:{z:C.<>,>>:wf(R)}.g(x,y,z):D(x,y,z); \ +\ [| d=g(a,b,c); g(a,b,c) : D(a,b,c) |] ==> P; \ +\ a:A; b:B; c:C; <>,>> : wf(R) |] ==> \ +\ P"; +brs (prems) 1; +brs (prems RL [sym]) 1; +br rcall3T 1; +by (REPEAT (ares_tac prems 1)); +val hyprcall3T = result(); + +val hyprcallTs = [hyprcallT,hyprcall2T,hyprcall3T]; + +(*** Rules to Remove Induction Hypotheses after Type Checking ***) + +val prems = goal Wfd.thy + "[| ALL x:{x:A.:wf(R)}.g(x):D(x); P |] ==> \ +\ P"; +by (REPEAT (ares_tac prems 1)); +val rmIH1 = result(); + +val prems = goal Wfd.thy + "[| ALL x:A.ALL y:{y:B.<,>:wf(R)}.g(x,y):D(x,y); P |] ==> \ +\ P"; +by (REPEAT (ares_tac prems 1)); +val rmIH2 = result(); + +val prems = goal Wfd.thy + "[| ALL x:A.ALL y:B.ALL z:{z:C.<>,>>:wf(R)}.g(x,y,z):D(x,y,z); \ +\ P |] ==> \ +\ P"; +by (REPEAT (ares_tac prems 1)); +val rmIH3 = result(); + +val rmIHs = [rmIH1,rmIH2,rmIH3]; + diff -r 000000000000 -r a5a9c433f639 src/CCL/gfp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/gfp.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,133 @@ +(* Title: CCL/gfp + ID: $Id$ + +Modified version of + Title: HOL/gfp + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For gfp.thy. The Knaster-Tarski Theorem for greatest fixed points. +*) + +open Gfp; + +(*** Proof of Knaster-Tarski Theorem using gfp ***) + +(* gfp(f) is the least upper bound of {u. u <= f(u)} *) + +val prems = goalw Gfp.thy [gfp_def] "[| A <= f(A) |] ==> A <= gfp(f)"; +by (rtac (CollectI RS Union_upper) 1); +by (resolve_tac prems 1); +val gfp_upperbound = result(); + +val prems = goalw Gfp.thy [gfp_def] + "[| !!u. u <= f(u) ==> u<=A |] ==> gfp(f) <= A"; +by (REPEAT (ares_tac ([Union_least]@prems) 1)); +by (etac CollectD 1); +val gfp_least = result(); + +val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) <= f(gfp(f))"; +by (EVERY1 [rtac gfp_least, rtac subset_trans, atac, + rtac (mono RS monoD), rtac gfp_upperbound, atac]); +val gfp_lemma2 = result(); + +val [mono] = goal Gfp.thy "mono(f) ==> f(gfp(f)) <= gfp(f)"; +by (EVERY1 [rtac gfp_upperbound, rtac (mono RS monoD), + rtac gfp_lemma2, rtac mono]); +val gfp_lemma3 = result(); + +val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) = f(gfp(f))"; +by (REPEAT (resolve_tac [equalityI,gfp_lemma2,gfp_lemma3,mono] 1)); +val gfp_Tarski = result(); + +(*** Coinduction rules for greatest fixed points ***) + +(*weak version*) +val prems = goal Gfp.thy + "[| a: A; A <= f(A) |] ==> a : gfp(f)"; +by (rtac (gfp_upperbound RS subsetD) 1); +by (REPEAT (ares_tac prems 1)); +val coinduct = result(); + +val [prem,mono] = goal Gfp.thy + "[| A <= f(A) Un gfp(f); mono(f) |] ==> \ +\ A Un gfp(f) <= f(A Un gfp(f))"; +by (rtac subset_trans 1); +by (rtac (mono RS mono_Un) 2); +by (rtac (mono RS gfp_Tarski RS subst) 1); +by (rtac (prem RS Un_least) 1); +by (rtac Un_upper2 1); +val coinduct2_lemma = result(); + +(*strong version, thanks to Martin Coen*) +val prems = goal Gfp.thy + "[| a: A; A <= f(A) Un gfp(f); mono(f) |] ==> a : gfp(f)"; +by (rtac (coinduct2_lemma RSN (2,coinduct)) 1); +by (REPEAT (resolve_tac (prems@[UnI1]) 1)); +val coinduct2 = result(); + +(*** Even Stronger version of coinduct [by Martin Coen] + - instead of the condition A <= f(A) + consider A <= (f(A) Un f(f(A)) ...) Un gfp(A) ***) + +val [prem] = goal Gfp.thy "mono(f) ==> mono(%x.f(x) Un A Un B)"; +by (REPEAT (ares_tac [subset_refl, monoI, Un_mono, prem RS monoD] 1)); +val coinduct3_mono_lemma= result(); + +val [prem,mono] = goal Gfp.thy + "[| A <= f(lfp(%x.f(x) Un A Un gfp(f))); mono(f) |] ==> \ +\ lfp(%x.f(x) Un A Un gfp(f)) <= f(lfp(%x.f(x) Un A Un gfp(f)))"; +by (rtac subset_trans 1); +br (mono RS coinduct3_mono_lemma RS lfp_lemma3) 1; +by (rtac (Un_least RS Un_least) 1); +br subset_refl 1; +br prem 1; +br (mono RS gfp_Tarski RS equalityD1 RS subset_trans) 1; +by (rtac (mono RS monoD) 1); +by (rtac (mono RS coinduct3_mono_lemma RS lfp_Tarski RS ssubst) 1); +by (rtac Un_upper2 1); +val coinduct3_lemma = result(); + +val prems = goal Gfp.thy + "[| a:A; A <= f(lfp(%x.f(x) Un A Un gfp(f))); mono(f) |] ==> a : gfp(f)"; +by (rtac (coinduct3_lemma RSN (2,coinduct)) 1); +brs (prems RL [coinduct3_mono_lemma RS lfp_Tarski RS ssubst]) 1; +br (UnI2 RS UnI1) 1; +by (REPEAT (resolve_tac prems 1)); +val coinduct3 = result(); + + +(** Definition forms of gfp_Tarski, to control unfolding **) + +val [rew,mono] = goal Gfp.thy "[| h==gfp(f); mono(f) |] ==> h = f(h)"; +by (rewtac rew); +by (rtac (mono RS gfp_Tarski) 1); +val def_gfp_Tarski = result(); + +val rew::prems = goal Gfp.thy + "[| h==gfp(f); a:A; A <= f(A) |] ==> a: h"; +by (rewtac rew); +by (REPEAT (ares_tac (prems @ [coinduct]) 1)); +val def_coinduct = result(); + +val rew::prems = goal Gfp.thy + "[| h==gfp(f); a:A; A <= f(A) Un h; mono(f) |] ==> a: h"; +by (rewtac rew); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct2]) 1)); +val def_coinduct2 = result(); + +val rew::prems = goal Gfp.thy + "[| h==gfp(f); a:A; A <= f(lfp(%x.f(x) Un A Un h)); mono(f) |] ==> a: h"; +by (rewtac rew); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct3]) 1)); +val def_coinduct3 = result(); + +(*Monotonicity of gfp!*) +val prems = goal Gfp.thy + "[| mono(f); !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)"; +by (rtac gfp_upperbound 1); +by (rtac subset_trans 1); +by (rtac gfp_lemma2 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +val gfp_mono = result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/gfp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/gfp.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,14 @@ +(* Title: HOL/gfp.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Greatest fixed points +*) + +Gfp = Lfp + +consts gfp :: "['a set=>'a set] => 'a set" +rules + (*greatest fixed point*) + gfp_def "gfp(f) == Union({u. u <= f(u)})" +end diff -r 000000000000 -r a5a9c433f639 src/CCL/hered.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/hered.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,196 @@ +(* Title: CCL/hered + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For hered.thy. +*) + +open Hered; + +fun type_of_terms (Const("Trueprop",_) $ (Const("op =",(Type ("fun", [t,_])))$_$_)) = t; + +val cong_rls = ccl_mk_congs Hered.thy ["HTTgen"]; + +(*** Hereditary Termination ***) + +goalw Hered.thy [HTTgen_def] "mono(%X.HTTgen(X))"; +br monoI 1; +by (fast_tac set_cs 1); +val HTTgen_mono = result(); + +goalw Hered.thy [HTTgen_def] + "t : HTTgen(A) <-> t=true | t=false | (EX a b.t= & a : A & b : A) | \ +\ (EX f.t=lam x.f(x) & (ALL x.f(x) : A))"; +by (fast_tac set_cs 1); +val HTTgenXH = result(); + +goal Hered.thy + "t : HTT <-> t=true | t=false | (EX a b.t= & a : HTT & b : HTT) | \ +\ (EX f.t=lam x.f(x) & (ALL x.f(x) : HTT))"; +br (rewrite_rule [HTTgen_def] + (HTTgen_mono RS (HTT_def RS def_gfp_Tarski) RS XHlemma1)) 1; +by (fast_tac set_cs 1); +val HTTXH = result(); + +(*** Introduction Rules for HTT ***) + +goal Hered.thy "~ bot : HTT"; +by (fast_tac (term_cs addDs [XH_to_D HTTXH]) 1); +val HTT_bot = result(); + +goal Hered.thy "true : HTT"; +by (fast_tac (term_cs addIs [XH_to_I HTTXH]) 1); +val HTT_true = result(); + +goal Hered.thy "false : HTT"; +by (fast_tac (term_cs addIs [XH_to_I HTTXH]) 1); +val HTT_false = result(); + +goal Hered.thy " : HTT <-> a : HTT & b : HTT"; +br (HTTXH RS iff_trans) 1; +by (fast_tac term_cs 1); +val HTT_pair = result(); + +goal Hered.thy "lam x.f(x) : HTT <-> (ALL x. f(x) : HTT)"; +br (HTTXH RS iff_trans) 1; +by (SIMP_TAC term_ss 1); +by (safe_tac term_cs); +by (ASM_SIMP_TAC term_ss 1); +by (fast_tac term_cs 1); +val HTT_lam = result(); + +local + val raw_HTTrews = [HTT_bot,HTT_true,HTT_false,HTT_pair,HTT_lam]; + fun mk_thm s = prove_goalw Hered.thy data_defs s (fn _ => + [SIMP_TAC (term_ss addrews raw_HTTrews) 1]); +in + val HTT_rews = raw_HTTrews @ + map mk_thm ["one : HTT", + "inl(a) : HTT <-> a : HTT", + "inr(b) : HTT <-> b : HTT", + "zero : HTT", + "succ(n) : HTT <-> n : HTT", + "[] : HTT", + "x.xs : HTT <-> x : HTT & xs : HTT"]; +end; + +val HTT_Is = HTT_rews @ (HTT_rews RL [iffD2]); + +(*** Coinduction for HTT ***) + +val prems = goal Hered.thy "[| t : R; R <= HTTgen(R) |] ==> t : HTT"; +br (HTT_def RS def_coinduct) 1; +by (REPEAT (ares_tac prems 1)); +val HTT_coinduct = result(); + +fun HTT_coinduct_tac s i = res_inst_tac [("R",s)] HTT_coinduct i; + +val prems = goal Hered.thy + "[| t : R; R <= HTTgen(lfp(%x. HTTgen(x) Un R Un HTT)) |] ==> t : HTT"; +br (HTTgen_mono RSN(3,HTT_def RS def_coinduct3)) 1; +by (REPEAT (ares_tac prems 1)); +val HTT_coinduct3 = result(); +val HTT_coinduct3_raw = rewrite_rule [HTTgen_def] HTT_coinduct3; + +fun HTT_coinduct3_tac s i = res_inst_tac [("R",s)] HTT_coinduct3 i; + +val HTTgenIs = map (mk_genIs Hered.thy data_defs HTTgenXH HTTgen_mono) + ["true : HTTgen(R)", + "false : HTTgen(R)", + "[| a : R; b : R |] ==> : HTTgen(R)", + "[| !!x. b(x) : R |] ==> lam x.b(x) : HTTgen(R)", + "one : HTTgen(R)", + "a : lfp(%x. HTTgen(x) Un R Un HTT) ==> \ +\ inl(a) : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))", + "b : lfp(%x. HTTgen(x) Un R Un HTT) ==> \ +\ inr(b) : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))", + "zero : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))", + "n : lfp(%x. HTTgen(x) Un R Un HTT) ==> \ +\ succ(n) : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))", + "[] : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))", + "[| h : lfp(%x. HTTgen(x) Un R Un HTT); t : lfp(%x. HTTgen(x) Un R Un HTT) |] ==>\ +\ h.t : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))"]; + +(*** Formation Rules for Types ***) + +goal Hered.thy "Unit <= HTT"; +by (SIMP_TAC (CCL_ss addrews ([subsetXH,UnitXH] @ HTT_rews)) 1); +val UnitF = result(); + +goal Hered.thy "Bool <= HTT"; +by (SIMP_TAC (CCL_ss addrews ([subsetXH,BoolXH] @ HTT_rews)) 1); +by (fast_tac (set_cs addIs HTT_Is @ (prems RL [subsetD])) 1); +val BoolF = result(); + +val prems = goal Hered.thy "[| A <= HTT; B <= HTT |] ==> A + B <= HTT"; +by (SIMP_TAC (CCL_ss addrews ([subsetXH,PlusXH] @ HTT_rews)) 1); +by (fast_tac (set_cs addIs HTT_Is @ (prems RL [subsetD])) 1); +val PlusF = result(); + +val prems = goal Hered.thy + "[| A <= HTT; !!x.x:A ==> B(x) <= HTT |] ==> SUM x:A.B(x) <= HTT"; +by (SIMP_TAC (CCL_ss addrews ([subsetXH,SgXH] @ HTT_rews)) 1); +by (fast_tac (set_cs addIs HTT_Is @ (prems RL [subsetD])) 1); +val SigmaF = result(); + +(*** Formation Rules for Recursive types - using coinduction these only need ***) +(*** exhaution rule for type-former ***) + +(*Proof by induction - needs induction rule for type*) +goal Hered.thy "Nat <= HTT"; +by (SIMP_TAC (term_ss addrews [subsetXH]) 1); +by (safe_tac set_cs); +be Nat_ind 1; +by (ALLGOALS (fast_tac (set_cs addIs HTT_Is @ (prems RL [subsetD])))); +val NatF = result(); + +goal Hered.thy "Nat <= HTT"; +by (safe_tac set_cs); +be HTT_coinduct3 1; +by (fast_tac (set_cs addIs HTTgenIs + addSEs [HTTgen_mono RS ci3_RI] addEs [XH_to_E NatXH]) 1); +val NatF = result(); + +val [prem] = goal Hered.thy "A <= HTT ==> List(A) <= HTT"; +by (safe_tac set_cs); +be HTT_coinduct3 1; +by (fast_tac (set_cs addSIs HTTgenIs + addSEs [HTTgen_mono RS ci3_RI,prem RS subsetD RS (HTTgen_mono RS ci3_AI)] + addEs [XH_to_E ListXH]) 1); +val ListF = result(); + +val [prem] = goal Hered.thy "A <= HTT ==> Lists(A) <= HTT"; +by (safe_tac set_cs); +be HTT_coinduct3 1; +by (fast_tac (set_cs addSIs HTTgenIs + addSEs [HTTgen_mono RS ci3_RI,prem RS subsetD RS (HTTgen_mono RS ci3_AI)] + addEs [XH_to_E ListsXH]) 1); +val ListsF = result(); + +val [prem] = goal Hered.thy "A <= HTT ==> ILists(A) <= HTT"; +by (safe_tac set_cs); +be HTT_coinduct3 1; +by (fast_tac (set_cs addSIs HTTgenIs + addSEs [HTTgen_mono RS ci3_RI,prem RS subsetD RS (HTTgen_mono RS ci3_AI)] + addEs [XH_to_E IListsXH]) 1); +val IListsF = result(); + +(*** A possible use for this predicate is proving equality from pre-order ***) +(*** but it seems as easy (and more general) to do this directly by coinduction ***) +(* +val prems = goal Hered.thy "[| t : HTT; t [= u |] ==> u [= t"; +by (po_coinduct_tac "{p. EX a b.p= & b : HTT & b [= a}" 1); +by (fast_tac (ccl_cs addIs prems) 1); +by (safe_tac ccl_cs); +bd (poXH RS iffD1) 1; +by (safe_tac (set_cs addSEs [HTT_bot RS notE])); +by (REPEAT_SOME (rtac (POgenXH RS iffD2) ORELSE' etac rev_mp)); +by (ALLGOALS (SIMP_TAC (term_ss addrews HTT_rews))); +by (ALLGOALS (fast_tac ccl_cs)); +val HTT_po_op = result(); + +val prems = goal Hered.thy "[| t : HTT; t [= u |] ==> t = u"; +by (REPEAT (ares_tac (prems @ [conjI RS (eq_iff RS iffD2),HTT_po_op]) 1)); +val HTT_po_eq = result(); +*) diff -r 000000000000 -r a5a9c433f639 src/CCL/hered.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/hered.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,30 @@ +(* Title: CCL/hered.thy + ID: $Id$ + Author: Martin Coen + Copyright 1993 University of Cambridge + +Hereditary Termination - cf. Martin Lo\"f + +Note that this is based on an untyped equality and so lam x.b(x) is only +hereditarily terminating if ALL x.b(x) is. Not so useful for functions! + +*) + +Hered = Type + + +consts + (*** Predicates ***) + HTTgen :: "i set => i set" + HTT :: "i set" + + +rules + + (*** Definitions of Hereditary Termination ***) + + HTTgen_def + "HTTgen(R) == {t. t=true | t=false | (EX a b.t= & a : R & b : R) | \ +\ (EX f. t=lam x.f(x) & (ALL x.f(x) : R))}" + HTT_def "HTT == gfp(HTTgen)" + +end diff -r 000000000000 -r a5a9c433f639 src/CCL/lfp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/lfp.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,82 @@ +(* Title: CCL/lfp + ID: $Id$ + +Modified version of + Title: HOL/lfp.ML + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For lfp.thy. The Knaster-Tarski Theorem +*) + +open Lfp; + +(*** Proof of Knaster-Tarski Theorem ***) + +(* lfp(f) is the greatest lower bound of {u. f(u) <= u} *) + +val prems = goalw Lfp.thy [lfp_def] "[| f(A) <= A |] ==> lfp(f) <= A"; +by (rtac (CollectI RS Inter_lower) 1); +by (resolve_tac prems 1); +val lfp_lowerbound = result(); + +val prems = goalw Lfp.thy [lfp_def] + "[| !!u. f(u) <= u ==> A<=u |] ==> A <= lfp(f)"; +by (REPEAT (ares_tac ([Inter_greatest]@prems) 1)); +by (etac CollectD 1); +val lfp_greatest = result(); + +val [mono] = goal Lfp.thy "mono(f) ==> f(lfp(f)) <= lfp(f)"; +by (EVERY1 [rtac lfp_greatest, rtac subset_trans, + rtac (mono RS monoD), rtac lfp_lowerbound, atac, atac]); +val lfp_lemma2 = result(); + +val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) <= f(lfp(f))"; +by (EVERY1 [rtac lfp_lowerbound, rtac (mono RS monoD), + rtac lfp_lemma2, rtac mono]); +val lfp_lemma3 = result(); + +val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) = f(lfp(f))"; +by (REPEAT (resolve_tac [equalityI,lfp_lemma2,lfp_lemma3,mono] 1)); +val lfp_Tarski = result(); + + +(*** General induction rule for least fixed points ***) + +val [lfp,mono,indhyp] = goal Lfp.thy + "[| a: lfp(f); mono(f); \ +\ !!x. [| x: f(lfp(f) Int {x.P(x)}) |] ==> P(x) \ +\ |] ==> P(a)"; +by (res_inst_tac [("a","a")] (Int_lower2 RS subsetD RS CollectD) 1); +by (rtac (lfp RSN (2, lfp_lowerbound RS subsetD)) 1); +by (EVERY1 [rtac Int_greatest, rtac subset_trans, + rtac (Int_lower1 RS (mono RS monoD)), + rtac (mono RS lfp_lemma2), + rtac (CollectI RS subsetI), rtac indhyp, atac]); +val induct = result(); + +(** Definition forms of lfp_Tarski and induct, to control unfolding **) + +val [rew,mono] = goal Lfp.thy "[| h==lfp(f); mono(f) |] ==> h = f(h)"; +by (rewtac rew); +by (rtac (mono RS lfp_Tarski) 1); +val def_lfp_Tarski = result(); + +val rew::prems = goal Lfp.thy + "[| A == lfp(f); a:A; mono(f); \ +\ !!x. [| x: f(A Int {x.P(x)}) |] ==> P(x) \ +\ |] ==> P(a)"; +by (EVERY1 [rtac induct, (*backtracking to force correct induction*) + REPEAT1 o (ares_tac (map (rewrite_rule [rew]) prems))]); +val def_induct = result(); + +(*Monotonicity of lfp!*) +val prems = goal Lfp.thy + "[| mono(g); !!Z. f(Z)<=g(Z) |] ==> lfp(f) <= lfp(g)"; +by (rtac lfp_lowerbound 1); +by (rtac subset_trans 1); +by (resolve_tac prems 1); +by (rtac lfp_lemma2 1); +by (resolve_tac prems 1); +val lfp_mono = result(); + diff -r 000000000000 -r a5a9c433f639 src/CCL/lfp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/lfp.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,14 @@ +(* Title: HOL/lfp.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +The Knaster-Tarski Theorem +*) + +Lfp = Set + +consts lfp :: "['a set=>'a set] => 'a set" +rules + (*least fixed point*) + lfp_def "lfp(f) == Inter({u. f(u) <= u})" +end diff -r 000000000000 -r a5a9c433f639 src/CCL/mono.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/mono.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,46 @@ +(* Title: CCL/mono + ID: $Id$ + +Modified version of + Title: HOL/mono + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Monotonicity of various operations +*) + +writeln"File HOL/mono"; + +val prems = goal Set.thy "A<=B ==> Union(A) <= Union(B)"; +by (cfast_tac prems 1); +val Union_mono = result(); + +val prems = goal Set.thy "[| B<=A |] ==> Inter(A) <= Inter(B)"; +by (cfast_tac prems 1); +val Inter_anti_mono = result(); + +val prems = goal Set.thy + "[| A<=B; !!x. x:A ==> f(x)<=g(x) |] ==> \ +\ (UN x:A. f(x)) <= (UN x:B. g(x))"; +by (REPEAT (eresolve_tac [UN_E,ssubst] 1 + ORELSE ares_tac ((prems RL [subsetD]) @ [subsetI,UN_I]) 1)); +val UN_mono = result(); + +val prems = goal Set.thy + "[| B<=A; !!x. x:A ==> f(x)<=g(x) |] ==> \ +\ (INT x:A. f(x)) <= (INT x:A. g(x))"; +by (REPEAT (ares_tac ((prems RL [subsetD]) @ [subsetI,INT_I]) 1 + ORELSE etac INT_D 1)); +val INT_anti_mono = result(); + +val prems = goal Set.thy "[| A<=C; B<=D |] ==> A Un B <= C Un D"; +by (cfast_tac prems 1); +val Un_mono = result(); + +val prems = goal Set.thy "[| A<=C; B<=D |] ==> A Int B <= C Int D"; +by (cfast_tac prems 1); +val Int_mono = result(); + +val prems = goal Set.thy "[| A<=B |] ==> Compl(B) <= Compl(A)"; +by (cfast_tac prems 1); +val Compl_anti_mono = result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/set.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/set.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,355 @@ +(* Title: set/set + ID: $Id$ + +For set.thy. + +Modified version of + Title: HOL/set + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For set.thy. Set theory for higher-order logic. A set is simply a predicate. +*) + +open Set; + +val [prem] = goal Set.thy "[| P(a) |] ==> a : {x.P(x)}"; +by (rtac (mem_Collect_iff RS iffD2) 1); +by (rtac prem 1); +val CollectI = result(); + +val prems = goal Set.thy "[| a : {x.P(x)} |] ==> P(a)"; +by (resolve_tac (prems RL [mem_Collect_iff RS iffD1]) 1); +val CollectD = result(); + +val [prem] = goal Set.thy "[| !!x. x:A <-> x:B |] ==> A = B"; +by (rtac (set_extension RS iffD2) 1); +by (rtac (prem RS allI) 1); +val set_ext = result(); + +val prems = goal Set.thy "[| !!x. P(x) <-> Q(x) |] ==> {x. P(x)} = {x. Q(x)}"; +by (REPEAT (ares_tac [set_ext,iffI,CollectI] 1 ORELSE + eresolve_tac ([CollectD] RL (prems RL [iffD1,iffD2])) 1)); +val Collect_cong = result(); + +val CollectE = make_elim CollectD; + +(*** Bounded quantifiers ***) + +val prems = goalw Set.thy [Ball_def] + "[| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)"; +by (REPEAT (ares_tac (prems @ [allI,impI]) 1)); +val ballI = result(); + +val [major,minor] = goalw Set.thy [Ball_def] + "[| ALL x:A. P(x); x:A |] ==> P(x)"; +by (rtac (minor RS (major RS spec RS mp)) 1); +val bspec = result(); + +val major::prems = goalw Set.thy [Ball_def] + "[| ALL x:A. P(x); P(x) ==> Q; ~ x:A ==> Q |] ==> Q"; +by (rtac (major RS spec RS impCE) 1); +by (REPEAT (eresolve_tac prems 1)); +val ballE = result(); + +(*Takes assumptions ALL x:A.P(x) and a:A; creates assumption P(a)*) +fun ball_tac i = etac ballE i THEN contr_tac (i+1); + +val prems = goalw Set.thy [Bex_def] + "[| P(x); x:A |] ==> EX x:A. P(x)"; +by (REPEAT (ares_tac (prems @ [exI,conjI]) 1)); +val bexI = result(); + +val bexCI = prove_goal Set.thy + "[| EX x:A. ~P(x) ==> P(a); a:A |] ==> EX x:A.P(x)" + (fn prems=> + [ (rtac classical 1), + (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ]); + +val major::prems = goalw Set.thy [Bex_def] + "[| EX x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q"; +by (rtac (major RS exE) 1); +by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)); +val bexE = result(); + +(*Trival rewrite rule; (! x:A.P)=P holds only if A is nonempty!*) +val prems = goal Set.thy + "(ALL x:A. True) <-> True"; +by (REPEAT (ares_tac [TrueI,ballI,iffI] 1)); +val ball_rew = result(); + +(** Congruence rules **) + +val prems = goal Set.thy + "[| A=A'; !!x. x:A' ==> P(x) <-> P'(x) |] ==> \ +\ (ALL x:A. P(x)) <-> (ALL x:A'. P'(x))"; +by (resolve_tac (prems RL [ssubst,iffD2]) 1); +by (REPEAT (ares_tac [ballI,iffI] 1 + ORELSE eresolve_tac ([make_elim bspec, mp] @ (prems RL [iffE])) 1)); +val ball_cong = result(); + +val prems = goal Set.thy + "[| A=A'; !!x. x:A' ==> P(x) <-> P'(x) |] ==> \ +\ (EX x:A. P(x)) <-> (EX x:A'. P'(x))"; +by (resolve_tac (prems RL [ssubst,iffD2]) 1); +by (REPEAT (etac bexE 1 + ORELSE ares_tac ([bexI,iffI] @ (prems RL [iffD1,iffD2])) 1)); +val bex_cong = result(); + +(*** Rules for subsets ***) + +val prems = goalw Set.thy [subset_def] "(!!x.x:A ==> x:B) ==> A <= B"; +by (REPEAT (ares_tac (prems @ [ballI]) 1)); +val subsetI = result(); + +(*Rule in Modus Ponens style*) +val major::prems = goalw Set.thy [subset_def] "[| A <= B; c:A |] ==> c:B"; +by (rtac (major RS bspec) 1); +by (resolve_tac prems 1); +val subsetD = result(); + +(*Classical elimination rule*) +val major::prems = goalw Set.thy [subset_def] + "[| A <= B; ~(c:A) ==> P; c:B ==> P |] ==> P"; +by (rtac (major RS ballE) 1); +by (REPEAT (eresolve_tac prems 1)); +val subsetCE = result(); + +(*Takes assumptions A<=B; c:A and creates the assumption c:B *) +fun set_mp_tac i = etac subsetCE i THEN mp_tac i; + +val subset_refl = prove_goal Set.thy "A <= A" + (fn _=> [ (REPEAT (ares_tac [subsetI] 1)) ]); + +goal Set.thy "!!A B C. [| A<=B; B<=C |] ==> A<=C"; +br subsetI 1; +by (REPEAT (eresolve_tac [asm_rl, subsetD] 1)); +val subset_trans = result(); + + +(*** Rules for equality ***) + +(*Anti-symmetry of the subset relation*) +val prems = goal Set.thy "[| A <= B; B <= A |] ==> A = B"; +by (rtac (iffI RS set_ext) 1); +by (REPEAT (ares_tac (prems RL [subsetD]) 1)); +val subset_antisym = result(); +val equalityI = subset_antisym; + +(* Equality rules from ZF set theory -- are they appropriate here? *) +val prems = goal Set.thy "A = B ==> A<=B"; +by (resolve_tac (prems RL [subst]) 1); +by (rtac subset_refl 1); +val equalityD1 = result(); + +val prems = goal Set.thy "A = B ==> B<=A"; +by (resolve_tac (prems RL [subst]) 1); +by (rtac subset_refl 1); +val equalityD2 = result(); + +val prems = goal Set.thy + "[| A = B; [| A<=B; B<=A |] ==> P |] ==> P"; +by (resolve_tac prems 1); +by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1)); +val equalityE = result(); + +val major::prems = goal Set.thy + "[| A = B; [| c:A; c:B |] ==> P; [| ~ c:A; ~ c:B |] ==> P |] ==> P"; +by (rtac (major RS equalityE) 1); +by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1)); +val equalityCE = result(); + +(*Lemma for creating induction formulae -- for "pattern matching" on p + To make the induction hypotheses usable, apply "spec" or "bspec" to + put universal quantifiers over the free variables in p. *) +val prems = goal Set.thy + "[| p:A; !!z. z:A ==> p=z --> R |] ==> R"; +by (rtac mp 1); +by (REPEAT (resolve_tac (refl::prems) 1)); +val setup_induction = result(); + +goal Set.thy "{x.x:A} = A"; +by (REPEAT (ares_tac [equalityI,subsetI,CollectI] 1 ORELSE eresolve_tac [CollectD] 1)); +val trivial_set = result(); + +(*** Rules for binary union -- Un ***) + +val prems = goalw Set.thy [Un_def] "c:A ==> c : A Un B"; +by (REPEAT (resolve_tac (prems @ [CollectI,disjI1]) 1)); +val UnI1 = result(); + +val prems = goalw Set.thy [Un_def] "c:B ==> c : A Un B"; +by (REPEAT (resolve_tac (prems @ [CollectI,disjI2]) 1)); +val UnI2 = result(); + +(*Classical introduction rule: no commitment to A vs B*) +val UnCI = prove_goal Set.thy "(~c:B ==> c:A) ==> c : A Un B" + (fn prems=> + [ (rtac classical 1), + (REPEAT (ares_tac (prems@[UnI1,notI]) 1)), + (REPEAT (ares_tac (prems@[UnI2,notE]) 1)) ]); + +val major::prems = goalw Set.thy [Un_def] + "[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P"; +by (rtac (major RS CollectD RS disjE) 1); +by (REPEAT (eresolve_tac prems 1)); +val UnE = result(); + + +(*** Rules for small intersection -- Int ***) + +val prems = goalw Set.thy [Int_def] + "[| c:A; c:B |] ==> c : A Int B"; +by (REPEAT (resolve_tac (prems @ [CollectI,conjI]) 1)); +val IntI = result(); + +val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:A"; +by (rtac (major RS CollectD RS conjunct1) 1); +val IntD1 = result(); + +val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:B"; +by (rtac (major RS CollectD RS conjunct2) 1); +val IntD2 = result(); + +val [major,minor] = goal Set.thy + "[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P"; +by (rtac minor 1); +by (rtac (major RS IntD1) 1); +by (rtac (major RS IntD2) 1); +val IntE = result(); + + +(*** Rules for set complement -- Compl ***) + +val prems = goalw Set.thy [Compl_def] + "[| c:A ==> False |] ==> c : Compl(A)"; +by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1)); +val ComplI = result(); + +(*This form, with negated conclusion, works well with the Classical prover. + Negated assumptions behave like formulae on the right side of the notional + turnstile...*) +val major::prems = goalw Set.thy [Compl_def] + "[| c : Compl(A) |] ==> ~c:A"; +by (rtac (major RS CollectD) 1); +val ComplD = result(); + +val ComplE = make_elim ComplD; + + +(*** Empty sets ***) + +goalw Set.thy [empty_def] "{x.False} = {}"; +br refl 1; +val empty_eq = result(); + +val [prem] = goalw Set.thy [empty_def] "a : {} ==> P"; +by (rtac (prem RS CollectD RS FalseE) 1); +val emptyD = result(); + +val emptyE = make_elim emptyD; + +val [prem] = goal Set.thy "~ A={} ==> (EX x.x:A)"; +br (prem RS swap) 1; +br equalityI 1; +by (ALLGOALS (fast_tac (FOL_cs addSIs [subsetI] addSEs [emptyD]))); +val not_emptyD = result(); + +(*** Singleton sets ***) + +goalw Set.thy [singleton_def] "a : {a}"; +by (rtac CollectI 1); +by (rtac refl 1); +val singletonI = result(); + +val [major] = goalw Set.thy [singleton_def] "b : {a} ==> b=a"; +by (rtac (major RS CollectD) 1); +val singletonD = result(); + +val singletonE = make_elim singletonD; + +(*** Unions of families ***) + +(*The order of the premises presupposes that A is rigid; b may be flexible*) +val prems = goalw Set.thy [UNION_def] + "[| a:A; b: B(a) |] ==> b: (UN x:A. B(x))"; +by (REPEAT (resolve_tac (prems @ [bexI,CollectI]) 1)); +val UN_I = result(); + +val major::prems = goalw Set.thy [UNION_def] + "[| b : (UN x:A. B(x)); !!x.[| x:A; b: B(x) |] ==> R |] ==> R"; +by (rtac (major RS CollectD RS bexE) 1); +by (REPEAT (ares_tac prems 1)); +val UN_E = result(); + +val prems = goal Set.thy + "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ +\ (UN x:A. C(x)) = (UN x:B. D(x))"; +by (REPEAT (etac UN_E 1 + ORELSE ares_tac ([UN_I,equalityI,subsetI] @ + (prems RL [equalityD1,equalityD2] RL [subsetD])) 1)); +val UN_cong = result(); + +(*** Intersections of families -- INTER x:A. B(x) is Inter(B)``A ) *) + +val prems = goalw Set.thy [INTER_def] + "(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))"; +by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1)); +val INT_I = result(); + +val major::prems = goalw Set.thy [INTER_def] + "[| b : (INT x:A. B(x)); a:A |] ==> b: B(a)"; +by (rtac (major RS CollectD RS bspec) 1); +by (resolve_tac prems 1); +val INT_D = result(); + +(*"Classical" elimination rule -- does not require proving X:C *) +val major::prems = goalw Set.thy [INTER_def] + "[| b : (INT x:A. B(x)); b: B(a) ==> R; ~ a:A ==> R |] ==> R"; +by (rtac (major RS CollectD RS ballE) 1); +by (REPEAT (eresolve_tac prems 1)); +val INT_E = result(); + +val prems = goal Set.thy + "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ +\ (INT x:A. C(x)) = (INT x:B. D(x))"; +by (REPEAT_FIRST (resolve_tac [INT_I,equalityI,subsetI])); +by (REPEAT (dtac INT_D 1 + ORELSE ares_tac (prems RL [equalityD1,equalityD2] RL [subsetD]) 1)); +val INT_cong = result(); + +(*** Rules for Unions ***) + +(*The order of the premises presupposes that C is rigid; A may be flexible*) +val prems = goalw Set.thy [Union_def] + "[| X:C; A:X |] ==> A : Union(C)"; +by (REPEAT (resolve_tac (prems @ [UN_I]) 1)); +val UnionI = result(); + +val major::prems = goalw Set.thy [Union_def] + "[| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R"; +by (rtac (major RS UN_E) 1); +by (REPEAT (ares_tac prems 1)); +val UnionE = result(); + +(*** Rules for Inter ***) + +val prems = goalw Set.thy [Inter_def] + "[| !!X. X:C ==> A:X |] ==> A : Inter(C)"; +by (REPEAT (ares_tac ([INT_I] @ prems) 1)); +val InterI = result(); + +(*A "destruct" rule -- every X in C contains A as an element, but + A:X can hold when X:C does not! This rule is analogous to "spec". *) +val major::prems = goalw Set.thy [Inter_def] + "[| A : Inter(C); X:C |] ==> A:X"; +by (rtac (major RS INT_D) 1); +by (resolve_tac prems 1); +val InterD = result(); + +(*"Classical" elimination rule -- does not require proving X:C *) +val major::prems = goalw Set.thy [Inter_def] + "[| A : Inter(C); A:X ==> R; ~ X:C ==> R |] ==> R"; +by (rtac (major RS INT_E) 1); +by (REPEAT (eresolve_tac prems 1)); +val InterE = result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/set.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/set.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,71 @@ +(* Title: CCL/set.thy + ID: $Id$ + +Modified version of HOL/set.thy that extends FOL + +*) + +Set = FOL + + +types + set 1 + +arities + set :: (term) term + +consts + Collect :: "['a => o] => 'a set" (*comprehension*) + Compl :: "('a set) => 'a set" (*complement*) + Int :: "['a set, 'a set] => 'a set" (infixl 70) + Un :: "['a set, 'a set] => 'a set" (infixl 65) + Union, Inter :: "(('a set)set) => 'a set" (*...of a set*) + UNION, INTER :: "['a set, 'a => 'b set] => 'b set" (*general*) + Ball, Bex :: "['a set, 'a => o] => o" (*bounded quants*) + mono :: "['a set => 'b set] => o" (*monotonicity*) + ":" :: "['a, 'a set] => o" (infixl 50) (*membership*) + "<=" :: "['a set, 'a set] => o" (infixl 50) + singleton :: "'a => 'a set" ("{_}") + empty :: "'a set" ("{}") + "oo" :: "['b => 'c, 'a => 'b, 'a] => 'c" (infixr 50) (*composition*) + + "@Coll" :: "[idt, o] => 'a set" ("(1{_./ _})") (*collection*) + + (* Big Intersection / Union *) + + "@INTER" :: "[idt, 'a set, 'b set] => 'b set" ("(INT _:_./ _)" [0, 0, 0] 10) + "@UNION" :: "[idt, 'a set, 'b set] => 'b set" ("(UN _:_./ _)" [0, 0, 0] 10) + + (* Bounded Quantifiers *) + + "@Ball" :: "[idt, 'a set, o] => o" ("(ALL _:_./ _)" [0, 0, 0] 10) + "@Bex" :: "[idt, 'a set, o] => o" ("(EX _:_./ _)" [0, 0, 0] 10) + + +translations + "{x. P}" == "Collect(%x. P)" + "INT x:A. B" == "INTER(A, %x. B)" + "UN x:A. B" == "UNION(A, %x. B)" + "ALL x:A. P" == "Ball(A, %x. P)" + "EX x:A. P" == "Bex(A, %x. P)" + + +rules + mem_Collect_iff "(a : {x.P(x)}) <-> P(a)" + set_extension "A=B <-> (ALL x.x:A <-> x:B)" + + Ball_def "Ball(A, P) == ALL x. x:A --> P(x)" + Bex_def "Bex(A, P) == EX x. x:A & P(x)" + mono_def "mono(f) == (ALL A B. A <= B --> f(A) <= f(B))" + subset_def "A <= B == ALL x:A. x:B" + singleton_def "{a} == {x.x=a}" + empty_def "{} == {x.False}" + Un_def "A Un B == {x.x:A | x:B}" + Int_def "A Int B == {x.x:A & x:B}" + Compl_def "Compl(A) == {x. ~x:A}" + INTER_def "INTER(A, B) == {y. ALL x:A. y: B(x)}" + UNION_def "UNION(A, B) == {y. EX x:A. y: B(x)}" + Inter_def "Inter(S) == (INT x:S. x)" + Union_def "Union(S) == (UN x:S. x)" + +end + diff -r 000000000000 -r a5a9c433f639 src/CCL/subset.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/subset.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,129 @@ +(* Title: CCL/subset + ID: $Id$ + +Modified version of + Title: HOL/subset + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Derived rules involving subsets +Union and Intersection as lattice operations +*) + +(*** Big Union -- least upper bound of a set ***) + +val prems = goal Set.thy + "B:A ==> B <= Union(A)"; +by (REPEAT (ares_tac (prems@[subsetI,UnionI]) 1)); +val Union_upper = result(); + +val prems = goal Set.thy + "[| !!X. X:A ==> X<=C |] ==> Union(A) <= C"; +by (REPEAT (ares_tac [subsetI] 1 + ORELSE eresolve_tac ([UnionE] @ (prems RL [subsetD])) 1)); +val Union_least = result(); + + +(*** Big Intersection -- greatest lower bound of a set ***) + +val prems = goal Set.thy + "B:A ==> Inter(A) <= B"; +by (REPEAT (resolve_tac (prems@[subsetI]) 1 + ORELSE etac InterD 1)); +val Inter_lower = result(); + +val prems = goal Set.thy + "[| !!X. X:A ==> C<=X |] ==> C <= Inter(A)"; +by (REPEAT (ares_tac [subsetI,InterI] 1 + ORELSE eresolve_tac (prems RL [subsetD]) 1)); +val Inter_greatest = result(); + +(*** Finite Union -- the least upper bound of 2 sets ***) + +goal Set.thy "A <= A Un B"; +by (REPEAT (ares_tac [subsetI,UnI1] 1)); +val Un_upper1 = result(); + +goal Set.thy "B <= A Un B"; +by (REPEAT (ares_tac [subsetI,UnI2] 1)); +val Un_upper2 = result(); + +val prems = goal Set.thy "[| A<=C; B<=C |] ==> A Un B <= C"; +by (cut_facts_tac prems 1); +by (DEPTH_SOLVE (ares_tac [subsetI] 1 + ORELSE eresolve_tac [UnE,subsetD] 1)); +val Un_least = result(); + +(*** Finite Intersection -- the greatest lower bound of 2 sets *) + +goal Set.thy "A Int B <= A"; +by (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)); +val Int_lower1 = result(); + +goal Set.thy "A Int B <= B"; +by (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)); +val Int_lower2 = result(); + +val prems = goal Set.thy "[| C<=A; C<=B |] ==> C <= A Int B"; +by (cut_facts_tac prems 1); +by (REPEAT (ares_tac [subsetI,IntI] 1 + ORELSE etac subsetD 1)); +val Int_greatest = result(); + +(*** Monotonicity ***) + +val [prem] = goalw Set.thy [mono_def] + "[| !!A B. A <= B ==> f(A) <= f(B) |] ==> mono(f)"; +by (REPEAT (ares_tac [allI, impI, prem] 1)); +val monoI = result(); + +val [major,minor] = goalw Set.thy [mono_def] + "[| mono(f); A <= B |] ==> f(A) <= f(B)"; +by (rtac (major RS spec RS spec RS mp) 1); +by (rtac minor 1); +val monoD = result(); + +val [prem] = goalw Set.thy [mono_def] "(!!x.f(x) = g(x)) ==> mono(f) <-> mono(g)"; +by (REPEAT (resolve_tac (iff_refl::FOL_congs @ [prem RS subst]) 1)); +val mono_cong = result(); + +val [prem] = goal Set.thy "mono(f) ==> f(A) Un f(B) <= f(A Un B)"; +by (rtac Un_least 1); +by (rtac (Un_upper1 RS (prem RS monoD)) 1); +by (rtac (Un_upper2 RS (prem RS monoD)) 1); +val mono_Un = result(); + +val [prem] = goal Set.thy "mono(f) ==> f(A Int B) <= f(A) Int f(B)"; +by (rtac Int_greatest 1); +by (rtac (Int_lower1 RS (prem RS monoD)) 1); +by (rtac (Int_lower2 RS (prem RS monoD)) 1); +val mono_Int = result(); + +(****) + +val set_cs = FOL_cs + addSIs [ballI, subsetI, InterI, INT_I, CollectI, + ComplI, IntI, UnCI, singletonI] + addIs [bexI, UnionI, UN_I] + addSEs [bexE, UnionE, UN_E, + CollectE, ComplE, IntE, UnE, emptyE, singletonE] + addEs [ballE, InterD, InterE, INT_D, INT_E, subsetD, subsetCE]; + +fun cfast_tac prems = cut_facts_tac prems THEN' fast_tac set_cs; + +fun prover s = prove_goal Set.thy s (fn _=>[fast_tac set_cs 1]); + +val mem_rews = [trivial_set,empty_eq] @ (map prover + [ "(a : A Un B) <-> (a:A | a:B)", + "(a : A Int B) <-> (a:A & a:B)", + "(a : Compl(B)) <-> (~a:B)", + "(a : {b}) <-> (a=b)", + "(a : {}) <-> False", + "(a : {x.P(x)}) <-> P(a)" ]); + +val set_congs = + [Collect_cong, ball_cong, bex_cong, INT_cong, UN_cong, mono_cong] @ + mk_congs Set.thy ["Compl", "op Int", "op Un", "Union", "Inter", + "op :", "op <=", "singleton"]; + +val set_ss = FOL_ss addcongs set_congs addrews mem_rews; diff -r 000000000000 -r a5a9c433f639 src/CCL/term.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/term.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,146 @@ +(* Title: CCL/terms + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For terms.thy. +*) + +open Term; + +val simp_can_defs = [one_def,inl_def,inr_def]; +val simp_ncan_defs = [if_def,when_def,split_def,fst_def,snd_def,thd_def]; +val simp_defs = simp_can_defs @ simp_ncan_defs; + +val ind_can_defs = [zero_def,succ_def,nil_def,cons_def]; +val ind_ncan_defs = [ncase_def,nrec_def,lcase_def,lrec_def]; +val ind_defs = ind_can_defs @ ind_ncan_defs; + +val data_defs = simp_defs @ ind_defs @ [napply_def]; +val genrec_defs = [letrec_def,letrec2_def,letrec3_def]; + +val term_congs = ccl_mk_congs Term.thy + ["inl","inr","succ","op .","split","if","when","ncase","nrec","lcase","lrec", + "fst","snd","thd","let","letrec","letrec2","letrec3","napply"]; + +(*** Beta Rules, including strictness ***) + +goalw Term.thy [let_def] "~ t=bot--> let x be t in f(x) = f(t)"; +by (res_inst_tac [("t","t")] term_case 1); +by (ALLGOALS(SIMP_TAC(CCL_ss addrews [caseBtrue,caseBfalse,caseBpair,caseBlam]))); +val letB = result() RS mp; + +goalw Term.thy [let_def] "let x be bot in f(x) = bot"; +br caseBbot 1; +val letBabot = result(); + +goalw Term.thy [let_def] "let x be t in bot = bot"; +brs ([caseBbot] RL [term_case]) 1; +by (ALLGOALS(SIMP_TAC(CCL_ss addrews [caseBtrue,caseBfalse,caseBpair,caseBlam]))); +val letBbbot = result(); + +goalw Term.thy [apply_def] "(lam x.b(x)) ` a = b(a)"; +by (ALLGOALS(SIMP_TAC(CCL_ss addrews [caseBtrue,caseBfalse,caseBpair,caseBlam]))); +val applyB = result(); + +goalw Term.thy [apply_def] "bot ` a = bot"; +br caseBbot 1; +val applyBbot = result(); + +goalw Term.thy [fix_def] "fix(f) = f(fix(f))"; +by (resolve_tac [applyB RS ssubst] 1 THEN resolve_tac [refl] 1); +val fixB = result(); + +goalw Term.thy [letrec_def] + "letrec g x be h(x,g) in g(a) = h(a,%y.letrec g x be h(x,g) in g(y))"; +by (resolve_tac [fixB RS ssubst] 1 THEN + resolve_tac [applyB RS ssubst] 1 THEN resolve_tac [refl] 1); +val letrecB = result(); + +val rawBs = caseBs @ [applyB,applyBbot,letrecB]; + +fun raw_mk_beta_rl defs s = prove_goalw Term.thy defs s + (fn _ => [SIMP_TAC (CCL_ss addrews rawBs addcongs term_congs) 1]); +fun mk_beta_rl s = raw_mk_beta_rl data_defs s; + +val ifBtrue = mk_beta_rl "if true then t else u = t"; +val ifBfalse = mk_beta_rl "if false then t else u = u"; +val ifBbot = mk_beta_rl "if bot then t else u = bot"; + +val whenBinl = mk_beta_rl "when(inl(a),t,u) = t(a)"; +val whenBinr = mk_beta_rl "when(inr(a),t,u) = u(a)"; +val whenBbot = mk_beta_rl "when(bot,t,u) = bot"; + +val splitB = mk_beta_rl "split(,h) = h(a,b)"; +val splitBbot = mk_beta_rl "split(bot,h) = bot"; +val fstB = mk_beta_rl "fst() = a"; +val fstBbot = mk_beta_rl "fst(bot) = bot"; +val sndB = mk_beta_rl "snd() = b"; +val sndBbot = mk_beta_rl "snd(bot) = bot"; +val thdB = mk_beta_rl "thd(>) = c"; +val thdBbot = mk_beta_rl "thd(bot) = bot"; + +val ncaseBzero = mk_beta_rl "ncase(zero,t,u) = t"; +val ncaseBsucc = mk_beta_rl "ncase(succ(n),t,u) = u(n)"; +val ncaseBbot = mk_beta_rl "ncase(bot,t,u) = bot"; +val nrecBzero = mk_beta_rl "nrec(zero,t,u) = t"; +val nrecBsucc = mk_beta_rl "nrec(succ(n),t,u) = u(n,nrec(n,t,u))"; +val nrecBbot = mk_beta_rl "nrec(bot,t,u) = bot"; + +val lcaseBnil = mk_beta_rl "lcase([],t,u) = t"; +val lcaseBcons = mk_beta_rl "lcase(x.xs,t,u) = u(x,xs)"; +val lcaseBbot = mk_beta_rl "lcase(bot,t,u) = bot"; +val lrecBnil = mk_beta_rl "lrec([],t,u) = t"; +val lrecBcons = mk_beta_rl "lrec(x.xs,t,u) = u(x,xs,lrec(xs,t,u))"; +val lrecBbot = mk_beta_rl "lrec(bot,t,u) = bot"; + +val letrec2B = raw_mk_beta_rl (data_defs @ [letrec2_def]) + "letrec g x y be h(x,y,g) in g(p,q) = \ +\ h(p,q,%u v.letrec g x y be h(x,y,g) in g(u,v))"; +val letrec3B = raw_mk_beta_rl (data_defs @ [letrec3_def]) + "letrec g x y z be h(x,y,z,g) in g(p,q,r) = \ +\ h(p,q,r,%u v w.letrec g x y z be h(x,y,z,g) in g(u,v,w))"; + +val napplyBzero = mk_beta_rl "f^zero`a = a"; +val napplyBsucc = mk_beta_rl "f^succ(n)`a = f(f^n`a)"; + +val termBs = [letB,applyB,applyBbot,splitB,splitBbot, + fstB,fstBbot,sndB,sndBbot,thdB,thdBbot, + ifBtrue,ifBfalse,ifBbot,whenBinl,whenBinr,whenBbot, + ncaseBzero,ncaseBsucc,ncaseBbot,nrecBzero,nrecBsucc,nrecBbot, + lcaseBnil,lcaseBcons,lcaseBbot,lrecBnil,lrecBcons,lrecBbot, + napplyBzero,napplyBsucc]; + +(*** Constructors are injective ***) + +val term_injs = map (mk_inj_rl Term.thy + [applyB,splitB,whenBinl,whenBinr,ncaseBsucc,lcaseBcons] + (ccl_congs @ term_congs)) + ["(inl(a) = inl(a')) <-> (a=a')", + "(inr(a) = inr(a')) <-> (a=a')", + "(succ(a) = succ(a')) <-> (a=a')", + "(a.b = a'.b') <-> (a=a' & b=b')"]; + +(*** Constructors are distinct ***) + +val term_dstncts = mkall_dstnct_thms Term.thy data_defs (ccl_injs @ term_injs) + [["bot","inl","inr"],["bot","zero","succ"],["bot","nil","op ."]]; + +(*** Rules for pre-order [= ***) + +local + fun mk_thm s = prove_goalw Term.thy data_defs s (fn _ => + [SIMP_TAC (ccl_ss addrews (ccl_porews)) 1]); +in + val term_porews = map mk_thm ["inl(a) [= inl(a') <-> a [= a'", + "inr(b) [= inr(b') <-> b [= b'", + "succ(n) [= succ(n') <-> n [= n'", + "x.xs [= x'.xs' <-> x [= x' & xs [= xs'"]; +end; + +(*** Rewriting and Proving ***) + +val term_rews = termBs @ term_injs @ term_dstncts @ ccl_porews @ term_porews; +val term_ss = ccl_ss addrews term_rews addcongs term_congs; + +val term_cs = ccl_cs addSEs (term_dstncts RL [notE]) addSDs (XH_to_Ds term_injs); diff -r 000000000000 -r a5a9c433f639 src/CCL/term.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/term.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,131 @@ +(* Title: CCL/terms.thy + ID: $Id$ + Author: Martin Coen + Copyright 1993 University of Cambridge + +Definitions of usual program constructs in CCL. + +*) + +Term = CCL + + +consts + + one :: "i" + + if :: "[i,i,i]=>i" ("(3if _/ then _/ else _)" [] 60) + + inl,inr :: "i=>i" + when :: "[i,i=>i,i=>i]=>i" + + split :: "[i,[i,i]=>i]=>i" + fst,snd, + thd :: "i=>i" + + zero :: "i" + succ :: "i=>i" + ncase :: "[i,i,i=>i]=>i" + nrec :: "[i,i,[i,i]=>i]=>i" + + nil :: "i" ("([])") + "." :: "[i,i]=>i" (infixr 80) + lcase :: "[i,i,[i,i]=>i]=>i" + lrec :: "[i,i,[i,i,i]=>i]=>i" + + let :: "[i,i=>i]=>i" + letrec :: "[[i,i=>i]=>i,(i=>i)=>i]=>i" + letrec2 :: "[[i,i,i=>i=>i]=>i,(i=>i=>i)=>i]=>i" + letrec3 :: "[[i,i,i,i=>i=>i=>i]=>i,(i=>i=>i=>i)=>i]=>i" + + "@let" :: "[id,i,i]=>i" ("(3let _ be _/ in _)" [] 60) + "@letrec" :: "[id,id,i,i]=>i" ("(3letrec _ _ be _/ in _)" [] 60) + "@letrec2" :: "[id,id,id,i,i]=>i" ("(3letrec _ _ _ be _/ in _)" [] 60) + "@letrec3" :: "[id,id,id,id,i,i]=>i" ("(3letrec _ _ _ _ be _/ in _)" [] 60) + + napply :: "[i=>i,i,i]=>i" ("(_ ^ _ ` _)") + +rules + + one_def "one == true" + if_def "if b then t else u == case(b,t,u,% x y.bot,%v.bot)" + inl_def "inl(a) == " + inr_def "inr(b) == " + when_def "when(t,f,g) == split(t,%b x.if b then f(x) else g(x))" + split_def "split(t,f) == case(t,bot,bot,f,%u.bot)" + fst_def "fst(t) == split(t,%x y.x)" + snd_def "snd(t) == split(t,%x y.y)" + thd_def "thd(t) == split(t,%x p.split(p,%y z.z))" + zero_def "zero == inl(one)" + succ_def "succ(n) == inr(n)" + ncase_def "ncase(n,b,c) == when(n,%x.b,%y.c(y))" + nrec_def " nrec(n,b,c) == letrec g x be ncase(x,b,%y.c(y,g(y))) in g(n)" + nil_def "[] == inl(one)" + cons_def "h.t == inr()" + lcase_def "lcase(l,b,c) == when(l,%x.b,%y.split(y,c))" + lrec_def "lrec(l,b,c) == letrec g x be lcase(x,b,%h t.c(h,t,g(t))) in g(l)" + + let_def "let x be t in f(x) == case(t,f(true),f(false),%x y.f(),%u.f(lam x.u(x)))" + letrec_def + "letrec g x be h(x,g) in b(g) == b(%x.fix(%f.lam x.h(x,%y.f`y))`x)" + + letrec2_def "letrec g x y be h(x,y,g) in f(g)== \ +\ letrec g' p be split(p,%x y.h(x,y,%u v.g'())) \ +\ in f(%x y.g'())" + + letrec3_def "letrec g x y z be h(x,y,z,g) in f(g) == \ +\ letrec g' p be split(p,%x xs.split(xs,%y z.h(x,y,z,%u v w.g'(>)))) \ +\ in f(%x y z.g'(>))" + + napply_def "f ^n` a == nrec(n,a,%x g.f(g))" + +end + +ML + +(** Quantifier translations: variable binding **) + +fun let_tr [Free(id,T),a,b] = Const("let",dummyT) $ a $ absfree(id,T,b); +fun let_tr' [a,Abs(id,T,b)] = + let val (id',b') = variant_abs(id,T,b) + in Const("@let",dummyT) $ Free(id',T) $ a $ b' end; + +fun letrec_tr [Free(f,S),Free(x,T),a,b] = + Const("letrec",dummyT) $ absfree(x,T,absfree(f,S,a)) $ absfree(f,S,b); +fun letrec2_tr [Free(f,S),Free(x,T),Free(y,U),a,b] = + Const("letrec2",dummyT) $ absfree(x,T,absfree(y,U,absfree(f,S,a))) $ absfree(f,S,b); +fun letrec3_tr [Free(f,S),Free(x,T),Free(y,U),Free(z,V),a,b] = + Const("letrec3",dummyT) $ absfree(x,T,absfree(y,U,absfree(z,U,absfree(f,S,a)))) $ absfree(f,S,b); + +fun letrec_tr' [Abs(x,T,Abs(f,S,a)),Abs(ff,SS,b)] = + let val (f',b') = variant_abs(ff,SS,b) + val (_,a'') = variant_abs(f,S,a) + val (x',a') = variant_abs(x,T,a'') + in Const("@letrec",dummyT) $ Free(f',SS) $ Free(x',T) $ a' $ b' end; +fun letrec2_tr' [Abs(x,T,Abs(y,U,Abs(f,S,a))),Abs(ff,SS,b)] = + let val (f',b') = variant_abs(ff,SS,b) + val ( _,a1) = variant_abs(f,S,a) + val (y',a2) = variant_abs(y,U,a1) + val (x',a') = variant_abs(x,T,a2) + in Const("@letrec2",dummyT) $ Free(f',SS) $ Free(x',T) $ Free(y',U) $ a' $ b' + end; +fun letrec3_tr' [Abs(x,T,Abs(y,U,Abs(z,V,Abs(f,S,a)))),Abs(ff,SS,b)] = + let val (f',b') = variant_abs(ff,SS,b) + val ( _,a1) = variant_abs(f,S,a) + val (z',a2) = variant_abs(z,V,a1) + val (y',a3) = variant_abs(y,U,a2) + val (x',a') = variant_abs(x,T,a3) + in Const("@letrec3",dummyT) $ Free(f',SS) $ Free(x',T) $ Free(y',U) $ Free(z',V) $ a' $ b' + end; + +val parse_translation= + [("@let", let_tr), + ("@letrec", letrec_tr), + ("@letrec2", letrec2_tr), + ("@letrec3", letrec3_tr) + ]; +val print_translation= + [("let", let_tr'), + ("letrec", letrec_tr'), + ("letrec2", letrec2_tr'), + ("letrec3", letrec3_tr') + ]; diff -r 000000000000 -r a5a9c433f639 src/CCL/terms.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/terms.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,146 @@ +(* Title: CCL/terms + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For terms.thy. +*) + +open Term; + +val simp_can_defs = [one_def,inl_def,inr_def]; +val simp_ncan_defs = [if_def,when_def,split_def,fst_def,snd_def,thd_def]; +val simp_defs = simp_can_defs @ simp_ncan_defs; + +val ind_can_defs = [zero_def,succ_def,nil_def,cons_def]; +val ind_ncan_defs = [ncase_def,nrec_def,lcase_def,lrec_def]; +val ind_defs = ind_can_defs @ ind_ncan_defs; + +val data_defs = simp_defs @ ind_defs @ [napply_def]; +val genrec_defs = [letrec_def,letrec2_def,letrec3_def]; + +val term_congs = ccl_mk_congs Term.thy + ["inl","inr","succ","op .","split","if","when","ncase","nrec","lcase","lrec", + "fst","snd","thd","let","letrec","letrec2","letrec3","napply"]; + +(*** Beta Rules, including strictness ***) + +goalw Term.thy [let_def] "~ t=bot--> let x be t in f(x) = f(t)"; +by (res_inst_tac [("t","t")] term_case 1); +by (ALLGOALS(SIMP_TAC(CCL_ss addrews [caseBtrue,caseBfalse,caseBpair,caseBlam]))); +val letB = result() RS mp; + +goalw Term.thy [let_def] "let x be bot in f(x) = bot"; +br caseBbot 1; +val letBabot = result(); + +goalw Term.thy [let_def] "let x be t in bot = bot"; +brs ([caseBbot] RL [term_case]) 1; +by (ALLGOALS(SIMP_TAC(CCL_ss addrews [caseBtrue,caseBfalse,caseBpair,caseBlam]))); +val letBbbot = result(); + +goalw Term.thy [apply_def] "(lam x.b(x)) ` a = b(a)"; +by (ALLGOALS(SIMP_TAC(CCL_ss addrews [caseBtrue,caseBfalse,caseBpair,caseBlam]))); +val applyB = result(); + +goalw Term.thy [apply_def] "bot ` a = bot"; +br caseBbot 1; +val applyBbot = result(); + +goalw Term.thy [fix_def] "fix(f) = f(fix(f))"; +by (resolve_tac [applyB RS ssubst] 1 THEN resolve_tac [refl] 1); +val fixB = result(); + +goalw Term.thy [letrec_def] + "letrec g x be h(x,g) in g(a) = h(a,%y.letrec g x be h(x,g) in g(y))"; +by (resolve_tac [fixB RS ssubst] 1 THEN + resolve_tac [applyB RS ssubst] 1 THEN resolve_tac [refl] 1); +val letrecB = result(); + +val rawBs = caseBs @ [applyB,applyBbot,letrecB]; + +fun raw_mk_beta_rl defs s = prove_goalw Term.thy defs s + (fn _ => [SIMP_TAC (CCL_ss addrews rawBs addcongs term_congs) 1]); +fun mk_beta_rl s = raw_mk_beta_rl data_defs s; + +val ifBtrue = mk_beta_rl "if true then t else u = t"; +val ifBfalse = mk_beta_rl "if false then t else u = u"; +val ifBbot = mk_beta_rl "if bot then t else u = bot"; + +val whenBinl = mk_beta_rl "when(inl(a),t,u) = t(a)"; +val whenBinr = mk_beta_rl "when(inr(a),t,u) = u(a)"; +val whenBbot = mk_beta_rl "when(bot,t,u) = bot"; + +val splitB = mk_beta_rl "split(,h) = h(a,b)"; +val splitBbot = mk_beta_rl "split(bot,h) = bot"; +val fstB = mk_beta_rl "fst() = a"; +val fstBbot = mk_beta_rl "fst(bot) = bot"; +val sndB = mk_beta_rl "snd() = b"; +val sndBbot = mk_beta_rl "snd(bot) = bot"; +val thdB = mk_beta_rl "thd(>) = c"; +val thdBbot = mk_beta_rl "thd(bot) = bot"; + +val ncaseBzero = mk_beta_rl "ncase(zero,t,u) = t"; +val ncaseBsucc = mk_beta_rl "ncase(succ(n),t,u) = u(n)"; +val ncaseBbot = mk_beta_rl "ncase(bot,t,u) = bot"; +val nrecBzero = mk_beta_rl "nrec(zero,t,u) = t"; +val nrecBsucc = mk_beta_rl "nrec(succ(n),t,u) = u(n,nrec(n,t,u))"; +val nrecBbot = mk_beta_rl "nrec(bot,t,u) = bot"; + +val lcaseBnil = mk_beta_rl "lcase([],t,u) = t"; +val lcaseBcons = mk_beta_rl "lcase(x.xs,t,u) = u(x,xs)"; +val lcaseBbot = mk_beta_rl "lcase(bot,t,u) = bot"; +val lrecBnil = mk_beta_rl "lrec([],t,u) = t"; +val lrecBcons = mk_beta_rl "lrec(x.xs,t,u) = u(x,xs,lrec(xs,t,u))"; +val lrecBbot = mk_beta_rl "lrec(bot,t,u) = bot"; + +val letrec2B = raw_mk_beta_rl (data_defs @ [letrec2_def]) + "letrec g x y be h(x,y,g) in g(p,q) = \ +\ h(p,q,%u v.letrec g x y be h(x,y,g) in g(u,v))"; +val letrec3B = raw_mk_beta_rl (data_defs @ [letrec3_def]) + "letrec g x y z be h(x,y,z,g) in g(p,q,r) = \ +\ h(p,q,r,%u v w.letrec g x y z be h(x,y,z,g) in g(u,v,w))"; + +val napplyBzero = mk_beta_rl "f^zero`a = a"; +val napplyBsucc = mk_beta_rl "f^succ(n)`a = f(f^n`a)"; + +val termBs = [letB,applyB,applyBbot,splitB,splitBbot, + fstB,fstBbot,sndB,sndBbot,thdB,thdBbot, + ifBtrue,ifBfalse,ifBbot,whenBinl,whenBinr,whenBbot, + ncaseBzero,ncaseBsucc,ncaseBbot,nrecBzero,nrecBsucc,nrecBbot, + lcaseBnil,lcaseBcons,lcaseBbot,lrecBnil,lrecBcons,lrecBbot, + napplyBzero,napplyBsucc]; + +(*** Constructors are injective ***) + +val term_injs = map (mk_inj_rl Term.thy + [applyB,splitB,whenBinl,whenBinr,ncaseBsucc,lcaseBcons] + (ccl_congs @ term_congs)) + ["(inl(a) = inl(a')) <-> (a=a')", + "(inr(a) = inr(a')) <-> (a=a')", + "(succ(a) = succ(a')) <-> (a=a')", + "(a.b = a'.b') <-> (a=a' & b=b')"]; + +(*** Constructors are distinct ***) + +val term_dstncts = mkall_dstnct_thms Term.thy data_defs (ccl_injs @ term_injs) + [["bot","inl","inr"],["bot","zero","succ"],["bot","nil","op ."]]; + +(*** Rules for pre-order [= ***) + +local + fun mk_thm s = prove_goalw Term.thy data_defs s (fn _ => + [SIMP_TAC (ccl_ss addrews (ccl_porews)) 1]); +in + val term_porews = map mk_thm ["inl(a) [= inl(a') <-> a [= a'", + "inr(b) [= inr(b') <-> b [= b'", + "succ(n) [= succ(n') <-> n [= n'", + "x.xs [= x'.xs' <-> x [= x' & xs [= xs'"]; +end; + +(*** Rewriting and Proving ***) + +val term_rews = termBs @ term_injs @ term_dstncts @ ccl_porews @ term_porews; +val term_ss = ccl_ss addrews term_rews addcongs term_congs; + +val term_cs = ccl_cs addSEs (term_dstncts RL [notE]) addSDs (XH_to_Ds term_injs); diff -r 000000000000 -r a5a9c433f639 src/CCL/terms.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/terms.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,131 @@ +(* Title: CCL/terms.thy + ID: $Id$ + Author: Martin Coen + Copyright 1993 University of Cambridge + +Definitions of usual program constructs in CCL. + +*) + +Term = CCL + + +consts + + one :: "i" + + if :: "[i,i,i]=>i" ("(3if _/ then _/ else _)" [] 60) + + inl,inr :: "i=>i" + when :: "[i,i=>i,i=>i]=>i" + + split :: "[i,[i,i]=>i]=>i" + fst,snd, + thd :: "i=>i" + + zero :: "i" + succ :: "i=>i" + ncase :: "[i,i,i=>i]=>i" + nrec :: "[i,i,[i,i]=>i]=>i" + + nil :: "i" ("([])") + "." :: "[i,i]=>i" (infixr 80) + lcase :: "[i,i,[i,i]=>i]=>i" + lrec :: "[i,i,[i,i,i]=>i]=>i" + + let :: "[i,i=>i]=>i" + letrec :: "[[i,i=>i]=>i,(i=>i)=>i]=>i" + letrec2 :: "[[i,i,i=>i=>i]=>i,(i=>i=>i)=>i]=>i" + letrec3 :: "[[i,i,i,i=>i=>i=>i]=>i,(i=>i=>i=>i)=>i]=>i" + + "@let" :: "[id,i,i]=>i" ("(3let _ be _/ in _)" [] 60) + "@letrec" :: "[id,id,i,i]=>i" ("(3letrec _ _ be _/ in _)" [] 60) + "@letrec2" :: "[id,id,id,i,i]=>i" ("(3letrec _ _ _ be _/ in _)" [] 60) + "@letrec3" :: "[id,id,id,id,i,i]=>i" ("(3letrec _ _ _ _ be _/ in _)" [] 60) + + napply :: "[i=>i,i,i]=>i" ("(_ ^ _ ` _)") + +rules + + one_def "one == true" + if_def "if b then t else u == case(b,t,u,% x y.bot,%v.bot)" + inl_def "inl(a) == " + inr_def "inr(b) == " + when_def "when(t,f,g) == split(t,%b x.if b then f(x) else g(x))" + split_def "split(t,f) == case(t,bot,bot,f,%u.bot)" + fst_def "fst(t) == split(t,%x y.x)" + snd_def "snd(t) == split(t,%x y.y)" + thd_def "thd(t) == split(t,%x p.split(p,%y z.z))" + zero_def "zero == inl(one)" + succ_def "succ(n) == inr(n)" + ncase_def "ncase(n,b,c) == when(n,%x.b,%y.c(y))" + nrec_def " nrec(n,b,c) == letrec g x be ncase(x,b,%y.c(y,g(y))) in g(n)" + nil_def "[] == inl(one)" + cons_def "h.t == inr()" + lcase_def "lcase(l,b,c) == when(l,%x.b,%y.split(y,c))" + lrec_def "lrec(l,b,c) == letrec g x be lcase(x,b,%h t.c(h,t,g(t))) in g(l)" + + let_def "let x be t in f(x) == case(t,f(true),f(false),%x y.f(),%u.f(lam x.u(x)))" + letrec_def + "letrec g x be h(x,g) in b(g) == b(%x.fix(%f.lam x.h(x,%y.f`y))`x)" + + letrec2_def "letrec g x y be h(x,y,g) in f(g)== \ +\ letrec g' p be split(p,%x y.h(x,y,%u v.g'())) \ +\ in f(%x y.g'())" + + letrec3_def "letrec g x y z be h(x,y,z,g) in f(g) == \ +\ letrec g' p be split(p,%x xs.split(xs,%y z.h(x,y,z,%u v w.g'(>)))) \ +\ in f(%x y z.g'(>))" + + napply_def "f ^n` a == nrec(n,a,%x g.f(g))" + +end + +ML + +(** Quantifier translations: variable binding **) + +fun let_tr [Free(id,T),a,b] = Const("let",dummyT) $ a $ absfree(id,T,b); +fun let_tr' [a,Abs(id,T,b)] = + let val (id',b') = variant_abs(id,T,b) + in Const("@let",dummyT) $ Free(id',T) $ a $ b' end; + +fun letrec_tr [Free(f,S),Free(x,T),a,b] = + Const("letrec",dummyT) $ absfree(x,T,absfree(f,S,a)) $ absfree(f,S,b); +fun letrec2_tr [Free(f,S),Free(x,T),Free(y,U),a,b] = + Const("letrec2",dummyT) $ absfree(x,T,absfree(y,U,absfree(f,S,a))) $ absfree(f,S,b); +fun letrec3_tr [Free(f,S),Free(x,T),Free(y,U),Free(z,V),a,b] = + Const("letrec3",dummyT) $ absfree(x,T,absfree(y,U,absfree(z,U,absfree(f,S,a)))) $ absfree(f,S,b); + +fun letrec_tr' [Abs(x,T,Abs(f,S,a)),Abs(ff,SS,b)] = + let val (f',b') = variant_abs(ff,SS,b) + val (_,a'') = variant_abs(f,S,a) + val (x',a') = variant_abs(x,T,a'') + in Const("@letrec",dummyT) $ Free(f',SS) $ Free(x',T) $ a' $ b' end; +fun letrec2_tr' [Abs(x,T,Abs(y,U,Abs(f,S,a))),Abs(ff,SS,b)] = + let val (f',b') = variant_abs(ff,SS,b) + val ( _,a1) = variant_abs(f,S,a) + val (y',a2) = variant_abs(y,U,a1) + val (x',a') = variant_abs(x,T,a2) + in Const("@letrec2",dummyT) $ Free(f',SS) $ Free(x',T) $ Free(y',U) $ a' $ b' + end; +fun letrec3_tr' [Abs(x,T,Abs(y,U,Abs(z,V,Abs(f,S,a)))),Abs(ff,SS,b)] = + let val (f',b') = variant_abs(ff,SS,b) + val ( _,a1) = variant_abs(f,S,a) + val (z',a2) = variant_abs(z,V,a1) + val (y',a3) = variant_abs(y,U,a2) + val (x',a') = variant_abs(x,T,a3) + in Const("@letrec3",dummyT) $ Free(f',SS) $ Free(x',T) $ Free(y',U) $ Free(z',V) $ a' $ b' + end; + +val parse_translation= + [("@let", let_tr), + ("@letrec", letrec_tr), + ("@letrec2", letrec2_tr), + ("@letrec3", letrec3_tr) + ]; +val print_translation= + [("let", let_tr'), + ("letrec", letrec_tr'), + ("letrec2", letrec2_tr'), + ("letrec3", letrec3_tr') + ]; diff -r 000000000000 -r a5a9c433f639 src/CCL/trancl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/trancl.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,215 @@ +(* Title: CCL/trancl + ID: $Id$ + +For trancl.thy. + +Modified version of + Title: HOL/trancl.ML + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +*) + +open Trancl; + +(** Natural deduction for trans(r) **) + +val prems = goalw Trancl.thy [trans_def] + "(!! x y z. [| :r; :r |] ==> :r) ==> trans(r)"; +by (REPEAT (ares_tac (prems@[allI,impI]) 1)); +val transI = result(); + +val major::prems = goalw Trancl.thy [trans_def] + "[| trans(r); :r; :r |] ==> :r"; +by (cut_facts_tac [major] 1); +by (fast_tac (FOL_cs addIs prems) 1); +val transD = result(); + +(** Identity relation **) + +goalw Trancl.thy [id_def] " : id"; +by (rtac CollectI 1); +by (rtac exI 1); +by (rtac refl 1); +val idI = result(); + +val major::prems = goalw Trancl.thy [id_def] + "[| p: id; !!x.[| p = |] ==> P \ +\ |] ==> P"; +by (rtac (major RS CollectE) 1); +by (etac exE 1); +by (eresolve_tac prems 1); +val idE = result(); + +(** Composition of two relations **) + +val prems = goalw Trancl.thy [comp_def] + "[| :s; :r |] ==> : r O s"; +by (fast_tac (set_cs addIs prems) 1); +val compI = result(); + +(*proof requires higher-level assumptions or a delaying of hyp_subst_tac*) +val prems = goalw Trancl.thy [comp_def] + "[| xz : r O s; \ +\ !!x y z. [| xz = ; :s; :r |] ==> P \ +\ |] ==> P"; +by (cut_facts_tac prems 1); +by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1)); +val compE = result(); + +val prems = goal Trancl.thy + "[| : r O s; \ +\ !!y. [| :s; :r |] ==> P \ +\ |] ==> P"; +by (rtac compE 1); +by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [pair_inject,ssubst] 1)); +val compEpair = result(); + +val comp_cs = set_cs addIs [compI,idI] + addEs [compE,idE] + addSEs [pair_inject]; + +val prems = goal Trancl.thy + "[| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; +by (cut_facts_tac prems 1); +by (fast_tac comp_cs 1); +val comp_mono = result(); + +(** The relation rtrancl **) + +goal Trancl.thy "mono(%s. id Un (r O s))"; +by (rtac monoI 1); +by (REPEAT (ares_tac [monoI, subset_refl, comp_mono, Un_mono] 1)); +val rtrancl_fun_mono = result(); + +val rtrancl_unfold = rtrancl_fun_mono RS (rtrancl_def RS def_lfp_Tarski); + +(*Reflexivity of rtrancl*) +goal Trancl.thy " : r^*"; +br (rtrancl_unfold RS ssubst) 1; +by (fast_tac comp_cs 1); +val rtrancl_refl = result(); + +(*Closure under composition with r*) +val prems = goal Trancl.thy + "[| : r^*; : r |] ==> : r^*"; +br (rtrancl_unfold RS ssubst) 1; +by (fast_tac (comp_cs addIs prems) 1); +val rtrancl_into_rtrancl = result(); + +(*rtrancl of r contains r*) +val [prem] = goal Trancl.thy "[| : r |] ==> : r^*"; +by (rtac (rtrancl_refl RS rtrancl_into_rtrancl) 1); +by (rtac prem 1); +val r_into_rtrancl = result(); + + +(** standard induction rule **) + +val major::prems = goal Trancl.thy + "[| : r^*; \ +\ !!x. P(); \ +\ !!x y z.[| P(); : r^*; : r |] ==> P() |] \ +\ ==> P()"; +by (rtac (major RS (rtrancl_def RS def_induct)) 1); +by (rtac rtrancl_fun_mono 1); +by (fast_tac (comp_cs addIs prems) 1); +val rtrancl_full_induct = result(); + +(*nice induction rule*) +val major::prems = goal Trancl.thy + "[| : r^*; \ +\ P(a); \ +\ !!y z.[| : r^*; : r; P(y) |] ==> P(z) |] \ +\ ==> P(b)"; +(*by induction on this formula*) +by (subgoal_tac "ALL y. = --> P(y)" 1); +(*now solve first subgoal: this formula is sufficient*) +by (fast_tac FOL_cs 1); +(*now do the induction*) +by (resolve_tac [major RS rtrancl_full_induct] 1); +by (fast_tac (comp_cs addIs prems) 1); +by (fast_tac (comp_cs addIs prems) 1); +val rtrancl_induct = result(); + +(*transitivity of transitive closure!! -- by induction.*) +goal Trancl.thy "trans(r^*)"; +by (rtac transI 1); +by (res_inst_tac [("b","z")] rtrancl_induct 1); +by (DEPTH_SOLVE (eresolve_tac [asm_rl, rtrancl_into_rtrancl] 1)); +val trans_rtrancl = result(); + +(*elimination of rtrancl -- by induction on a special formula*) +val major::prems = goal Trancl.thy + "[| : r^*; (a = b) ==> P; \ +\ !!y.[| : r^*; : r |] ==> P |] \ +\ ==> P"; +by (subgoal_tac "a = b | (EX y. : r^* & : r)" 1); +by (rtac (major RS rtrancl_induct) 2); +by (fast_tac (set_cs addIs prems) 2); +by (fast_tac (set_cs addIs prems) 2); +by (REPEAT (eresolve_tac ([asm_rl,exE,disjE,conjE]@prems) 1)); +val rtranclE = result(); + + +(**** The relation trancl ****) + +(** Conversions between trancl and rtrancl **) + +val [major] = goalw Trancl.thy [trancl_def] + "[| : r^+ |] ==> : r^*"; +by (resolve_tac [major RS compEpair] 1); +by (REPEAT (ares_tac [rtrancl_into_rtrancl] 1)); +val trancl_into_rtrancl = result(); + +(*r^+ contains r*) +val [prem] = goalw Trancl.thy [trancl_def] + "[| : r |] ==> : r^+"; +by (REPEAT (ares_tac [prem,compI,rtrancl_refl] 1)); +val r_into_trancl = result(); + +(*intro rule by definition: from rtrancl and r*) +val prems = goalw Trancl.thy [trancl_def] + "[| : r^*; : r |] ==> : r^+"; +by (REPEAT (resolve_tac ([compI]@prems) 1)); +val rtrancl_into_trancl1 = result(); + +(*intro rule from r and rtrancl*) +val prems = goal Trancl.thy + "[| : r; : r^* |] ==> : r^+"; +by (resolve_tac (prems RL [rtranclE]) 1); +by (etac subst 1); +by (resolve_tac (prems RL [r_into_trancl]) 1); +by (rtac (trans_rtrancl RS transD RS rtrancl_into_trancl1) 1); +by (REPEAT (ares_tac (prems@[r_into_rtrancl]) 1)); +val rtrancl_into_trancl2 = result(); + +(*elimination of r^+ -- NOT an induction rule*) +val major::prems = goal Trancl.thy + "[| : r^+; \ +\ : r ==> P; \ +\ !!y.[| : r^+; : r |] ==> P \ +\ |] ==> P"; +by (subgoal_tac " : r | (EX y. : r^+ & : r)" 1); +by (REPEAT (eresolve_tac ([asm_rl,disjE,exE,conjE]@prems) 1)); +by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); +by (etac rtranclE 1); +by (fast_tac comp_cs 1); +by (fast_tac (comp_cs addSIs [rtrancl_into_trancl1]) 1); +val tranclE = result(); + +(*Transitivity of r^+. + Proved by unfolding since it uses transitivity of rtrancl. *) +goalw Trancl.thy [trancl_def] "trans(r^+)"; +by (rtac transI 1); +by (REPEAT (etac compEpair 1)); +by (rtac (rtrancl_into_rtrancl RS (trans_rtrancl RS transD RS compI)) 1); +by (REPEAT (assume_tac 1)); +val trans_trancl = result(); + +val prems = goal Trancl.thy + "[| : r; : r^+ |] ==> : r^+"; +by (rtac (r_into_trancl RS (trans_trancl RS transD)) 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +val trancl_into_trancl2 = result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/trancl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/trancl.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,28 @@ +(* Title: CCL/trancl.thy + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Transitive closure of a relation +*) + +Trancl = CCL + + +consts + trans :: "i set => o" (*transitivity predicate*) + id :: "i set" + rtrancl :: "i set => i set" ("(_^*)" [100] 100) + trancl :: "i set => i set" ("(_^+)" [100] 100) + O :: "[i set,i set] => i set" (infixr 60) + +rules + +trans_def "trans(r) == (ALL x y z. :r --> :r --> :r)" +comp_def (*composition of relations*) + "r O s == {xz. EX x y z. xz = & :s & :r}" +id_def (*the identity relation*) + "id == {p. EX x. p = }" +rtrancl_def "r^* == lfp(%s. id Un (r O s))" +trancl_def "r^+ == r O rtrancl(r)" + +end diff -r 000000000000 -r a5a9c433f639 src/CCL/type.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/type.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,308 @@ +(* Title: CCL/types + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For types.thy. +*) + +open Type; + +val simp_type_defs = [Subtype_def,Unit_def,Bool_def,Plus_def,Sigma_def,Pi_def, + Lift_def,Tall_def,Tex_def]; +val ind_type_defs = [Nat_def,List_def]; + +val simp_data_defs = [one_def,inl_def,inr_def]; +val ind_data_defs = [zero_def,succ_def,nil_def,cons_def]; + +goal Set.thy "A <= B <-> (ALL x.x:A --> x:B)"; +by (fast_tac set_cs 1); +val subsetXH = result(); + +(*** Exhaustion Rules ***) + +fun mk_XH_tac thy defs rls s = prove_goalw thy defs s (fn _ => [cfast_tac rls 1]); +val XH_tac = mk_XH_tac Type.thy simp_type_defs []; + +val EmptyXH = XH_tac "a : {} <-> False"; +val SubtypeXH = XH_tac "a : {x:A.P(x)} <-> (a:A & P(a))"; +val UnitXH = XH_tac "a : Unit <-> a=one"; +val BoolXH = XH_tac "a : Bool <-> a=true | a=false"; +val PlusXH = XH_tac "a : A+B <-> (EX x:A.a=inl(x)) | (EX x:B.a=inr(x))"; +val PiXH = XH_tac "a : PROD x:A.B(x) <-> (EX b.a=lam x.b(x) & (ALL x:A.b(x):B(x)))"; +val SgXH = XH_tac "a : SUM x:A.B(x) <-> (EX x:A.EX y:B(x).a=)"; + +val XHs = [EmptyXH,SubtypeXH,UnitXH,BoolXH,PlusXH,PiXH,SgXH]; + +val LiftXH = XH_tac "a : [A] <-> (a=bot | a:A)"; +val TallXH = XH_tac "a : TALL X.B(X) <-> (ALL X. a:B(X))"; +val TexXH = XH_tac "a : TEX X.B(X) <-> (EX X. a:B(X))"; + +val case_rls = XH_to_Es XHs; + +(*** Canonical Type Rules ***) + +fun mk_canT_tac thy xhs s = prove_goal thy s + (fn prems => [fast_tac (set_cs addIs (prems @ (xhs RL [iffD2]))) 1]); +val canT_tac = mk_canT_tac Type.thy XHs; + +val oneT = canT_tac "one : Unit"; +val trueT = canT_tac "true : Bool"; +val falseT = canT_tac "false : Bool"; +val lamT = canT_tac "[| !!x.x:A ==> b(x):B(x) |] ==> lam x.b(x) : Pi(A,B)"; +val pairT = canT_tac "[| a:A; b:B(a) |] ==> :Sigma(A,B)"; +val inlT = canT_tac "a:A ==> inl(a) : A+B"; +val inrT = canT_tac "b:B ==> inr(b) : A+B"; + +val canTs = [oneT,trueT,falseT,pairT,lamT,inlT,inrT]; + +(*** Non-Canonical Type Rules ***) + +local +val lemma = prove_goal Type.thy "[| a:B(u); u=v |] ==> a : B(v)" + (fn prems => [cfast_tac prems 1]); +in +fun mk_ncanT_tac thy defs top_crls crls s = prove_goalw thy defs s + (fn major::prems => [(resolve_tac ([major] RL top_crls) 1), + (REPEAT_SOME (eresolve_tac (crls @ [exE,bexE,conjE,disjE]))), + (ALLGOALS (ASM_SIMP_TAC term_ss)), + (ALLGOALS (ares_tac (prems RL [lemma]) ORELSE' + eresolve_tac [bspec])), + (safe_tac (ccl_cs addSIs prems))]); +end; + +val ncanT_tac = mk_ncanT_tac Type.thy [] case_rls case_rls; + +val ifT = ncanT_tac + "[| b:Bool; b=true ==> t:A(true); b=false ==> u:A(false) |] ==> \ +\ if b then t else u : A(b)"; + +val applyT = ncanT_tac + "[| f : Pi(A,B); a:A |] ==> f ` a : B(a)"; + +val splitT = ncanT_tac + "[| p:Sigma(A,B); !!x y. [| x:A; y:B(x); p= |] ==> c(x,y):C() |] ==> \ +\ split(p,c):C(p)"; + +val whenT = ncanT_tac + "[| p:A+B; !!x.[| x:A; p=inl(x) |] ==> a(x):C(inl(x)); \ +\ !!y.[| y:B; p=inr(y) |] ==> b(y):C(inr(y)) |] ==> \ +\ when(p,a,b) : C(p)"; + +val ncanTs = [ifT,applyT,splitT,whenT]; + +(*** Subtypes ***) + +val SubtypeD1 = standard ((SubtypeXH RS iffD1) RS conjunct1); +val SubtypeD2 = standard ((SubtypeXH RS iffD1) RS conjunct2); + +val prems = goal Type.thy + "[| a:A; P(a) |] ==> a : {x:A. P(x)}"; +by (REPEAT (resolve_tac (prems@[SubtypeXH RS iffD2,conjI]) 1)); +val SubtypeI = result(); + +val prems = goal Type.thy + "[| a : {x:A. P(x)}; [| a:A; P(a) |] ==> Q |] ==> Q"; +by (REPEAT (resolve_tac (prems@[SubtypeD1,SubtypeD2]) 1)); +val SubtypeE = result(); + +(*** Monotonicity ***) + +goal Type.thy "mono (%X.X)"; +by (REPEAT (ares_tac [monoI] 1)); +val idM = result(); + +goal Type.thy "mono(%X.A)"; +by (REPEAT (ares_tac [monoI,subset_refl] 1)); +val constM = result(); + +val major::prems = goal Type.thy + "mono(%X.A(X)) ==> mono(%X.[A(X)])"; +br (subsetI RS monoI) 1; +bd (LiftXH RS iffD1) 1; +be disjE 1; +be (disjI1 RS (LiftXH RS iffD2)) 1; +br (disjI2 RS (LiftXH RS iffD2)) 1; +be (major RS monoD RS subsetD) 1; +ba 1; +val LiftM = result(); + +val prems = goal Type.thy + "[| mono(%X.A(X)); !!x X. x:A(X) ==> mono(%X.B(X,x)) |] ==> \ +\ mono(%X.Sigma(A(X),B(X)))"; +by (REPEAT (ares_tac ([subsetI RS monoI] @ canTs) 1 ORELSE + eresolve_tac ([bspec,exE,conjE,disjE,bexE] @ case_rls) 1 ORELSE + (resolve_tac (prems RL [monoD RS subsetD]) 1 THEN assume_tac 1) ORELSE + hyp_subst_tac 1)); +val SgM = result(); + +val prems = goal Type.thy + "[| !!x. x:A ==> mono(%X.B(X,x)) |] ==> mono(%X.Pi(A,B(X)))"; +by (REPEAT (ares_tac ([subsetI RS monoI] @ canTs) 1 ORELSE + eresolve_tac ([bspec,exE,conjE,disjE,bexE] @ case_rls) 1 ORELSE + (resolve_tac (prems RL [monoD RS subsetD]) 1 THEN assume_tac 1) ORELSE + hyp_subst_tac 1)); +val PiM = result(); + +val prems = goal Type.thy + "[| mono(%X.A(X)); mono(%X.B(X)) |] ==> mono(%X.A(X)+B(X))"; +by (REPEAT (ares_tac ([subsetI RS monoI] @ canTs) 1 ORELSE + eresolve_tac ([bspec,exE,conjE,disjE,bexE] @ case_rls) 1 ORELSE + (resolve_tac (prems RL [monoD RS subsetD]) 1 THEN assume_tac 1) ORELSE + hyp_subst_tac 1)); +val PlusM = result(); + +(**************** RECURSIVE TYPES ******************) + +(*** Conversion Rules for Fixed Points via monotonicity and Tarski ***) + +goal Type.thy "mono(%X.Unit+X)"; +by (REPEAT (ares_tac [PlusM,constM,idM] 1)); +val NatM = result(); +val def_NatB = result() RS (Nat_def RS def_lfp_Tarski); + +goal Type.thy "mono(%X.(Unit+Sigma(A,%y.X)))"; +by (REPEAT (ares_tac [PlusM,SgM,constM,idM] 1)); +val ListM = result(); +val def_ListB = result() RS (List_def RS def_lfp_Tarski); +val def_ListsB = result() RS (Lists_def RS def_gfp_Tarski); + +goal Type.thy "mono(%X.({} + Sigma(A,%y.X)))"; +by (REPEAT (ares_tac [PlusM,SgM,constM,idM] 1)); +val IListsM = result(); +val def_IListsB = result() RS (ILists_def RS def_gfp_Tarski); + +val ind_type_eqs = [def_NatB,def_ListB,def_ListsB,def_IListsB]; + +(*** Exhaustion Rules ***) + +fun mk_iXH_tac teqs ddefs rls s = prove_goalw Type.thy ddefs s + (fn _ => [resolve_tac (teqs RL [XHlemma1]) 1, + fast_tac (set_cs addSIs canTs addSEs case_rls) 1]); + +val iXH_tac = mk_iXH_tac ind_type_eqs ind_data_defs []; + +val NatXH = iXH_tac "a : Nat <-> (a=zero | (EX x:Nat.a=succ(x)))"; +val ListXH = iXH_tac "a : List(A) <-> (a=[] | (EX x:A.EX xs:List(A).a=x.xs))"; +val ListsXH = iXH_tac "a : Lists(A) <-> (a=[] | (EX x:A.EX xs:Lists(A).a=x.xs))"; +val IListsXH = iXH_tac "a : ILists(A) <-> (EX x:A.EX xs:ILists(A).a=x.xs)"; + +val iXHs = [NatXH,ListXH]; +val icase_rls = XH_to_Es iXHs; + +(*** Type Rules ***) + +val icanT_tac = mk_canT_tac Type.thy iXHs; +val incanT_tac = mk_ncanT_tac Type.thy [] icase_rls case_rls; + +val zeroT = icanT_tac "zero : Nat"; +val succT = icanT_tac "n:Nat ==> succ(n) : Nat"; +val nilT = icanT_tac "[] : List(A)"; +val consT = icanT_tac "[| h:A; t:List(A) |] ==> h.t : List(A)"; + +val icanTs = [zeroT,succT,nilT,consT]; + +val ncaseT = incanT_tac + "[| n:Nat; n=zero ==> b:C(zero); \ +\ !!x.[| x:Nat; n=succ(x) |] ==> c(x):C(succ(x)) |] ==> \ +\ ncase(n,b,c) : C(n)"; + +val lcaseT = incanT_tac + "[| l:List(A); l=[] ==> b:C([]); \ +\ !!h t.[| h:A; t:List(A); l=h.t |] ==> c(h,t):C(h.t) |] ==> \ +\ lcase(l,b,c) : C(l)"; + +val incanTs = [ncaseT,lcaseT]; + +(*** Induction Rules ***) + +val ind_Ms = [NatM,ListM]; + +fun mk_ind_tac ddefs tdefs Ms canTs case_rls s = prove_goalw Type.thy ddefs s + (fn major::prems => [resolve_tac (Ms RL ([major] RL (tdefs RL [def_induct]))) 1, + fast_tac (set_cs addSIs (prems @ canTs) addSEs case_rls) 1]); + +val ind_tac = mk_ind_tac ind_data_defs ind_type_defs ind_Ms canTs case_rls; + +val Nat_ind = ind_tac + "[| n:Nat; P(zero); !!x.[| x:Nat; P(x) |] ==> P(succ(x)) |] ==> \ +\ P(n)"; + +val List_ind = ind_tac + "[| l:List(A); P([]); \ +\ !!x xs.[| x:A; xs:List(A); P(xs) |] ==> P(x.xs) |] ==> \ +\ P(l)"; + +val inds = [Nat_ind,List_ind]; + +(*** Primitive Recursive Rules ***) + +fun mk_prec_tac inds s = prove_goal Type.thy s + (fn major::prems => [resolve_tac ([major] RL inds) 1, + ALLGOALS (SIMP_TAC term_ss THEN' + fast_tac (set_cs addSIs prems))]); +val prec_tac = mk_prec_tac inds; + +val nrecT = prec_tac + "[| n:Nat; b:C(zero); \ +\ !!x g.[| x:Nat; g:C(x) |] ==> c(x,g):C(succ(x)) |] ==> \ +\ nrec(n,b,c) : C(n)"; + +val lrecT = prec_tac + "[| l:List(A); b:C([]); \ +\ !!x xs g.[| x:A; xs:List(A); g:C(xs) |] ==> c(x,xs,g):C(x.xs) |] ==> \ +\ lrec(l,b,c) : C(l)"; + +val precTs = [nrecT,lrecT]; + + +(*** Theorem proving ***) + +val [major,minor] = goal Type.thy + "[| : Sigma(A,B); [| a:A; b:B(a) |] ==> P \ +\ |] ==> P"; +br (major RS (XH_to_E SgXH)) 1; +br minor 1; +by (ALLGOALS (fast_tac term_cs)); +val SgE2 = result(); + +(* General theorem proving ignores non-canonical term-formers, *) +(* - intro rules are type rules for canonical terms *) +(* - elim rules are case rules (no non-canonical terms appear) *) + +val type_cs = term_cs addSIs (SubtypeI::(canTs @ icanTs)) + addSEs (SubtypeE::(XH_to_Es XHs)); + + +(*** Infinite Data Types ***) + +val [mono] = goal Type.thy "mono(f) ==> lfp(f) <= gfp(f)"; +br (lfp_lowerbound RS subset_trans) 1; +br (mono RS gfp_lemma3) 1; +br subset_refl 1; +val lfp_subset_gfp = result(); + +val prems = goal Type.thy + "[| a:A; !!x X.[| x:A; ALL y:A.t(y):X |] ==> t(x) : B(X) |] ==> \ +\ t(a) : gfp(B)"; +br coinduct 1; +by (res_inst_tac [("P","%x.EX y:A.x=t(y)")] CollectI 1); +by (ALLGOALS (fast_tac (ccl_cs addSIs prems))); +val gfpI = result(); + +val rew::prem::prems = goal Type.thy + "[| C==gfp(B); a:A; !!x X.[| x:A; ALL y:A.t(y):X |] ==> t(x) : B(X) |] ==> \ +\ t(a) : C"; +by (rewtac rew); +by (REPEAT (ares_tac ((prem RS gfpI)::prems) 1)); +val def_gfpI = result(); + +(* EG *) + +val prems = goal Type.thy + "letrec g x be zero.g(x) in g(bot) : Lists(Nat)"; +by (rtac (refl RS (XH_to_I UnitXH) RS (Lists_def RS def_gfpI)) 1); +br (letrecB RS ssubst) 1; +bw cons_def; +by (fast_tac type_cs 1); +result(); diff -r 000000000000 -r a5a9c433f639 src/CCL/type.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/type.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,73 @@ +(* Title: CCL/types.thy + ID: $Id$ + Author: Martin Coen + Copyright 1993 University of Cambridge + +Types in CCL are defined as sets of terms. + +*) + +Type = Term + + +consts + + Subtype :: "['a set, 'a => o] => 'a set" + Bool :: "i set" + Unit :: "i set" + "+" :: "[i set, i set] => i set" (infixr 55) + Pi :: "[i set, i => i set] => i set" + Sigma :: "[i set, i => i set] => i set" + Nat :: "i set" + List :: "i set => i set" + Lists :: "i set => i set" + ILists :: "i set => i set" + TAll :: "(i set => i set) => i set" (binder "TALL " 55) + TEx :: "(i set => i set) => i set" (binder "TEX " 55) + Lift :: "i set => i set" ("(3[_])") + + SPLIT :: "[i, [i, i] => i set] => i set" + + "@Pi" :: "[idt, i set, i set] => i set" ("(3PROD _:_./ _)" [] 60) + "@Sigma" :: "[idt, i set, i set] => i set" ("(3SUM _:_./ _)" [] 60) + "@->" :: "[i set, i set] => i set" ("(_ ->/ _)" [54, 53] 53) + "@*" :: "[i set, i set] => i set" ("(_ */ _)" [56, 55] 55) + "@Subtype" :: "[idt, 'a set, o] => 'a set" ("(1{_: _ ./ _})") + +translations + "PROD x:A. B" => "Pi(A, %x. B)" + "SUM x:A. B" => "Sigma(A, %x. B)" + "{x: A. B}" == "Subtype(A, %x. B)" + +rules + + Subtype_def "{x:A.P(x)} == {x.x:A & P(x)}" + Unit_def "Unit == {x.x=one}" + Bool_def "Bool == {x.x=true | x=false}" + Plus_def "A+B == {x. (EX a:A.x=inl(a)) | (EX b:B.x=inr(b))}" + Pi_def "Pi(A,B) == {x.EX b.x=lam x.b(x) & (ALL x:A.b(x):B(x))}" + Sigma_def "Sigma(A,B) == {x.EX a:A.EX b:B(a).x=}" + Nat_def "Nat == lfp(% X.Unit + X)" + List_def "List(A) == lfp(% X.Unit + A*X)" + + Lists_def "Lists(A) == gfp(% X.Unit + A*X)" + ILists_def "ILists(A) == gfp(% X.{} + A*X)" + + Tall_def "TALL X.B(X) == Inter({X.EX Y.X=B(Y)})" + Tex_def "TEX X.B(X) == Union({X.EX Y.X=B(Y)})" + Lift_def "[A] == A Un {bot}" + + SPLIT_def "SPLIT(p,B) == Union({A.EX x y.p= & A=B(x,y)})" + +end + + +ML + +val parse_translation = + [("@->", ndependent_tr "Pi"), + ("@*", ndependent_tr "Sigma")]; + +val print_translation = + [("Pi", dependent_tr' ("@Pi", "@->")), + ("Sigma", dependent_tr' ("@Sigma", "@*"))]; + diff -r 000000000000 -r a5a9c433f639 src/CCL/typecheck.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/typecheck.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,142 @@ +(* Title: CCL/typecheck + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +*) + +(*** Lemmas for constructors and subtypes ***) + +(* 0-ary constructors do not need additional rules as they are handled *) +(* correctly by applying SubtypeI *) +(* +val Subtype_canTs = + let fun tac prems = cut_facts_tac prems 1 THEN + REPEAT (ares_tac (SubtypeI::canTs@icanTs) 1 ORELSE + eresolve_tac [SubtypeE] 1) + fun solve s = prove_goal Type.thy s (fn prems => [tac prems]) + in map solve + ["P(one) ==> one : {x:Unit.P(x)}", + "P(true) ==> true : {x:Bool.P(x)}", + "P(false) ==> false : {x:Bool.P(x)}", + "a : {x:A. b:{y:B(a).P()}} ==> : {x:Sigma(A,B).P(x)}", + "a : {x:A.P(inl(x))} ==> inl(a) : {x:A+B.P(x)}", + "b : {x:B.P(inr(x))} ==> inr(b) : {x:A+B.P(x)}", + "P(zero) ==> zero : {x:Nat.P(x)}", + "a : {x:Nat.P(succ(x))} ==> succ(a) : {x:Nat.P(x)}", + "P([]) ==> [] : {x:List(A).P(x)}", + "h : {x:A. t : {y:List(A).P(x.y)}} ==> h.t : {x:List(A).P(x)}" + ] end; +*) +val Subtype_canTs = + let fun tac prems = cut_facts_tac prems 1 THEN + REPEAT (ares_tac (SubtypeI::canTs@icanTs) 1 ORELSE + eresolve_tac [SubtypeE] 1) + fun solve s = prove_goal Type.thy s (fn prems => [tac prems]) + in map solve + ["a : {x:A. b:{y:B(a).P()}} ==> : {x:Sigma(A,B).P(x)}", + "a : {x:A.P(inl(x))} ==> inl(a) : {x:A+B.P(x)}", + "b : {x:B.P(inr(x))} ==> inr(b) : {x:A+B.P(x)}", + "a : {x:Nat.P(succ(x))} ==> succ(a) : {x:Nat.P(x)}", + "h : {x:A. t : {y:List(A).P(x.y)}} ==> h.t : {x:List(A).P(x)}" + ] end; + +val prems = goal Type.thy + "[| f(t):B; ~t=bot |] ==> let x be t in f(x) : B"; +by (cut_facts_tac prems 1); +be (letB RS ssubst) 1; +ba 1; +val letT = result(); + +val prems = goal Type.thy + "[| a:A; f : Pi(A,B) |] ==> f ` a : B(a)"; +by (REPEAT (ares_tac (applyT::prems) 1)); +val applyT2 = result(); + +val prems = goal Type.thy + "[| a:A; a:A ==> P(a) |] ==> a : {x:A.P(x)}"; +by (fast_tac (type_cs addSIs prems) 1); +val rcall_lemma1 = result(); + +val prems = goal Type.thy + "[| a:{x:A.Q(x)}; [| a:A; Q(a) |] ==> P(a) |] ==> a : {x:A.P(x)}"; +by (cut_facts_tac prems 1); +by (fast_tac (type_cs addSIs prems) 1); +val rcall_lemma2 = result(); + +val rcall_lemmas = [asm_rl,rcall_lemma1,SubtypeD1,rcall_lemma2]; + +(***********************************TYPECHECKING*************************************) + +fun bvars (Const("all",_) $ Abs(s,_,t)) l = bvars t (s::l) + | bvars _ l = l; + +fun get_bno l n (Const("all",_) $ Abs(s,_,t)) = get_bno (s::l) n t + | get_bno l n (Const("Trueprop",_) $ t) = get_bno l n t + | get_bno l n (Const("Ball",_) $ _ $ Abs(s,_,t)) = get_bno (s::l) (n+1) t + | get_bno l n (Const("op :",_) $ t $ _) = get_bno l n t + | get_bno l n (t $ s) = get_bno l n t + | get_bno l n (Bound m) = (m-length(l),n); + +(* Not a great way of identifying induction hypothesis! *) +fun could_IH x = could_unify(x,hd (prems_of rcallT)) orelse + could_unify(x,hd (prems_of rcall2T)) orelse + could_unify(x,hd (prems_of rcall3T)); + +fun IHinst tac rls i = STATE (fn state => + let val (_,_,Bi,_) = dest_state(state,i); + val bvs = bvars Bi []; + val ihs = filter could_IH (Logic.strip_assums_hyp Bi); + val rnames = map (fn x=> + let val (a,b) = get_bno [] 0 x + in (nth_elem(a,bvs),b) end) ihs; + fun try_IHs [] = no_tac + | try_IHs ((x,y)::xs) = tac [("g",x)] (nth_elem(y-1,rls)) i ORELSE (try_IHs xs); + in try_IHs rnames end); + +(*****) + +val type_rls = canTs@icanTs@(applyT2::ncanTs)@incanTs@ + precTs@letrecTs@[letT]@Subtype_canTs; + +fun is_rigid_prog t = + case (Logic.strip_assums_concl t) of + (Const("Trueprop",_) $ (Const("op :",_) $ a $ _)) => (term_vars a = []) + | _ => false; + +fun rcall_tac i = let fun tac ps rl i = res_inst_tac ps rl i THEN atac i; + in IHinst tac rcallTs i end THEN + eresolve_tac rcall_lemmas i; + +fun raw_step_tac prems i = ares_tac (prems@type_rls) i ORELSE + rcall_tac i ORELSE + ematch_tac [SubtypeE] i ORELSE + match_tac [SubtypeI] i; + +fun tc_step_tac prems i = STATE(fn state => + if (i>length(prems_of state)) + then no_tac + else let val (_,_,c,_) = dest_state(state,i) + in if is_rigid_prog c then raw_step_tac prems i else no_tac + end); + +fun typechk_tac rls i = SELECT_GOAL (REPEAT_FIRST (tc_step_tac rls)) i; + +val tac = typechk_tac [] 1; + + +(*** Clean up Correctness Condictions ***) + +val clean_ccs_tac = REPEAT_FIRST (eresolve_tac ([SubtypeE]@rmIHs) ORELSE' + hyp_subst_tac); + +val clean_ccs_tac = + let fun tac ps rl i = eres_inst_tac ps rl i THEN atac i; + in TRY (REPEAT_FIRST (IHinst tac hyprcallTs ORELSE' + eresolve_tac ([asm_rl,SubtypeE]@rmIHs) ORELSE' + hyp_subst_tac)) end; + +fun gen_ccs_tac rls i = SELECT_GOAL (REPEAT_FIRST (tc_step_tac rls) THEN + clean_ccs_tac) i; + + diff -r 000000000000 -r a5a9c433f639 src/CCL/wfd.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/wfd.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,208 @@ +(* Title: CCL/wf + ID: $Id$ + +For wf.thy. + +Based on + Titles: ZF/wf.ML and HOL/ex/lex-prod + Authors: Lawrence C Paulson and Tobias Nipkow + Copyright 1992 University of Cambridge + +*) + +open Wfd; + +(***********) + +val wfd_congs = mk_congs Wfd.thy ["Wfd","wf","op **","wmap","ListPR"]; + +(***********) + +val [major,prem] = goalw Wfd.thy [Wfd_def] + "[| Wfd(R); \ +\ !!x.[| ALL y. : R --> P(y) |] ==> P(x) |] ==> \ +\ P(a)"; +by (rtac (major RS spec RS mp RS spec RS CollectD) 1); +by (fast_tac (set_cs addSIs [prem RS CollectI]) 1); +val wfd_induct = result(); + +val [p1,p2,p3] = goal Wfd.thy + "[| !!x y. : R ==> Q(x); \ +\ ALL x. (ALL y. : R --> y : P) --> x : P; \ +\ !!x.Q(x) ==> x:P |] ==> a:P"; +br (p2 RS spec RS mp) 1; +by (fast_tac (set_cs addSIs [p1 RS p3]) 1); +val wfd_strengthen_lemma = result(); + +fun wfd_strengthen_tac s i = res_inst_tac [("Q",s)] wfd_strengthen_lemma i THEN + assume_tac (i+1); + +val wfd::prems = goal Wfd.thy "[| Wfd(r); :r; :r |] ==> P"; +by (subgoal_tac "ALL x. :r --> :r --> P" 1); +by (fast_tac (FOL_cs addIs prems) 1); +br (wfd RS wfd_induct) 1; +by (ALLGOALS (fast_tac (ccl_cs addSIs prems))); +val wf_anti_sym = result(); + +val prems = goal Wfd.thy "[| Wfd(r); : r |] ==> P"; +by (rtac wf_anti_sym 1); +by (REPEAT (resolve_tac prems 1)); +val wf_anti_refl = result(); + +(*** Irreflexive transitive closure ***) + +val [prem] = goal Wfd.thy "Wfd(R) ==> Wfd(R^+)"; +by (rewtac Wfd_def); +by (REPEAT (ares_tac [allI,ballI,impI] 1)); +(*must retain the universal formula for later use!*) +by (rtac allE 1 THEN assume_tac 1); +by (etac mp 1); +br (prem RS wfd_induct) 1; +by (rtac (impI RS allI) 1); +by (etac tranclE 1); +by (fast_tac ccl_cs 1); +be (spec RS mp RS spec RS mp) 1; +by (REPEAT (atac 1)); +val trancl_wf = result(); + +(*** Lexicographic Ordering ***) + +goalw Wfd.thy [lex_def] + "p : ra**rb <-> (EX a a' b b'.p = <,> & ( : ra | a=a' & : rb))"; +by (fast_tac ccl_cs 1); +val lexXH = result(); + +val prems = goal Wfd.thy + " : ra ==> <,> : ra**rb"; +by (fast_tac (ccl_cs addSIs (prems @ [lexXH RS iffD2])) 1); +val lexI1 = result(); + +val prems = goal Wfd.thy + " : rb ==> <,> : ra**rb"; +by (fast_tac (ccl_cs addSIs (prems @ [lexXH RS iffD2])) 1); +val lexI2 = result(); + +val major::prems = goal Wfd.thy + "[| p : ra**rb; \ +\ !!a a' b b'.[| : ra; p=<,> |] ==> R; \ +\ !!a b b'.[| : rb; p = <,> |] ==> R |] ==> \ +\ R"; +br (major RS (lexXH RS iffD1) RS exE) 1; +by (REPEAT_SOME (eresolve_tac ([exE,conjE,disjE]@prems))); +by (ALLGOALS (fast_tac ccl_cs)); +val lexE = result(); + +val [major,minor] = goal Wfd.thy + "[| p : r**s; !!a a' b b'. p = <,> ==> P |] ==>P"; +br (major RS lexE) 1; +by (ALLGOALS (fast_tac (set_cs addSEs [minor]))); +val lex_pair = result(); + +val [wfa,wfb] = goal Wfd.thy + "[| Wfd(R); Wfd(S) |] ==> Wfd(R**S)"; +bw Wfd_def; +by (safe_tac ccl_cs); +by (wfd_strengthen_tac "%x.EX a b.x=" 1); +by (fast_tac (term_cs addSEs [lex_pair]) 1); +by (subgoal_tac "ALL a b.:P" 1); +by (fast_tac ccl_cs 1); +br (wfa RS wfd_induct RS allI) 1; +br (wfb RS wfd_induct RS allI) 1;back(); +by (fast_tac (type_cs addSEs [lexE]) 1); +val lex_wf = result(); + +(*** Mapping ***) + +goalw Wfd.thy [wmap_def] + "p : wmap(f,r) <-> (EX x y. p= & : r)"; +by (fast_tac ccl_cs 1); +val wmapXH = result(); + +val prems = goal Wfd.thy + " : r ==> : wmap(f,r)"; +by (fast_tac (ccl_cs addSIs (prems @ [wmapXH RS iffD2])) 1); +val wmapI = result(); + +val major::prems = goal Wfd.thy + "[| p : wmap(f,r); !!a b.[| : r; p= |] ==> R |] ==> R"; +br (major RS (wmapXH RS iffD1) RS exE) 1; +by (REPEAT_SOME (eresolve_tac ([exE,conjE,disjE]@prems))); +by (ALLGOALS (fast_tac ccl_cs)); +val wmapE = result(); + +val [wf] = goal Wfd.thy + "Wfd(r) ==> Wfd(wmap(f,r))"; +bw Wfd_def; +by (safe_tac ccl_cs); +by (subgoal_tac "ALL b.ALL a.f(a)=b-->a:P" 1); +by (fast_tac ccl_cs 1); +br (wf RS wfd_induct RS allI) 1; +by (safe_tac ccl_cs); +be (spec RS mp) 1; +by (safe_tac (ccl_cs addSEs [wmapE])); +be (spec RS mp RS spec RS mp) 1; +ba 1; +br refl 1; +val wmap_wf = result(); + +(* Projections *) + +val prems = goal Wfd.thy " : r ==> <,> : wmap(fst,r)"; +br wmapI 1; +by (SIMP_TAC (term_ss addrews prems) 1); +val wfstI = result(); + +val prems = goal Wfd.thy " : r ==> <,> : wmap(snd,r)"; +br wmapI 1; +by (SIMP_TAC (term_ss addrews prems) 1); +val wsndI = result(); + +val prems = goal Wfd.thy " : r ==> <>,>> : wmap(thd,r)"; +br wmapI 1; +by (SIMP_TAC (term_ss addrews prems) 1); +val wthdI = result(); + +(*** Ground well-founded relations ***) + +val prems = goalw Wfd.thy [wf_def] + "[| Wfd(r); a : r |] ==> a : wf(r)"; +by (fast_tac (set_cs addSIs prems) 1); +val wfI = result(); + +val prems = goalw Wfd.thy [Wfd_def] "Wfd({})"; +by (fast_tac (set_cs addEs [EmptyXH RS iffD1 RS FalseE]) 1); +val Empty_wf = result(); + +val prems = goalw Wfd.thy [wf_def] "Wfd(wf(R))"; +by (res_inst_tac [("Q","Wfd(R)")] (excluded_middle RS disjE) 1); +by (ALLGOALS (ASM_SIMP_TAC (CCL_ss addcongs wfd_congs))); +br Empty_wf 1; +val wf_wf = result(); + +goalw Wfd.thy [NatPR_def] "p : NatPR <-> (EX x:Nat.p=)"; +by (fast_tac set_cs 1); +val NatPRXH = result(); + +goalw Wfd.thy [ListPR_def] "p : ListPR(A) <-> (EX h:A.EX t:List(A).p=)"; +by (fast_tac set_cs 1); +val ListPRXH = result(); + +val NatPRI = refl RS (bexI RS (NatPRXH RS iffD2)); +val ListPRI = refl RS (bexI RS (bexI RS (ListPRXH RS iffD2))); + +goalw Wfd.thy [Wfd_def] "Wfd(NatPR)"; +by (safe_tac set_cs); +by (wfd_strengthen_tac "%x.x:Nat" 1); +by (fast_tac (type_cs addSEs [XH_to_E NatPRXH]) 1); +be Nat_ind 1; +by (ALLGOALS (fast_tac (type_cs addEs [XH_to_E NatPRXH]))); +val NatPR_wf = result(); + +goalw Wfd.thy [Wfd_def] "Wfd(ListPR(A))"; +by (safe_tac set_cs); +by (wfd_strengthen_tac "%x.x:List(A)" 1); +by (fast_tac (type_cs addSEs [XH_to_E ListPRXH]) 1); +be List_ind 1; +by (ALLGOALS (fast_tac (type_cs addEs [XH_to_E ListPRXH]))); +val ListPR_wf = result(); + diff -r 000000000000 -r a5a9c433f639 src/CCL/wfd.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CCL/wfd.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,34 @@ +(* Title: CCL/wf.thy + ID: $Id$ + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Well-founded relations in CCL. +*) + +Wfd = Trancl + Type + + +consts + (*** Predicates ***) + Wfd :: "[i set] => o" + (*** Relations ***) + wf :: "[i set] => i set" + wmap :: "[i=>i,i set] => i set" + "**" :: "[i set,i set] => i set" (infixl 70) + NatPR :: "i set" + ListPR :: "i set => i set" + +rules + + Wfd_def + "Wfd(R) == ALL P.(ALL x.(ALL y. : R --> y:P) --> x:P) --> (ALL a.a:P)" + + wf_def "wf(R) == {x.x:R & Wfd(R)}" + + wmap_def "wmap(f,R) == {p. EX x y. p= & : R}" + lex_def + "ra**rb == {p. EX a a' b b'.p = <,> & ( : ra | (a=a' & : rb))}" + + NatPR_def "NatPR == {p.EX x:Nat. p=}" + ListPR_def "ListPR(A) == {p.EX h:A.EX t:List(A). p=}" +end diff -r 000000000000 -r a5a9c433f639 src/CTT/Arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/Arith.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,497 @@ +(* Title: CTT/arith + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Theorems for arith.thy (Arithmetic operators) + +Proofs about elementary arithmetic: addition, multiplication, etc. +Tests definitions and simplifier. +*) + +open Arith; +val arith_defs = [add_def, diff_def, absdiff_def, mult_def, mod_def, div_def]; + + +(** Addition *) + +(*typing of add: short and long versions*) + +val add_typing = prove_goal Arith.thy + "[| a:N; b:N |] ==> a #+ b : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (typechk_tac prems) ]); + +val add_typingL = prove_goal Arith.thy + "[| a=c:N; b=d:N |] ==> a #+ b = c #+ d : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (equal_tac prems) ]); + + +(*computation for add: 0 and successor cases*) + +val addC0 = prove_goal Arith.thy + "b:N ==> 0 #+ b = b : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (rew_tac prems) ]); + +val addC_succ = prove_goal Arith.thy + "[| a:N; b:N |] ==> succ(a) #+ b = succ(a #+ b) : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (rew_tac prems) ]); + + +(** Multiplication *) + +(*typing of mult: short and long versions*) + +val mult_typing = prove_goal Arith.thy + "[| a:N; b:N |] ==> a #* b : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (typechk_tac([add_typing]@prems)) ]); + +val mult_typingL = prove_goal Arith.thy + "[| a=c:N; b=d:N |] ==> a #* b = c #* d : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (equal_tac (prems@[add_typingL])) ]); + +(*computation for mult: 0 and successor cases*) + +val multC0 = prove_goal Arith.thy + "b:N ==> 0 #* b = 0 : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (rew_tac prems) ]); + +val multC_succ = prove_goal Arith.thy + "[| a:N; b:N |] ==> succ(a) #* b = b #+ (a #* b) : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (rew_tac prems) ]); + + +(** Difference *) + +(*typing of difference*) + +val diff_typing = prove_goal Arith.thy + "[| a:N; b:N |] ==> a - b : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (typechk_tac prems) ]); + +val diff_typingL = prove_goal Arith.thy + "[| a=c:N; b=d:N |] ==> a - b = c - d : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (equal_tac prems) ]); + + + +(*computation for difference: 0 and successor cases*) + +val diffC0 = prove_goal Arith.thy + "a:N ==> a - 0 = a : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (rew_tac prems) ]); + +(*Note: rec(a, 0, %z w.z) is pred(a). *) + +val diff_0_eq_0 = prove_goal Arith.thy + "b:N ==> 0 - b = 0 : N" + (fn prems=> + [ (NE_tac "b" 1), + (rewrite_goals_tac arith_defs), + (hyp_rew_tac prems) ]); + + +(*Essential to simplify FIRST!! (Else we get a critical pair) + succ(a) - succ(b) rewrites to pred(succ(a) - b) *) +val diff_succ_succ = prove_goal Arith.thy + "[| a:N; b:N |] ==> succ(a) - succ(b) = a - b : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (hyp_rew_tac prems), + (NE_tac "b" 1), + (hyp_rew_tac prems) ]); + + + +(*** Simplification *) + +val arith_typing_rls = + [add_typing, mult_typing, diff_typing]; + +val arith_congr_rls = + [add_typingL, mult_typingL, diff_typingL]; + +val congr_rls = arith_congr_rls@standard_congr_rls; + +val arithC_rls = + [addC0, addC_succ, + multC0, multC_succ, + diffC0, diff_0_eq_0, diff_succ_succ]; + + +structure Arith_simp_data: TSIMP_DATA = + struct + val refl = refl_elem + val sym = sym_elem + val trans = trans_elem + val refl_red = refl_red + val trans_red = trans_red + val red_if_equal = red_if_equal + val default_rls = arithC_rls @ comp_rls + val routine_tac = routine_tac (arith_typing_rls @ routine_rls) + end; + +structure Arith_simp = TSimpFun (Arith_simp_data); + +fun arith_rew_tac prems = make_rew_tac + (Arith_simp.norm_tac(congr_rls, prems)); + +fun hyp_arith_rew_tac prems = make_rew_tac + (Arith_simp.cond_norm_tac(prove_cond_tac, congr_rls, prems)); + + +(********** + Addition + **********) + +(*Associative law for addition*) +val add_assoc = prove_goal Arith.thy + "[| a:N; b:N; c:N |] ==> (a #+ b) #+ c = a #+ (b #+ c) : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac prems) ]); + + +(*Commutative law for addition. Can be proved using three inductions. + Must simplify after first induction! Orientation of rewrites is delicate*) +val add_commute = prove_goal Arith.thy + "[| a:N; b:N |] ==> a #+ b = b #+ a : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac prems), + (NE_tac "b" 2), + (resolve_tac [sym_elem] 1), + (NE_tac "b" 1), + (hyp_arith_rew_tac prems) ]); + + +(**************** + Multiplication + ****************) + +(*Commutative law for multiplication +val mult_commute = prove_goal Arith.thy + "[| a:N; b:N |] ==> a #* b = b #* a : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac prems), + (NE_tac "b" 2), + (resolve_tac [sym_elem] 1), + (NE_tac "b" 1), + (hyp_arith_rew_tac prems) ]); NEEDS COMMUTATIVE MATCHING +***************) + +(*right annihilation in product*) +val mult_0_right = prove_goal Arith.thy + "a:N ==> a #* 0 = 0 : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac prems) ]); + +(*right successor law for multiplication*) +val mult_succ_right = prove_goal Arith.thy + "[| a:N; b:N |] ==> a #* succ(b) = a #+ (a #* b) : N" + (fn prems=> + [ (NE_tac "a" 1), +(*swap round the associative law of addition*) + (hyp_arith_rew_tac (prems @ [add_assoc RS sym_elem])), +(*leaves a goal involving a commutative law*) + (REPEAT (assume_tac 1 ORELSE + resolve_tac + (prems@[add_commute,mult_typingL,add_typingL]@ + intrL_rls@[refl_elem]) 1)) ]); + +(*Commutative law for multiplication*) +val mult_commute = prove_goal Arith.thy + "[| a:N; b:N |] ==> a #* b = b #* a : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac (prems @ [mult_0_right, mult_succ_right])) ]); + +(*addition distributes over multiplication*) +val add_mult_distrib = prove_goal Arith.thy + "[| a:N; b:N; c:N |] ==> (a #+ b) #* c = (a #* c) #+ (b #* c) : N" + (fn prems=> + [ (NE_tac "a" 1), +(*swap round the associative law of addition*) + (hyp_arith_rew_tac (prems @ [add_assoc RS sym_elem])) ]); + + +(*Associative law for multiplication*) +val mult_assoc = prove_goal Arith.thy + "[| a:N; b:N; c:N |] ==> (a #* b) #* c = a #* (b #* c) : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac (prems @ [add_mult_distrib])) ]); + + +(************ + Difference + ************ + +Difference on natural numbers, without negative numbers + a - b = 0 iff a<=b a - b = succ(c) iff a>b *) + +val diff_self_eq_0 = prove_goal Arith.thy + "a:N ==> a - a = 0 : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac prems) ]); + + +(* [| c : N; 0 : N; c : N |] ==> c #+ 0 = c : N *) +val add_0_right = addC0 RSN (3, add_commute RS trans_elem); + +(*Addition is the inverse of subtraction: if b<=x then b#+(x-b) = x. + An example of induction over a quantified formula (a product). + Uses rewriting with a quantified, implicative inductive hypothesis.*) +val prems = +goal Arith.thy + "b:N ==> ?a : PROD x:N. Eq(N, b-x, 0) --> Eq(N, b #+ (x-b), x)"; +by (NE_tac "b" 1); +(*strip one "universal quantifier" but not the "implication"*) +by (resolve_tac intr_rls 3); +(*case analysis on x in + (succ(u) <= x) --> (succ(u)#+(x-succ(u)) = x) *) +by (NE_tac "x" 4 THEN assume_tac 4); +(*Prepare for simplification of types -- the antecedent succ(u)<=x *) +by (resolve_tac [replace_type] 5); +by (resolve_tac [replace_type] 4); +by (arith_rew_tac prems); +(*Solves first 0 goal, simplifies others. Two sugbgoals remain. + Both follow by rewriting, (2) using quantified induction hyp*) +by (intr_tac[]); (*strips remaining PRODs*) +by (hyp_arith_rew_tac (prems@[add_0_right])); +by (assume_tac 1); +val add_diff_inverse_lemma = result(); + + +(*Version of above with premise b-a=0 i.e. a >= b. + Using ProdE does not work -- for ?B(?a) is ambiguous. + Instead, add_diff_inverse_lemma states the desired induction scheme; + the use of RS below instantiates Vars in ProdE automatically. *) +val prems = +goal Arith.thy "[| a:N; b:N; b-a = 0 : N |] ==> b #+ (a-b) = a : N"; +by (resolve_tac [EqE] 1); +by (resolve_tac [ add_diff_inverse_lemma RS ProdE RS ProdE ] 1); +by (REPEAT (resolve_tac (prems@[EqI]) 1)); +val add_diff_inverse = result(); + + +(******************** + Absolute difference + ********************) + +(*typing of absolute difference: short and long versions*) + +val absdiff_typing = prove_goal Arith.thy + "[| a:N; b:N |] ==> a |-| b : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (typechk_tac prems) ]); + +val absdiff_typingL = prove_goal Arith.thy + "[| a=c:N; b=d:N |] ==> a |-| b = c |-| d : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (equal_tac prems) ]); + +val absdiff_self_eq_0 = prove_goal Arith.thy + "a:N ==> a |-| a = 0 : N" + (fn prems=> + [ (rewrite_goals_tac [absdiff_def]), + (arith_rew_tac (prems@[diff_self_eq_0])) ]); + +val absdiffC0 = prove_goal Arith.thy + "a:N ==> 0 |-| a = a : N" + (fn prems=> + [ (rewrite_goals_tac [absdiff_def]), + (hyp_arith_rew_tac prems) ]); + + +val absdiff_succ_succ = prove_goal Arith.thy + "[| a:N; b:N |] ==> succ(a) |-| succ(b) = a |-| b : N" + (fn prems=> + [ (rewrite_goals_tac [absdiff_def]), + (hyp_arith_rew_tac prems) ]); + +(*Note how easy using commutative laws can be? ...not always... *) +val prems = goal Arith.thy "[| a:N; b:N |] ==> a |-| b = b |-| a : N"; +by (rewrite_goals_tac [absdiff_def]); +by (resolve_tac [add_commute] 1); +by (typechk_tac ([diff_typing]@prems)); +val absdiff_commute = result(); + +(*If a+b=0 then a=0. Surprisingly tedious*) +val prems = +goal Arith.thy "[| a:N; b:N |] ==> ?c : PROD u: Eq(N,a#+b,0) . Eq(N,a,0)"; +by (NE_tac "a" 1); +by (resolve_tac [replace_type] 3); +by (arith_rew_tac prems); +by (intr_tac[]); (*strips remaining PRODs*) +by (resolve_tac [ zero_ne_succ RS FE ] 2); +by (etac (EqE RS sym_elem) 3); +by (typechk_tac ([add_typing] @prems)); +val add_eq0_lemma = result(); + +(*Version of above with the premise a+b=0. + Again, resolution instantiates variables in ProdE *) +val prems = +goal Arith.thy "[| a:N; b:N; a #+ b = 0 : N |] ==> a = 0 : N"; +by (resolve_tac [EqE] 1); +by (resolve_tac [add_eq0_lemma RS ProdE] 1); +by (resolve_tac [EqI] 3); +by (ALLGOALS (resolve_tac prems)); +val add_eq0 = result(); + +(*Here is a lemma to infer a-b=0 and b-a=0 from a|-|b=0, below. *) +val prems = goal Arith.thy + "[| a:N; b:N; a |-| b = 0 : N |] ==> \ +\ ?a : SUM v: Eq(N, a-b, 0) . Eq(N, b-a, 0)"; +by (intr_tac[]); +by eqintr_tac; +by (resolve_tac [add_eq0] 2); +by (resolve_tac [add_eq0] 1); +by (resolve_tac [add_commute RS trans_elem] 6); +by (typechk_tac (diff_typing:: map (rewrite_rule [absdiff_def]) prems)); +val absdiff_eq0_lem = result(); + +(*if a |-| b = 0 then a = b + proof: a-b=0 and b-a=0, so b = a+(b-a) = a+0 = a*) +val prems = +goal Arith.thy "[| a |-| b = 0 : N; a:N; b:N |] ==> a = b : N"; +by (resolve_tac [EqE] 1); +by (resolve_tac [absdiff_eq0_lem RS SumE] 1); +by (TRYALL (resolve_tac prems)); +by eqintr_tac; +by (resolve_tac [add_diff_inverse RS sym_elem RS trans_elem] 1); +by (resolve_tac [EqE] 3 THEN assume_tac 3); +by (hyp_arith_rew_tac (prems@[add_0_right])); +val absdiff_eq0 = result(); + +(*********************** + Remainder and Quotient + ***********************) + +(*typing of remainder: short and long versions*) + +val mod_typing = prove_goal Arith.thy + "[| a:N; b:N |] ==> a mod b : N" + (fn prems=> + [ (rewrite_goals_tac [mod_def]), + (typechk_tac (absdiff_typing::prems)) ]); + +val mod_typingL = prove_goal Arith.thy + "[| a=c:N; b=d:N |] ==> a mod b = c mod d : N" + (fn prems=> + [ (rewrite_goals_tac [mod_def]), + (equal_tac (prems@[absdiff_typingL])) ]); + + +(*computation for mod : 0 and successor cases*) + +val modC0 = prove_goal Arith.thy "b:N ==> 0 mod b = 0 : N" + (fn prems=> + [ (rewrite_goals_tac [mod_def]), + (rew_tac(absdiff_typing::prems)) ]); + +val modC_succ = prove_goal Arith.thy +"[| a:N; b:N |] ==> succ(a) mod b = rec(succ(a mod b) |-| b, 0, %x y.succ(a mod b)) : N" + (fn prems=> + [ (rewrite_goals_tac [mod_def]), + (rew_tac(absdiff_typing::prems)) ]); + + +(*typing of quotient: short and long versions*) + +val div_typing = prove_goal Arith.thy "[| a:N; b:N |] ==> a div b : N" + (fn prems=> + [ (rewrite_goals_tac [div_def]), + (typechk_tac ([absdiff_typing,mod_typing]@prems)) ]); + +val div_typingL = prove_goal Arith.thy + "[| a=c:N; b=d:N |] ==> a div b = c div d : N" + (fn prems=> + [ (rewrite_goals_tac [div_def]), + (equal_tac (prems @ [absdiff_typingL, mod_typingL])) ]); + +val div_typing_rls = [mod_typing, div_typing, absdiff_typing]; + + +(*computation for quotient: 0 and successor cases*) + +val divC0 = prove_goal Arith.thy "b:N ==> 0 div b = 0 : N" + (fn prems=> + [ (rewrite_goals_tac [div_def]), + (rew_tac([mod_typing, absdiff_typing] @ prems)) ]); + +val divC_succ = +prove_goal Arith.thy "[| a:N; b:N |] ==> succ(a) div b = \ +\ rec(succ(a) mod b, succ(a div b), %x y. a div b) : N" + (fn prems=> + [ (rewrite_goals_tac [div_def]), + (rew_tac([mod_typing]@prems)) ]); + + +(*Version of above with same condition as the mod one*) +val divC_succ2 = prove_goal Arith.thy + "[| a:N; b:N |] ==> \ +\ succ(a) div b =rec(succ(a mod b) |-| b, succ(a div b), %x y. a div b) : N" + (fn prems=> + [ (resolve_tac [ divC_succ RS trans_elem ] 1), + (rew_tac(div_typing_rls @ prems @ [modC_succ])), + (NE_tac "succ(a mod b)|-|b" 1), + (rew_tac ([mod_typing, div_typing, absdiff_typing] @prems)) ]); + +(*for case analysis on whether a number is 0 or a successor*) +val iszero_decidable = prove_goal Arith.thy + "a:N ==> rec(a, inl(eq), %ka kb.inr()) : \ +\ Eq(N,a,0) + (SUM x:N. Eq(N,a, succ(x)))" + (fn prems=> + [ (NE_tac "a" 1), + (resolve_tac [PlusI_inr] 3), + (resolve_tac [PlusI_inl] 2), + eqintr_tac, + (equal_tac prems) ]); + +(*Main Result. Holds when b is 0 since a mod 0 = a and a div 0 = 0 *) +val prems = +goal Arith.thy "[| a:N; b:N |] ==> a mod b #+ (a div b) #* b = a : N"; +by (NE_tac "a" 1); +by (arith_rew_tac (div_typing_rls@prems@[modC0,modC_succ,divC0,divC_succ2])); +by (resolve_tac [EqE] 1); +(*case analysis on succ(u mod b)|-|b *) +by (res_inst_tac [("a1", "succ(u mod b) |-| b")] + (iszero_decidable RS PlusE) 1); +by (etac SumE 3); +by (hyp_arith_rew_tac (prems @ div_typing_rls @ + [modC0,modC_succ, divC0, divC_succ2])); +(*Replace one occurence of b by succ(u mod b). Clumsy!*) +by (resolve_tac [ add_typingL RS trans_elem ] 1); +by (eresolve_tac [EqE RS absdiff_eq0 RS sym_elem] 1); +by (resolve_tac [refl_elem] 3); +by (hyp_arith_rew_tac (prems @ div_typing_rls)); +val mod_div_equality = result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/CTT/Arith.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/Arith.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,24 @@ +(* Title: CTT/arith + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Arithmetic operators and their definitions + +Proves about elementary arithmetic: addition, multiplication, etc. +Tests definitions and simplifier. +*) + +Arith = CTT + + +consts "#+","-","|-|" :: "[i,i]=>i" (infixr 65) + "#*",div,mod :: "[i,i]=>i" (infixr 70) + +rules + add_def "a#+b == rec(a, b, %u v.succ(v))" + diff_def "a-b == rec(b, a, %u v.rec(v, 0, %x y.x))" + absdiff_def "a|-|b == (a-b) #+ (b-a)" + mult_def "a#*b == rec(a, 0, %u v. b #+ v)" + mod_def "a mod b == rec(a, 0, %u v. rec(succ(v) |-| b, 0, %x y.succ(v)))" + div_def "a div b == rec(a, 0, %u v. rec(succ(u) mod b, succ(v), %x y.v))" +end diff -r 000000000000 -r a5a9c433f639 src/CTT/Bool.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/Bool.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,79 @@ +(* Title: CTT/bool + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Theorems for bool.thy (booleans and conditionals) +*) + +open Bool; + +val bool_defs = [Bool_def,true_def,false_def,cond_def]; + +(*Derivation of rules for the type Bool*) + +(*formation rule*) +val boolF = prove_goal Bool.thy + "Bool type" + (fn prems=> + [ (rewrite_goals_tac bool_defs), + (typechk_tac[]) ]); + + +(*introduction rules for true, false*) + +val boolI_true = prove_goal Bool.thy + "true : Bool" + (fn prems=> + [ (rewrite_goals_tac bool_defs), + (typechk_tac[]) ]); + +val boolI_false = prove_goal Bool.thy + "false : Bool" + (fn prems=> + [ (rewrite_goals_tac bool_defs), + (typechk_tac[]) ]); + +(*elimination rule: typing of cond*) +val boolE = prove_goal Bool.thy + "[| p:Bool; a : C(true); b : C(false) |] ==> cond(p,a,b) : C(p)" + (fn prems=> + [ (cut_facts_tac prems 1), + (rewrite_goals_tac bool_defs), + (typechk_tac prems), + (ALLGOALS (etac TE)), + (typechk_tac prems) ]); + +val boolEL = prove_goal Bool.thy + "[| p = q : Bool; a = c : C(true); b = d : C(false) |] ==> \ +\ cond(p,a,b) = cond(q,c,d) : C(p)" + (fn prems=> + [ (cut_facts_tac prems 1), + (rewrite_goals_tac bool_defs), + (resolve_tac [PlusEL] 1), + (REPEAT (eresolve_tac [asm_rl, refl_elem RS TEL] 1)) ]); + +(*computation rules for true, false*) + +val boolC_true = prove_goal Bool.thy + "[| a : C(true); b : C(false) |] ==> cond(true,a,b) = a : C(true)" + (fn prems=> + [ (cut_facts_tac prems 1), + (rewrite_goals_tac bool_defs), + (resolve_tac comp_rls 1), + (typechk_tac[]), + (ALLGOALS (etac TE)), + (typechk_tac prems) ]); + +val boolC_false = prove_goal Bool.thy + "[| a : C(true); b : C(false) |] ==> cond(false,a,b) = b : C(false)" + (fn prems=> + [ (cut_facts_tac prems 1), + (rewrite_goals_tac bool_defs), + (resolve_tac comp_rls 1), + (typechk_tac[]), + (ALLGOALS (etac TE)), + (typechk_tac prems) ]); + +writeln"Reached end of file."; + diff -r 000000000000 -r a5a9c433f639 src/CTT/Bool.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/Bool.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,19 @@ +(* Title: CTT/bool + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +The two-element type (booleans and conditionals) +*) + +Bool = CTT + + +consts Bool :: "t" + true,false :: "i" + cond :: "[i,i,i]=>i" +rules + Bool_def "Bool == T+T" + true_def "true == inl(tt)" + false_def "false == inr(tt)" + cond_def "cond(a,b,c) == when(a, %u.b, %u.c)" +end diff -r 000000000000 -r a5a9c433f639 src/CTT/CTT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/CTT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,249 @@ +(* Title: CTT/ctt.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Tactics and lemmas for ctt.thy (Constructive Type Theory) +*) + +open CTT; + +signature CTT_RESOLVE = + sig + val add_mp_tac: int -> tactic + val ASSUME: (int -> tactic) -> int -> tactic + val basic_defs: thm list + val comp_rls: thm list + val element_rls: thm list + val elimL_rls: thm list + val elim_rls: thm list + val eqintr_tac: tactic + val equal_tac: thm list -> tactic + val formL_rls: thm list + val form_rls: thm list + val form_tac: tactic + val intrL2_rls: thm list + val intrL_rls: thm list + val intr_rls: thm list + val intr_tac: thm list -> tactic + val mp_tac: int -> tactic + val NE_tac: string -> int -> tactic + val pc_tac: thm list -> int -> tactic + val PlusE_tac: string -> int -> tactic + val reduction_rls: thm list + val replace_type: thm + val routine_rls: thm list + val routine_tac: thm list -> thm list -> int -> tactic + val safe_brls: (bool * thm) list + val safestep_tac: thm list -> int -> tactic + val safe_tac: thm list -> int -> tactic + val step_tac: thm list -> int -> tactic + val subst_eqtyparg: thm + val subst_prodE: thm + val SumE_fst: thm + val SumE_snd: thm + val SumE_tac: string -> int -> tactic + val SumIL2: thm + val test_assume_tac: int -> tactic + val typechk_tac: thm list -> tactic + val unsafe_brls: (bool * thm) list + end; + + +structure CTT_Resolve : CTT_RESOLVE = +struct + +(*Formation rules*) +val form_rls = [NF, ProdF, SumF, PlusF, EqF, FF, TF] +and formL_rls = [ProdFL, SumFL, PlusFL, EqFL]; + + +(*Introduction rules + OMITTED: EqI, because its premise is an eqelem, not an elem*) +val intr_rls = [NI0, NI_succ, ProdI, SumI, PlusI_inl, PlusI_inr, TI] +and intrL_rls = [NI_succL, ProdIL, SumIL, PlusI_inlL, PlusI_inrL]; + + +(*Elimination rules + OMITTED: EqE, because its conclusion is an eqelem, not an elem + TE, because it does not involve a constructor *) +val elim_rls = [NE, ProdE, SumE, PlusE, FE] +and elimL_rls = [NEL, ProdEL, SumEL, PlusEL, FEL]; + +(*OMITTED: eqC are TC because they make rewriting loop: p = un = un = ... *) +val comp_rls = [NC0, NC_succ, ProdC, SumC, PlusC_inl, PlusC_inr]; + +(*rules with conclusion a:A, an elem judgement*) +val element_rls = intr_rls @ elim_rls; + +(*Definitions are (meta)equality axioms*) +val basic_defs = [fst_def,snd_def]; + +(*Compare with standard version: B is applied to UNSIMPLIFIED expression! *) +val SumIL2 = prove_goal CTT.thy + "[| c=a : A; d=b : B(a) |] ==> = : Sum(A,B)" + (fn prems=> + [ (resolve_tac [sym_elem] 1), + (resolve_tac [SumIL] 1), + (ALLGOALS (resolve_tac [sym_elem])), + (ALLGOALS (resolve_tac prems)) ]); + +val intrL2_rls = [NI_succL, ProdIL, SumIL2, PlusI_inlL, PlusI_inrL]; + +(*Exploit p:Prod(A,B) to create the assumption z:B(a). + A more natural form of product elimination. *) +val subst_prodE = prove_goal CTT.thy + "[| p: Prod(A,B); a: A; !!z. z: B(a) ==> c(z): C(z) \ +\ |] ==> c(p`a): C(p`a)" + (fn prems=> + [ (REPEAT (resolve_tac (prems@[ProdE]) 1)) ]); + +(** Tactics for type checking **) + +fun is_rigid_elem (Const("Elem",_) $ a $ _) = not (is_Var (head_of a)) + | is_rigid_elem _ = false; + +(*Try solving a:A by assumption provided a is rigid!*) +val test_assume_tac = SUBGOAL(fn (prem,i) => + if is_rigid_elem (Logic.strip_assums_concl prem) + then assume_tac i else no_tac); + +fun ASSUME tf i = test_assume_tac i ORELSE tf i; + + +(*For simplification: type formation and checking, + but no equalities between terms*) +val routine_rls = form_rls @ formL_rls @ [refl_type] @ element_rls; + +fun routine_tac rls prems = ASSUME (filt_resolve_tac (prems @ rls) 4); + + +(*Solve all subgoals "A type" using formation rules. *) +val form_tac = REPEAT_FIRST (filt_resolve_tac(form_rls) 1); + + +(*Type checking: solve a:A (a rigid, A flexible) by intro and elim rules. *) +fun typechk_tac thms = + let val tac = filt_resolve_tac (thms @ form_rls @ element_rls) 3 + in REPEAT_FIRST (ASSUME tac) end; + + +(*Solve a:A (a flexible, A rigid) by introduction rules. + Cannot use stringtrees (filt_resolve_tac) since + goals like ?a:SUM(A,B) have a trivial head-string *) +fun intr_tac thms = + let val tac = filt_resolve_tac(thms@form_rls@intr_rls) 1 + in REPEAT_FIRST (ASSUME tac) end; + + +(*Equality proving: solve a=b:A (where a is rigid) by long rules. *) +fun equal_tac thms = + let val rls = thms @ form_rls @ element_rls @ intrL_rls @ + elimL_rls @ [refl_elem] + in REPEAT_FIRST (ASSUME (filt_resolve_tac rls 3)) end; + +(*** Simplification ***) + +(*To simplify the type in a goal*) +val replace_type = prove_goal CTT.thy + "[| B = A; a : A |] ==> a : B" + (fn prems=> + [ (resolve_tac [equal_types] 1), + (resolve_tac [sym_type] 2), + (ALLGOALS (resolve_tac prems)) ]); + +(*Simplify the parameter of a unary type operator.*) +val subst_eqtyparg = prove_goal CTT.thy + "a=c : A ==> (!!z.z:A ==> B(z) type) ==> B(a)=B(c)" + (fn prems=> + [ (resolve_tac [subst_typeL] 1), + (resolve_tac [refl_type] 2), + (ALLGOALS (resolve_tac prems)), + (assume_tac 1) ]); + +(*Make a reduction rule for simplification. + A goal a=c becomes b=c, by virtue of a=b *) +fun resolve_trans rl = rl RS trans_elem; + +(*Simplification rules for Constructive Type Theory*) +val reduction_rls = map resolve_trans comp_rls; + +(*Converts each goal "e : Eq(A,a,b)" into "a=b:A" for simplification. + Uses other intro rules to avoid changing flexible goals.*) +val eqintr_tac = REPEAT_FIRST (ASSUME (filt_resolve_tac(EqI::intr_rls) 1)); + +(** Tactics that instantiate CTT-rules. + Vars in the given terms will be incremented! + The (resolve_tac [EqE] i) lets them apply to equality judgements. **) + +fun NE_tac (sp: string) i = + TRY (resolve_tac [EqE] i) THEN res_inst_tac [ ("p",sp) ] NE i; + +fun SumE_tac (sp: string) i = + TRY (resolve_tac [EqE] i) THEN res_inst_tac [ ("p",sp) ] SumE i; + +fun PlusE_tac (sp: string) i = + TRY (resolve_tac [EqE] i) THEN res_inst_tac [ ("p",sp) ] PlusE i; + +(** Predicate logic reasoning, WITH THINNING!! Procedures adapted from NJ. **) + +(*Finds f:Prod(A,B) and a:A in the assumptions, concludes there is z:B(a) *) +fun add_mp_tac i = + resolve_tac [subst_prodE] i THEN assume_tac i THEN assume_tac i; + +(*Finds P-->Q and P in the assumptions, replaces implication by Q *) +fun mp_tac i = eresolve_tac [subst_prodE] i THEN assume_tac i; + +(*"safe" when regarded as predicate calculus rules*) +val safe_brls = sort lessb + [ (true,FE), (true,asm_rl), + (false,ProdI), (true,SumE), (true,PlusE) ]; + +val unsafe_brls = + [ (false,PlusI_inl), (false,PlusI_inr), (false,SumI), + (true,subst_prodE) ]; + +(*0 subgoals vs 1 or more*) +val (safe0_brls, safep_brls) = + partition (apl(0,op=) o subgoals_of_brl) safe_brls; + +fun safestep_tac thms i = + form_tac ORELSE + resolve_tac thms i ORELSE + biresolve_tac safe0_brls i ORELSE mp_tac i ORELSE + DETERM (biresolve_tac safep_brls i); + +fun safe_tac thms i = DEPTH_SOLVE_1 (safestep_tac thms i); + +fun step_tac thms = safestep_tac thms ORELSE' biresolve_tac unsafe_brls; + +(*Fails unless it solves the goal!*) +fun pc_tac thms = DEPTH_SOLVE_1 o (step_tac thms); + +(** The elimination rules for fst/snd **) + +val SumE_fst = prove_goal CTT.thy + "p : Sum(A,B) ==> fst(p) : A" + (fn prems=> + [ (rewrite_goals_tac basic_defs), + (resolve_tac elim_rls 1), + (REPEAT (pc_tac prems 1)), + (fold_tac basic_defs) ]); + +(*The first premise must be p:Sum(A,B) !!*) +val SumE_snd = prove_goal CTT.thy + "[| p: Sum(A,B); A type; !!x. x:A ==> B(x) type \ +\ |] ==> snd(p) : B(fst(p))" + (fn prems=> + [ (rewrite_goals_tac basic_defs), + (resolve_tac elim_rls 1), + (resolve_tac prems 1), + (resolve_tac [replace_type] 1), + (resolve_tac [subst_eqtyparg] 1), (*like B(x) equality formation?*) + (resolve_tac comp_rls 1), + (typechk_tac prems), + (fold_tac basic_defs) ]); + +end; + +open CTT_Resolve; diff -r 000000000000 -r a5a9c433f639 src/CTT/CTT.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/CTT.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,253 @@ +(* Title: CTT/ctt.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Constructive Type Theory +*) + +CTT = Pure + + +types i,t,o 0 + +arities i,t,o :: logic + +consts + (*Types*) + F,T :: "t" (*F is empty, T contains one element*) + contr :: "i=>i" + tt :: "i" + (*Natural numbers*) + N :: "t" + succ :: "i=>i" + rec :: "[i, i, [i,i]=>i] => i" + (*Unions*) + inl,inr :: "i=>i" + when :: "[i, i=>i, i=>i]=>i" + (*General Sum and Binary Product*) + Sum :: "[t, i=>t]=>t" + fst,snd :: "i=>i" + split :: "[i, [i,i]=>i] =>i" + (*General Product and Function Space*) + Prod :: "[t, i=>t]=>t" + (*Equality type*) + Eq :: "[t,i,i]=>t" + eq :: "i" + (*Judgements*) + Type :: "t => prop" ("(_ type)" [10] 5) + Eqtype :: "[t,t]=>prop" ("(3_ =/ _)" [10,10] 5) + Elem :: "[i, t]=>prop" ("(_ /: _)" [10,10] 5) + Eqelem :: "[i,i,t]=>prop" ("(3_ =/ _ :/ _)" [10,10,10] 5) + Reduce :: "[i,i]=>prop" ("Reduce[_,_]") + (*Types*) + "@PROD" :: "[id,t,t]=>t" ("(3PROD _:_./ _)" 10) + "@SUM" :: "[id,t,t]=>t" ("(3SUM _:_./ _)" 10) + "+" :: "[t,t]=>t" (infixr 40) + (*Invisible infixes!*) + "@-->" :: "[t,t]=>t" ("(_ -->/ _)" [31,30] 30) + "@*" :: "[t,t]=>t" ("(_ */ _)" [51,50] 50) + (*Functions*) + lambda :: "(i => i) => i" (binder "lam " 10) + "`" :: "[i,i]=>i" (infixl 60) + (*Natural numbers*) + "0" :: "i" ("0") + (*Pairing*) + pair :: "[i,i]=>i" ("(1<_,/_>)") + +translations + "PROD x:A. B" => "Prod(A, %x. B)" + "SUM x:A. B" => "Sum(A, %x. B)" + +rules + + (*Reduction: a weaker notion than equality; a hack for simplification. + Reduce[a,b] means either that a=b:A for some A or else that "a" and "b" + are textually identical.*) + + (*does not verify a:A! Sound because only trans_red uses a Reduce premise + No new theorems can be proved about the standard judgements.*) + refl_red "Reduce[a,a]" + red_if_equal "a = b : A ==> Reduce[a,b]" + trans_red "[| a = b : A; Reduce[b,c] |] ==> a = c : A" + + (*Reflexivity*) + + refl_type "A type ==> A = A" + refl_elem "a : A ==> a = a : A" + + (*Symmetry*) + + sym_type "A = B ==> B = A" + sym_elem "a = b : A ==> b = a : A" + + (*Transitivity*) + + trans_type "[| A = B; B = C |] ==> A = C" + trans_elem "[| a = b : A; b = c : A |] ==> a = c : A" + + equal_types "[| a : A; A = B |] ==> a : B" + equal_typesL "[| a = b : A; A = B |] ==> a = b : B" + + (*Substitution*) + + subst_type "[| a : A; !!z. z:A ==> B(z) type |] ==> B(a) type" + subst_typeL "[| a = c : A; !!z. z:A ==> B(z) = D(z) |] ==> B(a) = D(c)" + + subst_elem "[| a : A; !!z. z:A ==> b(z):B(z) |] ==> b(a):B(a)" + subst_elemL + "[| a=c : A; !!z. z:A ==> b(z)=d(z) : B(z) |] ==> b(a)=d(c) : B(a)" + + + (*The type N -- natural numbers*) + + NF "N type" + NI0 "0 : N" + NI_succ "a : N ==> succ(a) : N" + NI_succL "a = b : N ==> succ(a) = succ(b) : N" + + NE + "[| p: N; a: C(0); !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] \ +\ ==> rec(p, a, %u v.b(u,v)) : C(p)" + + NEL + "[| p = q : N; a = c : C(0); \ +\ !!u v. [| u: N; v: C(u) |] ==> b(u,v) = d(u,v): C(succ(u)) |] \ +\ ==> rec(p, a, %u v.b(u,v)) = rec(q,c,d) : C(p)" + + NC0 + "[| a: C(0); !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] \ +\ ==> rec(0, a, %u v.b(u,v)) = a : C(0)" + + NC_succ + "[| p: N; a: C(0); \ +\ !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] ==> \ +\ rec(succ(p), a, %u v.b(u,v)) = b(p, rec(p, a, %u v.b(u,v))) : C(succ(p))" + + (*The fourth Peano axiom. See page 91 of Martin-Lof's book*) + zero_ne_succ + "[| a: N; 0 = succ(a) : N |] ==> 0: F" + + + (*The Product of a family of types*) + + ProdF "[| A type; !!x. x:A ==> B(x) type |] ==> PROD x:A.B(x) type" + + ProdFL + "[| A = C; !!x. x:A ==> B(x) = D(x) |] ==> \ +\ PROD x:A.B(x) = PROD x:C.D(x)" + + ProdI + "[| A type; !!x. x:A ==> b(x):B(x)|] ==> lam x.b(x) : PROD x:A.B(x)" + + ProdIL + "[| A type; !!x. x:A ==> b(x) = c(x) : B(x)|] ==> \ +\ lam x.b(x) = lam x.c(x) : PROD x:A.B(x)" + + ProdE "[| p : PROD x:A.B(x); a : A |] ==> p`a : B(a)" + ProdEL "[| p=q: PROD x:A.B(x); a=b : A |] ==> p`a = q`b : B(a)" + + ProdC + "[| a : A; !!x. x:A ==> b(x) : B(x)|] ==> \ +\ (lam x.b(x)) ` a = b(a) : B(a)" + + ProdC2 + "p : PROD x:A.B(x) ==> (lam x. p`x) = p : PROD x:A.B(x)" + + + (*The Sum of a family of types*) + + SumF "[| A type; !!x. x:A ==> B(x) type |] ==> SUM x:A.B(x) type" + SumFL + "[| A = C; !!x. x:A ==> B(x) = D(x) |] ==> SUM x:A.B(x) = SUM x:C.D(x)" + + SumI "[| a : A; b : B(a) |] ==> : SUM x:A.B(x)" + SumIL "[| a=c:A; b=d:B(a) |] ==> = : SUM x:A.B(x)" + + SumE + "[| p: SUM x:A.B(x); !!x y. [| x:A; y:B(x) |] ==> c(x,y): C() |] \ +\ ==> split(p, %x y.c(x,y)) : C(p)" + + SumEL + "[| p=q : SUM x:A.B(x); \ +\ !!x y. [| x:A; y:B(x) |] ==> c(x,y)=d(x,y): C()|] \ +\ ==> split(p, %x y.c(x,y)) = split(q, % x y.d(x,y)) : C(p)" + + SumC + "[| a: A; b: B(a); !!x y. [| x:A; y:B(x) |] ==> c(x,y): C() |] \ +\ ==> split(, %x y.c(x,y)) = c(a,b) : C()" + + fst_def "fst(a) == split(a, %x y.x)" + snd_def "snd(a) == split(a, %x y.y)" + + + (*The sum of two types*) + + PlusF "[| A type; B type |] ==> A+B type" + PlusFL "[| A = C; B = D |] ==> A+B = C+D" + + PlusI_inl "[| a : A; B type |] ==> inl(a) : A+B" + PlusI_inlL "[| a = c : A; B type |] ==> inl(a) = inl(c) : A+B" + + PlusI_inr "[| A type; b : B |] ==> inr(b) : A+B" + PlusI_inrL "[| A type; b = d : B |] ==> inr(b) = inr(d) : A+B" + + PlusE + "[| p: A+B; !!x. x:A ==> c(x): C(inl(x)); \ +\ !!y. y:B ==> d(y): C(inr(y)) |] \ +\ ==> when(p, %x.c(x), %y.d(y)) : C(p)" + + PlusEL + "[| p = q : A+B; !!x. x: A ==> c(x) = e(x) : C(inl(x)); \ +\ !!y. y: B ==> d(y) = f(y) : C(inr(y)) |] \ +\ ==> when(p, %x.c(x), %y.d(y)) = when(q, %x.e(x), %y.f(y)) : C(p)" + + PlusC_inl + "[| a: A; !!x. x:A ==> c(x): C(inl(x)); \ +\ !!y. y:B ==> d(y): C(inr(y)) |] \ +\ ==> when(inl(a), %x.c(x), %y.d(y)) = c(a) : C(inl(a))" + + PlusC_inr + "[| b: B; !!x. x:A ==> c(x): C(inl(x)); \ +\ !!y. y:B ==> d(y): C(inr(y)) |] \ +\ ==> when(inr(b), %x.c(x), %y.d(y)) = d(b) : C(inr(b))" + + + (*The type Eq*) + + EqF "[| A type; a : A; b : A |] ==> Eq(A,a,b) type" + EqFL "[| A=B; a=c: A; b=d : A |] ==> Eq(A,a,b) = Eq(B,c,d)" + EqI "a = b : A ==> eq : Eq(A,a,b)" + EqE "p : Eq(A,a,b) ==> a = b : A" + + (*By equality of types, can prove C(p) from C(eq), an elimination rule*) + EqC "p : Eq(A,a,b) ==> p = eq : Eq(A,a,b)" + + (*The type F*) + + FF "F type" + FE "[| p: F; C type |] ==> contr(p) : C" + FEL "[| p = q : F; C type |] ==> contr(p) = contr(q) : C" + + (*The type T + Martin-Lof's book (page 68) discusses elimination and computation. + Elimination can be derived by computation and equality of types, + but with an extra premise C(x) type x:T. + Also computation can be derived from elimination. *) + + TF "T type" + TI "tt : T" + TE "[| p : T; c : C(tt) |] ==> c : C(p)" + TEL "[| p = q : T; c = d : C(tt) |] ==> c = d : C(p)" + TC "p : T ==> p = tt : T" +end + + +ML + +val parse_translation = + [("@-->", ndependent_tr "Prod"), ("@*", ndependent_tr "Sum")]; + +val print_translation = + [("Prod", dependent_tr' ("@PROD", "@-->")), + ("Sum", dependent_tr' ("@SUM", "@*"))]; + diff -r 000000000000 -r a5a9c433f639 src/CTT/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/Makefile Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,43 @@ +######################################################################### +# # +# Makefile for Isabelle (CTT) # +# # +######################################################################### + +#To make the system, cd to this directory and type +# make -f Makefile +#To make the system and test it on standard examples, type +# make -f Makefile test + +#Environment variable ISABELLECOMP specifies the compiler. +#Environment variable ISABELLEBIN specifies the destination directory. +#For Poly/ML, ISABELLEBIN must begin with a / + +#Makes pure Isabelle (Pure) if this file is ABSENT -- but not +#if it is out of date, since this Makefile does not know its dependencies! + +BIN = $(ISABELLEBIN) +COMP = $(ISABELLECOMP) +FILES = ROOT.ML ctt.thy ctt.ML bool.thy bool.ML \ + arith.thy arith.ML rew.ML ../Provers/typedsimp.ML + +$(BIN)/CTT: $(BIN)/Pure $(FILES) + case "$(COMP)" in \ + poly*) echo 'make_database"$(BIN)/CTT"; quit();' \ + | $(COMP) $(BIN)/Pure;\ + echo 'open PolyML; use"ROOT";' | $(COMP) $(BIN)/CTT ;;\ + sml*) echo 'use"ROOT.ML"; xML"$(BIN)/CTT" banner;' | $(BIN)/Pure ;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +$(BIN)/Pure: + cd ../Pure; $(MAKE) + +test: ex/ROOT.ML $(BIN)/CTT + case "$(COMP)" in \ + poly*) echo 'use"ex/ROOT.ML"; quit();' | $(COMP) $(BIN)/CTT ;;\ + sml*) echo 'use"ex/ROOT.ML";' | $(BIN)/CTT;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +.PRECIOUS: $(BIN)/Pure $(BIN)/CTT diff -r 000000000000 -r a5a9c433f639 src/CTT/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/README Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,23 @@ + CTT: Constructive Type Theory + +This directory contains the Standard ML sources of the Isabelle system for +Constructive Type Theory (extensional equality, no universes). Important +files include + +ROOT.ML -- loads all source files. Enter an ML image containing Pure +Isabelle and type: use "ROOT.ML"; + +Makefile -- compiles the files under Poly/ML or SML of New Jersey + +ex -- subdirectory containing examples. To execute them, enter an ML image +containing CTT and type: use "ex/ROOT.ML"; + +Useful references on Constructive Type Theory: + + B. Nordstr\"om, K. Petersson and J. M. Smith, + Programming in Martin-L\"of's Type Theory + (Oxford University Press, 1990) + + Simon Thompson, + Type Theory and Functional Programming (Addison-Wesley, 1991) + diff -r 000000000000 -r a5a9c433f639 src/CTT/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,24 @@ +(* Title: CTT/ROOT + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Adds Constructive Type Theory to a database containing pure Isabelle. +Should be executed in the subdirectory CTT. +*) + +val banner = "Constructive Type Theory"; +writeln banner; + +print_depth 1; + +use_thy"ctt"; +use "../Provers/typedsimp.ML"; +use "rew.ML"; +use_thy "arith"; +use_thy "bool"; + +use "../Pure/install_pp.ML"; +print_depth 8; + +val CTT_build_completed = (); (*indicate successful build*) diff -r 000000000000 -r a5a9c433f639 src/CTT/arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/arith.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,497 @@ +(* Title: CTT/arith + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Theorems for arith.thy (Arithmetic operators) + +Proofs about elementary arithmetic: addition, multiplication, etc. +Tests definitions and simplifier. +*) + +open Arith; +val arith_defs = [add_def, diff_def, absdiff_def, mult_def, mod_def, div_def]; + + +(** Addition *) + +(*typing of add: short and long versions*) + +val add_typing = prove_goal Arith.thy + "[| a:N; b:N |] ==> a #+ b : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (typechk_tac prems) ]); + +val add_typingL = prove_goal Arith.thy + "[| a=c:N; b=d:N |] ==> a #+ b = c #+ d : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (equal_tac prems) ]); + + +(*computation for add: 0 and successor cases*) + +val addC0 = prove_goal Arith.thy + "b:N ==> 0 #+ b = b : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (rew_tac prems) ]); + +val addC_succ = prove_goal Arith.thy + "[| a:N; b:N |] ==> succ(a) #+ b = succ(a #+ b) : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (rew_tac prems) ]); + + +(** Multiplication *) + +(*typing of mult: short and long versions*) + +val mult_typing = prove_goal Arith.thy + "[| a:N; b:N |] ==> a #* b : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (typechk_tac([add_typing]@prems)) ]); + +val mult_typingL = prove_goal Arith.thy + "[| a=c:N; b=d:N |] ==> a #* b = c #* d : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (equal_tac (prems@[add_typingL])) ]); + +(*computation for mult: 0 and successor cases*) + +val multC0 = prove_goal Arith.thy + "b:N ==> 0 #* b = 0 : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (rew_tac prems) ]); + +val multC_succ = prove_goal Arith.thy + "[| a:N; b:N |] ==> succ(a) #* b = b #+ (a #* b) : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (rew_tac prems) ]); + + +(** Difference *) + +(*typing of difference*) + +val diff_typing = prove_goal Arith.thy + "[| a:N; b:N |] ==> a - b : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (typechk_tac prems) ]); + +val diff_typingL = prove_goal Arith.thy + "[| a=c:N; b=d:N |] ==> a - b = c - d : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (equal_tac prems) ]); + + + +(*computation for difference: 0 and successor cases*) + +val diffC0 = prove_goal Arith.thy + "a:N ==> a - 0 = a : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (rew_tac prems) ]); + +(*Note: rec(a, 0, %z w.z) is pred(a). *) + +val diff_0_eq_0 = prove_goal Arith.thy + "b:N ==> 0 - b = 0 : N" + (fn prems=> + [ (NE_tac "b" 1), + (rewrite_goals_tac arith_defs), + (hyp_rew_tac prems) ]); + + +(*Essential to simplify FIRST!! (Else we get a critical pair) + succ(a) - succ(b) rewrites to pred(succ(a) - b) *) +val diff_succ_succ = prove_goal Arith.thy + "[| a:N; b:N |] ==> succ(a) - succ(b) = a - b : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (hyp_rew_tac prems), + (NE_tac "b" 1), + (hyp_rew_tac prems) ]); + + + +(*** Simplification *) + +val arith_typing_rls = + [add_typing, mult_typing, diff_typing]; + +val arith_congr_rls = + [add_typingL, mult_typingL, diff_typingL]; + +val congr_rls = arith_congr_rls@standard_congr_rls; + +val arithC_rls = + [addC0, addC_succ, + multC0, multC_succ, + diffC0, diff_0_eq_0, diff_succ_succ]; + + +structure Arith_simp_data: TSIMP_DATA = + struct + val refl = refl_elem + val sym = sym_elem + val trans = trans_elem + val refl_red = refl_red + val trans_red = trans_red + val red_if_equal = red_if_equal + val default_rls = arithC_rls @ comp_rls + val routine_tac = routine_tac (arith_typing_rls @ routine_rls) + end; + +structure Arith_simp = TSimpFun (Arith_simp_data); + +fun arith_rew_tac prems = make_rew_tac + (Arith_simp.norm_tac(congr_rls, prems)); + +fun hyp_arith_rew_tac prems = make_rew_tac + (Arith_simp.cond_norm_tac(prove_cond_tac, congr_rls, prems)); + + +(********** + Addition + **********) + +(*Associative law for addition*) +val add_assoc = prove_goal Arith.thy + "[| a:N; b:N; c:N |] ==> (a #+ b) #+ c = a #+ (b #+ c) : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac prems) ]); + + +(*Commutative law for addition. Can be proved using three inductions. + Must simplify after first induction! Orientation of rewrites is delicate*) +val add_commute = prove_goal Arith.thy + "[| a:N; b:N |] ==> a #+ b = b #+ a : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac prems), + (NE_tac "b" 2), + (resolve_tac [sym_elem] 1), + (NE_tac "b" 1), + (hyp_arith_rew_tac prems) ]); + + +(**************** + Multiplication + ****************) + +(*Commutative law for multiplication +val mult_commute = prove_goal Arith.thy + "[| a:N; b:N |] ==> a #* b = b #* a : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac prems), + (NE_tac "b" 2), + (resolve_tac [sym_elem] 1), + (NE_tac "b" 1), + (hyp_arith_rew_tac prems) ]); NEEDS COMMUTATIVE MATCHING +***************) + +(*right annihilation in product*) +val mult_0_right = prove_goal Arith.thy + "a:N ==> a #* 0 = 0 : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac prems) ]); + +(*right successor law for multiplication*) +val mult_succ_right = prove_goal Arith.thy + "[| a:N; b:N |] ==> a #* succ(b) = a #+ (a #* b) : N" + (fn prems=> + [ (NE_tac "a" 1), +(*swap round the associative law of addition*) + (hyp_arith_rew_tac (prems @ [add_assoc RS sym_elem])), +(*leaves a goal involving a commutative law*) + (REPEAT (assume_tac 1 ORELSE + resolve_tac + (prems@[add_commute,mult_typingL,add_typingL]@ + intrL_rls@[refl_elem]) 1)) ]); + +(*Commutative law for multiplication*) +val mult_commute = prove_goal Arith.thy + "[| a:N; b:N |] ==> a #* b = b #* a : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac (prems @ [mult_0_right, mult_succ_right])) ]); + +(*addition distributes over multiplication*) +val add_mult_distrib = prove_goal Arith.thy + "[| a:N; b:N; c:N |] ==> (a #+ b) #* c = (a #* c) #+ (b #* c) : N" + (fn prems=> + [ (NE_tac "a" 1), +(*swap round the associative law of addition*) + (hyp_arith_rew_tac (prems @ [add_assoc RS sym_elem])) ]); + + +(*Associative law for multiplication*) +val mult_assoc = prove_goal Arith.thy + "[| a:N; b:N; c:N |] ==> (a #* b) #* c = a #* (b #* c) : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac (prems @ [add_mult_distrib])) ]); + + +(************ + Difference + ************ + +Difference on natural numbers, without negative numbers + a - b = 0 iff a<=b a - b = succ(c) iff a>b *) + +val diff_self_eq_0 = prove_goal Arith.thy + "a:N ==> a - a = 0 : N" + (fn prems=> + [ (NE_tac "a" 1), + (hyp_arith_rew_tac prems) ]); + + +(* [| c : N; 0 : N; c : N |] ==> c #+ 0 = c : N *) +val add_0_right = addC0 RSN (3, add_commute RS trans_elem); + +(*Addition is the inverse of subtraction: if b<=x then b#+(x-b) = x. + An example of induction over a quantified formula (a product). + Uses rewriting with a quantified, implicative inductive hypothesis.*) +val prems = +goal Arith.thy + "b:N ==> ?a : PROD x:N. Eq(N, b-x, 0) --> Eq(N, b #+ (x-b), x)"; +by (NE_tac "b" 1); +(*strip one "universal quantifier" but not the "implication"*) +by (resolve_tac intr_rls 3); +(*case analysis on x in + (succ(u) <= x) --> (succ(u)#+(x-succ(u)) = x) *) +by (NE_tac "x" 4 THEN assume_tac 4); +(*Prepare for simplification of types -- the antecedent succ(u)<=x *) +by (resolve_tac [replace_type] 5); +by (resolve_tac [replace_type] 4); +by (arith_rew_tac prems); +(*Solves first 0 goal, simplifies others. Two sugbgoals remain. + Both follow by rewriting, (2) using quantified induction hyp*) +by (intr_tac[]); (*strips remaining PRODs*) +by (hyp_arith_rew_tac (prems@[add_0_right])); +by (assume_tac 1); +val add_diff_inverse_lemma = result(); + + +(*Version of above with premise b-a=0 i.e. a >= b. + Using ProdE does not work -- for ?B(?a) is ambiguous. + Instead, add_diff_inverse_lemma states the desired induction scheme; + the use of RS below instantiates Vars in ProdE automatically. *) +val prems = +goal Arith.thy "[| a:N; b:N; b-a = 0 : N |] ==> b #+ (a-b) = a : N"; +by (resolve_tac [EqE] 1); +by (resolve_tac [ add_diff_inverse_lemma RS ProdE RS ProdE ] 1); +by (REPEAT (resolve_tac (prems@[EqI]) 1)); +val add_diff_inverse = result(); + + +(******************** + Absolute difference + ********************) + +(*typing of absolute difference: short and long versions*) + +val absdiff_typing = prove_goal Arith.thy + "[| a:N; b:N |] ==> a |-| b : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (typechk_tac prems) ]); + +val absdiff_typingL = prove_goal Arith.thy + "[| a=c:N; b=d:N |] ==> a |-| b = c |-| d : N" + (fn prems=> + [ (rewrite_goals_tac arith_defs), + (equal_tac prems) ]); + +val absdiff_self_eq_0 = prove_goal Arith.thy + "a:N ==> a |-| a = 0 : N" + (fn prems=> + [ (rewrite_goals_tac [absdiff_def]), + (arith_rew_tac (prems@[diff_self_eq_0])) ]); + +val absdiffC0 = prove_goal Arith.thy + "a:N ==> 0 |-| a = a : N" + (fn prems=> + [ (rewrite_goals_tac [absdiff_def]), + (hyp_arith_rew_tac prems) ]); + + +val absdiff_succ_succ = prove_goal Arith.thy + "[| a:N; b:N |] ==> succ(a) |-| succ(b) = a |-| b : N" + (fn prems=> + [ (rewrite_goals_tac [absdiff_def]), + (hyp_arith_rew_tac prems) ]); + +(*Note how easy using commutative laws can be? ...not always... *) +val prems = goal Arith.thy "[| a:N; b:N |] ==> a |-| b = b |-| a : N"; +by (rewrite_goals_tac [absdiff_def]); +by (resolve_tac [add_commute] 1); +by (typechk_tac ([diff_typing]@prems)); +val absdiff_commute = result(); + +(*If a+b=0 then a=0. Surprisingly tedious*) +val prems = +goal Arith.thy "[| a:N; b:N |] ==> ?c : PROD u: Eq(N,a#+b,0) . Eq(N,a,0)"; +by (NE_tac "a" 1); +by (resolve_tac [replace_type] 3); +by (arith_rew_tac prems); +by (intr_tac[]); (*strips remaining PRODs*) +by (resolve_tac [ zero_ne_succ RS FE ] 2); +by (etac (EqE RS sym_elem) 3); +by (typechk_tac ([add_typing] @prems)); +val add_eq0_lemma = result(); + +(*Version of above with the premise a+b=0. + Again, resolution instantiates variables in ProdE *) +val prems = +goal Arith.thy "[| a:N; b:N; a #+ b = 0 : N |] ==> a = 0 : N"; +by (resolve_tac [EqE] 1); +by (resolve_tac [add_eq0_lemma RS ProdE] 1); +by (resolve_tac [EqI] 3); +by (ALLGOALS (resolve_tac prems)); +val add_eq0 = result(); + +(*Here is a lemma to infer a-b=0 and b-a=0 from a|-|b=0, below. *) +val prems = goal Arith.thy + "[| a:N; b:N; a |-| b = 0 : N |] ==> \ +\ ?a : SUM v: Eq(N, a-b, 0) . Eq(N, b-a, 0)"; +by (intr_tac[]); +by eqintr_tac; +by (resolve_tac [add_eq0] 2); +by (resolve_tac [add_eq0] 1); +by (resolve_tac [add_commute RS trans_elem] 6); +by (typechk_tac (diff_typing:: map (rewrite_rule [absdiff_def]) prems)); +val absdiff_eq0_lem = result(); + +(*if a |-| b = 0 then a = b + proof: a-b=0 and b-a=0, so b = a+(b-a) = a+0 = a*) +val prems = +goal Arith.thy "[| a |-| b = 0 : N; a:N; b:N |] ==> a = b : N"; +by (resolve_tac [EqE] 1); +by (resolve_tac [absdiff_eq0_lem RS SumE] 1); +by (TRYALL (resolve_tac prems)); +by eqintr_tac; +by (resolve_tac [add_diff_inverse RS sym_elem RS trans_elem] 1); +by (resolve_tac [EqE] 3 THEN assume_tac 3); +by (hyp_arith_rew_tac (prems@[add_0_right])); +val absdiff_eq0 = result(); + +(*********************** + Remainder and Quotient + ***********************) + +(*typing of remainder: short and long versions*) + +val mod_typing = prove_goal Arith.thy + "[| a:N; b:N |] ==> a mod b : N" + (fn prems=> + [ (rewrite_goals_tac [mod_def]), + (typechk_tac (absdiff_typing::prems)) ]); + +val mod_typingL = prove_goal Arith.thy + "[| a=c:N; b=d:N |] ==> a mod b = c mod d : N" + (fn prems=> + [ (rewrite_goals_tac [mod_def]), + (equal_tac (prems@[absdiff_typingL])) ]); + + +(*computation for mod : 0 and successor cases*) + +val modC0 = prove_goal Arith.thy "b:N ==> 0 mod b = 0 : N" + (fn prems=> + [ (rewrite_goals_tac [mod_def]), + (rew_tac(absdiff_typing::prems)) ]); + +val modC_succ = prove_goal Arith.thy +"[| a:N; b:N |] ==> succ(a) mod b = rec(succ(a mod b) |-| b, 0, %x y.succ(a mod b)) : N" + (fn prems=> + [ (rewrite_goals_tac [mod_def]), + (rew_tac(absdiff_typing::prems)) ]); + + +(*typing of quotient: short and long versions*) + +val div_typing = prove_goal Arith.thy "[| a:N; b:N |] ==> a div b : N" + (fn prems=> + [ (rewrite_goals_tac [div_def]), + (typechk_tac ([absdiff_typing,mod_typing]@prems)) ]); + +val div_typingL = prove_goal Arith.thy + "[| a=c:N; b=d:N |] ==> a div b = c div d : N" + (fn prems=> + [ (rewrite_goals_tac [div_def]), + (equal_tac (prems @ [absdiff_typingL, mod_typingL])) ]); + +val div_typing_rls = [mod_typing, div_typing, absdiff_typing]; + + +(*computation for quotient: 0 and successor cases*) + +val divC0 = prove_goal Arith.thy "b:N ==> 0 div b = 0 : N" + (fn prems=> + [ (rewrite_goals_tac [div_def]), + (rew_tac([mod_typing, absdiff_typing] @ prems)) ]); + +val divC_succ = +prove_goal Arith.thy "[| a:N; b:N |] ==> succ(a) div b = \ +\ rec(succ(a) mod b, succ(a div b), %x y. a div b) : N" + (fn prems=> + [ (rewrite_goals_tac [div_def]), + (rew_tac([mod_typing]@prems)) ]); + + +(*Version of above with same condition as the mod one*) +val divC_succ2 = prove_goal Arith.thy + "[| a:N; b:N |] ==> \ +\ succ(a) div b =rec(succ(a mod b) |-| b, succ(a div b), %x y. a div b) : N" + (fn prems=> + [ (resolve_tac [ divC_succ RS trans_elem ] 1), + (rew_tac(div_typing_rls @ prems @ [modC_succ])), + (NE_tac "succ(a mod b)|-|b" 1), + (rew_tac ([mod_typing, div_typing, absdiff_typing] @prems)) ]); + +(*for case analysis on whether a number is 0 or a successor*) +val iszero_decidable = prove_goal Arith.thy + "a:N ==> rec(a, inl(eq), %ka kb.inr()) : \ +\ Eq(N,a,0) + (SUM x:N. Eq(N,a, succ(x)))" + (fn prems=> + [ (NE_tac "a" 1), + (resolve_tac [PlusI_inr] 3), + (resolve_tac [PlusI_inl] 2), + eqintr_tac, + (equal_tac prems) ]); + +(*Main Result. Holds when b is 0 since a mod 0 = a and a div 0 = 0 *) +val prems = +goal Arith.thy "[| a:N; b:N |] ==> a mod b #+ (a div b) #* b = a : N"; +by (NE_tac "a" 1); +by (arith_rew_tac (div_typing_rls@prems@[modC0,modC_succ,divC0,divC_succ2])); +by (resolve_tac [EqE] 1); +(*case analysis on succ(u mod b)|-|b *) +by (res_inst_tac [("a1", "succ(u mod b) |-| b")] + (iszero_decidable RS PlusE) 1); +by (etac SumE 3); +by (hyp_arith_rew_tac (prems @ div_typing_rls @ + [modC0,modC_succ, divC0, divC_succ2])); +(*Replace one occurence of b by succ(u mod b). Clumsy!*) +by (resolve_tac [ add_typingL RS trans_elem ] 1); +by (eresolve_tac [EqE RS absdiff_eq0 RS sym_elem] 1); +by (resolve_tac [refl_elem] 3); +by (hyp_arith_rew_tac (prems @ div_typing_rls)); +val mod_div_equality = result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/CTT/arith.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/arith.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,24 @@ +(* Title: CTT/arith + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Arithmetic operators and their definitions + +Proves about elementary arithmetic: addition, multiplication, etc. +Tests definitions and simplifier. +*) + +Arith = CTT + + +consts "#+","-","|-|" :: "[i,i]=>i" (infixr 65) + "#*",div,mod :: "[i,i]=>i" (infixr 70) + +rules + add_def "a#+b == rec(a, b, %u v.succ(v))" + diff_def "a-b == rec(b, a, %u v.rec(v, 0, %x y.x))" + absdiff_def "a|-|b == (a-b) #+ (b-a)" + mult_def "a#*b == rec(a, 0, %u v. b #+ v)" + mod_def "a mod b == rec(a, 0, %u v. rec(succ(v) |-| b, 0, %x y.succ(v)))" + div_def "a div b == rec(a, 0, %u v. rec(succ(u) mod b, succ(v), %x y.v))" +end diff -r 000000000000 -r a5a9c433f639 src/CTT/bool.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/bool.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,79 @@ +(* Title: CTT/bool + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Theorems for bool.thy (booleans and conditionals) +*) + +open Bool; + +val bool_defs = [Bool_def,true_def,false_def,cond_def]; + +(*Derivation of rules for the type Bool*) + +(*formation rule*) +val boolF = prove_goal Bool.thy + "Bool type" + (fn prems=> + [ (rewrite_goals_tac bool_defs), + (typechk_tac[]) ]); + + +(*introduction rules for true, false*) + +val boolI_true = prove_goal Bool.thy + "true : Bool" + (fn prems=> + [ (rewrite_goals_tac bool_defs), + (typechk_tac[]) ]); + +val boolI_false = prove_goal Bool.thy + "false : Bool" + (fn prems=> + [ (rewrite_goals_tac bool_defs), + (typechk_tac[]) ]); + +(*elimination rule: typing of cond*) +val boolE = prove_goal Bool.thy + "[| p:Bool; a : C(true); b : C(false) |] ==> cond(p,a,b) : C(p)" + (fn prems=> + [ (cut_facts_tac prems 1), + (rewrite_goals_tac bool_defs), + (typechk_tac prems), + (ALLGOALS (etac TE)), + (typechk_tac prems) ]); + +val boolEL = prove_goal Bool.thy + "[| p = q : Bool; a = c : C(true); b = d : C(false) |] ==> \ +\ cond(p,a,b) = cond(q,c,d) : C(p)" + (fn prems=> + [ (cut_facts_tac prems 1), + (rewrite_goals_tac bool_defs), + (resolve_tac [PlusEL] 1), + (REPEAT (eresolve_tac [asm_rl, refl_elem RS TEL] 1)) ]); + +(*computation rules for true, false*) + +val boolC_true = prove_goal Bool.thy + "[| a : C(true); b : C(false) |] ==> cond(true,a,b) = a : C(true)" + (fn prems=> + [ (cut_facts_tac prems 1), + (rewrite_goals_tac bool_defs), + (resolve_tac comp_rls 1), + (typechk_tac[]), + (ALLGOALS (etac TE)), + (typechk_tac prems) ]); + +val boolC_false = prove_goal Bool.thy + "[| a : C(true); b : C(false) |] ==> cond(false,a,b) = b : C(false)" + (fn prems=> + [ (cut_facts_tac prems 1), + (rewrite_goals_tac bool_defs), + (resolve_tac comp_rls 1), + (typechk_tac[]), + (ALLGOALS (etac TE)), + (typechk_tac prems) ]); + +writeln"Reached end of file."; + diff -r 000000000000 -r a5a9c433f639 src/CTT/bool.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/bool.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,19 @@ +(* Title: CTT/bool + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +The two-element type (booleans and conditionals) +*) + +Bool = CTT + + +consts Bool :: "t" + true,false :: "i" + cond :: "[i,i,i]=>i" +rules + Bool_def "Bool == T+T" + true_def "true == inl(tt)" + false_def "false == inr(tt)" + cond_def "cond(a,b,c) == when(a, %u.b, %u.c)" +end diff -r 000000000000 -r a5a9c433f639 src/CTT/ctt.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/ctt.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,249 @@ +(* Title: CTT/ctt.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Tactics and lemmas for ctt.thy (Constructive Type Theory) +*) + +open CTT; + +signature CTT_RESOLVE = + sig + val add_mp_tac: int -> tactic + val ASSUME: (int -> tactic) -> int -> tactic + val basic_defs: thm list + val comp_rls: thm list + val element_rls: thm list + val elimL_rls: thm list + val elim_rls: thm list + val eqintr_tac: tactic + val equal_tac: thm list -> tactic + val formL_rls: thm list + val form_rls: thm list + val form_tac: tactic + val intrL2_rls: thm list + val intrL_rls: thm list + val intr_rls: thm list + val intr_tac: thm list -> tactic + val mp_tac: int -> tactic + val NE_tac: string -> int -> tactic + val pc_tac: thm list -> int -> tactic + val PlusE_tac: string -> int -> tactic + val reduction_rls: thm list + val replace_type: thm + val routine_rls: thm list + val routine_tac: thm list -> thm list -> int -> tactic + val safe_brls: (bool * thm) list + val safestep_tac: thm list -> int -> tactic + val safe_tac: thm list -> int -> tactic + val step_tac: thm list -> int -> tactic + val subst_eqtyparg: thm + val subst_prodE: thm + val SumE_fst: thm + val SumE_snd: thm + val SumE_tac: string -> int -> tactic + val SumIL2: thm + val test_assume_tac: int -> tactic + val typechk_tac: thm list -> tactic + val unsafe_brls: (bool * thm) list + end; + + +structure CTT_Resolve : CTT_RESOLVE = +struct + +(*Formation rules*) +val form_rls = [NF, ProdF, SumF, PlusF, EqF, FF, TF] +and formL_rls = [ProdFL, SumFL, PlusFL, EqFL]; + + +(*Introduction rules + OMITTED: EqI, because its premise is an eqelem, not an elem*) +val intr_rls = [NI0, NI_succ, ProdI, SumI, PlusI_inl, PlusI_inr, TI] +and intrL_rls = [NI_succL, ProdIL, SumIL, PlusI_inlL, PlusI_inrL]; + + +(*Elimination rules + OMITTED: EqE, because its conclusion is an eqelem, not an elem + TE, because it does not involve a constructor *) +val elim_rls = [NE, ProdE, SumE, PlusE, FE] +and elimL_rls = [NEL, ProdEL, SumEL, PlusEL, FEL]; + +(*OMITTED: eqC are TC because they make rewriting loop: p = un = un = ... *) +val comp_rls = [NC0, NC_succ, ProdC, SumC, PlusC_inl, PlusC_inr]; + +(*rules with conclusion a:A, an elem judgement*) +val element_rls = intr_rls @ elim_rls; + +(*Definitions are (meta)equality axioms*) +val basic_defs = [fst_def,snd_def]; + +(*Compare with standard version: B is applied to UNSIMPLIFIED expression! *) +val SumIL2 = prove_goal CTT.thy + "[| c=a : A; d=b : B(a) |] ==> = : Sum(A,B)" + (fn prems=> + [ (resolve_tac [sym_elem] 1), + (resolve_tac [SumIL] 1), + (ALLGOALS (resolve_tac [sym_elem])), + (ALLGOALS (resolve_tac prems)) ]); + +val intrL2_rls = [NI_succL, ProdIL, SumIL2, PlusI_inlL, PlusI_inrL]; + +(*Exploit p:Prod(A,B) to create the assumption z:B(a). + A more natural form of product elimination. *) +val subst_prodE = prove_goal CTT.thy + "[| p: Prod(A,B); a: A; !!z. z: B(a) ==> c(z): C(z) \ +\ |] ==> c(p`a): C(p`a)" + (fn prems=> + [ (REPEAT (resolve_tac (prems@[ProdE]) 1)) ]); + +(** Tactics for type checking **) + +fun is_rigid_elem (Const("Elem",_) $ a $ _) = not (is_Var (head_of a)) + | is_rigid_elem _ = false; + +(*Try solving a:A by assumption provided a is rigid!*) +val test_assume_tac = SUBGOAL(fn (prem,i) => + if is_rigid_elem (Logic.strip_assums_concl prem) + then assume_tac i else no_tac); + +fun ASSUME tf i = test_assume_tac i ORELSE tf i; + + +(*For simplification: type formation and checking, + but no equalities between terms*) +val routine_rls = form_rls @ formL_rls @ [refl_type] @ element_rls; + +fun routine_tac rls prems = ASSUME (filt_resolve_tac (prems @ rls) 4); + + +(*Solve all subgoals "A type" using formation rules. *) +val form_tac = REPEAT_FIRST (filt_resolve_tac(form_rls) 1); + + +(*Type checking: solve a:A (a rigid, A flexible) by intro and elim rules. *) +fun typechk_tac thms = + let val tac = filt_resolve_tac (thms @ form_rls @ element_rls) 3 + in REPEAT_FIRST (ASSUME tac) end; + + +(*Solve a:A (a flexible, A rigid) by introduction rules. + Cannot use stringtrees (filt_resolve_tac) since + goals like ?a:SUM(A,B) have a trivial head-string *) +fun intr_tac thms = + let val tac = filt_resolve_tac(thms@form_rls@intr_rls) 1 + in REPEAT_FIRST (ASSUME tac) end; + + +(*Equality proving: solve a=b:A (where a is rigid) by long rules. *) +fun equal_tac thms = + let val rls = thms @ form_rls @ element_rls @ intrL_rls @ + elimL_rls @ [refl_elem] + in REPEAT_FIRST (ASSUME (filt_resolve_tac rls 3)) end; + +(*** Simplification ***) + +(*To simplify the type in a goal*) +val replace_type = prove_goal CTT.thy + "[| B = A; a : A |] ==> a : B" + (fn prems=> + [ (resolve_tac [equal_types] 1), + (resolve_tac [sym_type] 2), + (ALLGOALS (resolve_tac prems)) ]); + +(*Simplify the parameter of a unary type operator.*) +val subst_eqtyparg = prove_goal CTT.thy + "a=c : A ==> (!!z.z:A ==> B(z) type) ==> B(a)=B(c)" + (fn prems=> + [ (resolve_tac [subst_typeL] 1), + (resolve_tac [refl_type] 2), + (ALLGOALS (resolve_tac prems)), + (assume_tac 1) ]); + +(*Make a reduction rule for simplification. + A goal a=c becomes b=c, by virtue of a=b *) +fun resolve_trans rl = rl RS trans_elem; + +(*Simplification rules for Constructive Type Theory*) +val reduction_rls = map resolve_trans comp_rls; + +(*Converts each goal "e : Eq(A,a,b)" into "a=b:A" for simplification. + Uses other intro rules to avoid changing flexible goals.*) +val eqintr_tac = REPEAT_FIRST (ASSUME (filt_resolve_tac(EqI::intr_rls) 1)); + +(** Tactics that instantiate CTT-rules. + Vars in the given terms will be incremented! + The (resolve_tac [EqE] i) lets them apply to equality judgements. **) + +fun NE_tac (sp: string) i = + TRY (resolve_tac [EqE] i) THEN res_inst_tac [ ("p",sp) ] NE i; + +fun SumE_tac (sp: string) i = + TRY (resolve_tac [EqE] i) THEN res_inst_tac [ ("p",sp) ] SumE i; + +fun PlusE_tac (sp: string) i = + TRY (resolve_tac [EqE] i) THEN res_inst_tac [ ("p",sp) ] PlusE i; + +(** Predicate logic reasoning, WITH THINNING!! Procedures adapted from NJ. **) + +(*Finds f:Prod(A,B) and a:A in the assumptions, concludes there is z:B(a) *) +fun add_mp_tac i = + resolve_tac [subst_prodE] i THEN assume_tac i THEN assume_tac i; + +(*Finds P-->Q and P in the assumptions, replaces implication by Q *) +fun mp_tac i = eresolve_tac [subst_prodE] i THEN assume_tac i; + +(*"safe" when regarded as predicate calculus rules*) +val safe_brls = sort lessb + [ (true,FE), (true,asm_rl), + (false,ProdI), (true,SumE), (true,PlusE) ]; + +val unsafe_brls = + [ (false,PlusI_inl), (false,PlusI_inr), (false,SumI), + (true,subst_prodE) ]; + +(*0 subgoals vs 1 or more*) +val (safe0_brls, safep_brls) = + partition (apl(0,op=) o subgoals_of_brl) safe_brls; + +fun safestep_tac thms i = + form_tac ORELSE + resolve_tac thms i ORELSE + biresolve_tac safe0_brls i ORELSE mp_tac i ORELSE + DETERM (biresolve_tac safep_brls i); + +fun safe_tac thms i = DEPTH_SOLVE_1 (safestep_tac thms i); + +fun step_tac thms = safestep_tac thms ORELSE' biresolve_tac unsafe_brls; + +(*Fails unless it solves the goal!*) +fun pc_tac thms = DEPTH_SOLVE_1 o (step_tac thms); + +(** The elimination rules for fst/snd **) + +val SumE_fst = prove_goal CTT.thy + "p : Sum(A,B) ==> fst(p) : A" + (fn prems=> + [ (rewrite_goals_tac basic_defs), + (resolve_tac elim_rls 1), + (REPEAT (pc_tac prems 1)), + (fold_tac basic_defs) ]); + +(*The first premise must be p:Sum(A,B) !!*) +val SumE_snd = prove_goal CTT.thy + "[| p: Sum(A,B); A type; !!x. x:A ==> B(x) type \ +\ |] ==> snd(p) : B(fst(p))" + (fn prems=> + [ (rewrite_goals_tac basic_defs), + (resolve_tac elim_rls 1), + (resolve_tac prems 1), + (resolve_tac [replace_type] 1), + (resolve_tac [subst_eqtyparg] 1), (*like B(x) equality formation?*) + (resolve_tac comp_rls 1), + (typechk_tac prems), + (fold_tac basic_defs) ]); + +end; + +open CTT_Resolve; diff -r 000000000000 -r a5a9c433f639 src/CTT/ctt.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/ctt.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,253 @@ +(* Title: CTT/ctt.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Constructive Type Theory +*) + +CTT = Pure + + +types i,t,o 0 + +arities i,t,o :: logic + +consts + (*Types*) + F,T :: "t" (*F is empty, T contains one element*) + contr :: "i=>i" + tt :: "i" + (*Natural numbers*) + N :: "t" + succ :: "i=>i" + rec :: "[i, i, [i,i]=>i] => i" + (*Unions*) + inl,inr :: "i=>i" + when :: "[i, i=>i, i=>i]=>i" + (*General Sum and Binary Product*) + Sum :: "[t, i=>t]=>t" + fst,snd :: "i=>i" + split :: "[i, [i,i]=>i] =>i" + (*General Product and Function Space*) + Prod :: "[t, i=>t]=>t" + (*Equality type*) + Eq :: "[t,i,i]=>t" + eq :: "i" + (*Judgements*) + Type :: "t => prop" ("(_ type)" [10] 5) + Eqtype :: "[t,t]=>prop" ("(3_ =/ _)" [10,10] 5) + Elem :: "[i, t]=>prop" ("(_ /: _)" [10,10] 5) + Eqelem :: "[i,i,t]=>prop" ("(3_ =/ _ :/ _)" [10,10,10] 5) + Reduce :: "[i,i]=>prop" ("Reduce[_,_]") + (*Types*) + "@PROD" :: "[id,t,t]=>t" ("(3PROD _:_./ _)" 10) + "@SUM" :: "[id,t,t]=>t" ("(3SUM _:_./ _)" 10) + "+" :: "[t,t]=>t" (infixr 40) + (*Invisible infixes!*) + "@-->" :: "[t,t]=>t" ("(_ -->/ _)" [31,30] 30) + "@*" :: "[t,t]=>t" ("(_ */ _)" [51,50] 50) + (*Functions*) + lambda :: "(i => i) => i" (binder "lam " 10) + "`" :: "[i,i]=>i" (infixl 60) + (*Natural numbers*) + "0" :: "i" ("0") + (*Pairing*) + pair :: "[i,i]=>i" ("(1<_,/_>)") + +translations + "PROD x:A. B" => "Prod(A, %x. B)" + "SUM x:A. B" => "Sum(A, %x. B)" + +rules + + (*Reduction: a weaker notion than equality; a hack for simplification. + Reduce[a,b] means either that a=b:A for some A or else that "a" and "b" + are textually identical.*) + + (*does not verify a:A! Sound because only trans_red uses a Reduce premise + No new theorems can be proved about the standard judgements.*) + refl_red "Reduce[a,a]" + red_if_equal "a = b : A ==> Reduce[a,b]" + trans_red "[| a = b : A; Reduce[b,c] |] ==> a = c : A" + + (*Reflexivity*) + + refl_type "A type ==> A = A" + refl_elem "a : A ==> a = a : A" + + (*Symmetry*) + + sym_type "A = B ==> B = A" + sym_elem "a = b : A ==> b = a : A" + + (*Transitivity*) + + trans_type "[| A = B; B = C |] ==> A = C" + trans_elem "[| a = b : A; b = c : A |] ==> a = c : A" + + equal_types "[| a : A; A = B |] ==> a : B" + equal_typesL "[| a = b : A; A = B |] ==> a = b : B" + + (*Substitution*) + + subst_type "[| a : A; !!z. z:A ==> B(z) type |] ==> B(a) type" + subst_typeL "[| a = c : A; !!z. z:A ==> B(z) = D(z) |] ==> B(a) = D(c)" + + subst_elem "[| a : A; !!z. z:A ==> b(z):B(z) |] ==> b(a):B(a)" + subst_elemL + "[| a=c : A; !!z. z:A ==> b(z)=d(z) : B(z) |] ==> b(a)=d(c) : B(a)" + + + (*The type N -- natural numbers*) + + NF "N type" + NI0 "0 : N" + NI_succ "a : N ==> succ(a) : N" + NI_succL "a = b : N ==> succ(a) = succ(b) : N" + + NE + "[| p: N; a: C(0); !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] \ +\ ==> rec(p, a, %u v.b(u,v)) : C(p)" + + NEL + "[| p = q : N; a = c : C(0); \ +\ !!u v. [| u: N; v: C(u) |] ==> b(u,v) = d(u,v): C(succ(u)) |] \ +\ ==> rec(p, a, %u v.b(u,v)) = rec(q,c,d) : C(p)" + + NC0 + "[| a: C(0); !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] \ +\ ==> rec(0, a, %u v.b(u,v)) = a : C(0)" + + NC_succ + "[| p: N; a: C(0); \ +\ !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] ==> \ +\ rec(succ(p), a, %u v.b(u,v)) = b(p, rec(p, a, %u v.b(u,v))) : C(succ(p))" + + (*The fourth Peano axiom. See page 91 of Martin-Lof's book*) + zero_ne_succ + "[| a: N; 0 = succ(a) : N |] ==> 0: F" + + + (*The Product of a family of types*) + + ProdF "[| A type; !!x. x:A ==> B(x) type |] ==> PROD x:A.B(x) type" + + ProdFL + "[| A = C; !!x. x:A ==> B(x) = D(x) |] ==> \ +\ PROD x:A.B(x) = PROD x:C.D(x)" + + ProdI + "[| A type; !!x. x:A ==> b(x):B(x)|] ==> lam x.b(x) : PROD x:A.B(x)" + + ProdIL + "[| A type; !!x. x:A ==> b(x) = c(x) : B(x)|] ==> \ +\ lam x.b(x) = lam x.c(x) : PROD x:A.B(x)" + + ProdE "[| p : PROD x:A.B(x); a : A |] ==> p`a : B(a)" + ProdEL "[| p=q: PROD x:A.B(x); a=b : A |] ==> p`a = q`b : B(a)" + + ProdC + "[| a : A; !!x. x:A ==> b(x) : B(x)|] ==> \ +\ (lam x.b(x)) ` a = b(a) : B(a)" + + ProdC2 + "p : PROD x:A.B(x) ==> (lam x. p`x) = p : PROD x:A.B(x)" + + + (*The Sum of a family of types*) + + SumF "[| A type; !!x. x:A ==> B(x) type |] ==> SUM x:A.B(x) type" + SumFL + "[| A = C; !!x. x:A ==> B(x) = D(x) |] ==> SUM x:A.B(x) = SUM x:C.D(x)" + + SumI "[| a : A; b : B(a) |] ==> : SUM x:A.B(x)" + SumIL "[| a=c:A; b=d:B(a) |] ==> = : SUM x:A.B(x)" + + SumE + "[| p: SUM x:A.B(x); !!x y. [| x:A; y:B(x) |] ==> c(x,y): C() |] \ +\ ==> split(p, %x y.c(x,y)) : C(p)" + + SumEL + "[| p=q : SUM x:A.B(x); \ +\ !!x y. [| x:A; y:B(x) |] ==> c(x,y)=d(x,y): C()|] \ +\ ==> split(p, %x y.c(x,y)) = split(q, % x y.d(x,y)) : C(p)" + + SumC + "[| a: A; b: B(a); !!x y. [| x:A; y:B(x) |] ==> c(x,y): C() |] \ +\ ==> split(, %x y.c(x,y)) = c(a,b) : C()" + + fst_def "fst(a) == split(a, %x y.x)" + snd_def "snd(a) == split(a, %x y.y)" + + + (*The sum of two types*) + + PlusF "[| A type; B type |] ==> A+B type" + PlusFL "[| A = C; B = D |] ==> A+B = C+D" + + PlusI_inl "[| a : A; B type |] ==> inl(a) : A+B" + PlusI_inlL "[| a = c : A; B type |] ==> inl(a) = inl(c) : A+B" + + PlusI_inr "[| A type; b : B |] ==> inr(b) : A+B" + PlusI_inrL "[| A type; b = d : B |] ==> inr(b) = inr(d) : A+B" + + PlusE + "[| p: A+B; !!x. x:A ==> c(x): C(inl(x)); \ +\ !!y. y:B ==> d(y): C(inr(y)) |] \ +\ ==> when(p, %x.c(x), %y.d(y)) : C(p)" + + PlusEL + "[| p = q : A+B; !!x. x: A ==> c(x) = e(x) : C(inl(x)); \ +\ !!y. y: B ==> d(y) = f(y) : C(inr(y)) |] \ +\ ==> when(p, %x.c(x), %y.d(y)) = when(q, %x.e(x), %y.f(y)) : C(p)" + + PlusC_inl + "[| a: A; !!x. x:A ==> c(x): C(inl(x)); \ +\ !!y. y:B ==> d(y): C(inr(y)) |] \ +\ ==> when(inl(a), %x.c(x), %y.d(y)) = c(a) : C(inl(a))" + + PlusC_inr + "[| b: B; !!x. x:A ==> c(x): C(inl(x)); \ +\ !!y. y:B ==> d(y): C(inr(y)) |] \ +\ ==> when(inr(b), %x.c(x), %y.d(y)) = d(b) : C(inr(b))" + + + (*The type Eq*) + + EqF "[| A type; a : A; b : A |] ==> Eq(A,a,b) type" + EqFL "[| A=B; a=c: A; b=d : A |] ==> Eq(A,a,b) = Eq(B,c,d)" + EqI "a = b : A ==> eq : Eq(A,a,b)" + EqE "p : Eq(A,a,b) ==> a = b : A" + + (*By equality of types, can prove C(p) from C(eq), an elimination rule*) + EqC "p : Eq(A,a,b) ==> p = eq : Eq(A,a,b)" + + (*The type F*) + + FF "F type" + FE "[| p: F; C type |] ==> contr(p) : C" + FEL "[| p = q : F; C type |] ==> contr(p) = contr(q) : C" + + (*The type T + Martin-Lof's book (page 68) discusses elimination and computation. + Elimination can be derived by computation and equality of types, + but with an extra premise C(x) type x:T. + Also computation can be derived from elimination. *) + + TF "T type" + TI "tt : T" + TE "[| p : T; c : C(tt) |] ==> c : C(p)" + TEL "[| p = q : T; c = d : C(tt) |] ==> c = d : C(p)" + TC "p : T ==> p = tt : T" +end + + +ML + +val parse_translation = + [("@-->", ndependent_tr "Prod"), ("@*", ndependent_tr "Sum")]; + +val print_translation = + [("Prod", dependent_tr' ("@PROD", "@-->")), + ("Sum", dependent_tr' ("@SUM", "@*"))]; + diff -r 000000000000 -r a5a9c433f639 src/CTT/ex/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/ex/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,20 @@ +(* Title: CTT/ex/ROOT + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Executes all examples for Constructive Type Theory. +*) + +CTT_build_completed; (*Cause examples to fail if CTT did*) + +writeln"Root file for CTT examples"; + +print_depth 2; +proof_timing := true; +time_use "ex/typechk.ML"; +time_use "ex/elim.ML"; +time_use "ex/equal.ML"; +time_use "ex/synth.ML"; + +maketest"END: Root file for CTT examples"; diff -r 000000000000 -r a5a9c433f639 src/CTT/ex/elim.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/ex/elim.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,188 @@ +(* Title: CTT/ex/elim + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Some examples taken from P. Martin-L\"of, Intuitionistic type theory + (Bibliopolis, 1984). + +by (safe_tac prems 1); +by (step_tac prems 1); +by (pc_tac prems 1); +*) + +writeln"Examples with elimination rules"; + + +writeln"This finds the functions fst and snd!"; +val prems = goal CTT.thy "A type ==> ?a : (A*A) --> A"; +by (pc_tac prems 1 THEN fold_tac basic_defs); (*puts in fst and snd*) +result(); +writeln"first solution is fst; backtracking gives snd"; +back(); +back() handle ERROR => writeln"And there are indeed no others"; + + +writeln"Double negation of the Excluded Middle"; +val prems = goal CTT.thy "A type ==> ?a : ((A + (A-->F)) --> F) --> F"; +by (intr_tac prems); +by (rtac ProdE 1); +by (assume_tac 1); +by (pc_tac prems 1); +result(); + +val prems = goal CTT.thy + "[| A type; B type |] ==> ?a : (A*B) --> (B*A)"; +by (pc_tac prems 1); +result(); +(*The sequent version (ITT) could produce an interesting alternative + by backtracking. No longer.*) + +writeln"Binary sums and products"; +val prems = goal CTT.thy + "[| A type; B type; C type |] ==> ?a : (A+B --> C) --> (A-->C) * (B-->C)"; +by (pc_tac prems 1); +result(); + +(*A distributive law*) +val prems = goal CTT.thy + "[| A type; B type; C type |] ==> ?a : A * (B+C) --> (A*B + A*C)"; +by (pc_tac prems 1); +result(); + +(*more general version, same proof*) +val prems = goal CTT.thy + "[| A type; !!x. x:A ==> B(x) type; !!x. x:A ==> C(x) type|] ==> \ +\ ?a : (SUM x:A. B(x) + C(x)) --> (SUM x:A. B(x)) + (SUM x:A. C(x))"; +by (pc_tac prems 1); +result(); + +writeln"Construction of the currying functional"; +val prems = goal CTT.thy + "[| A type; B type; C type |] ==> ?a : (A*B --> C) --> (A--> (B-->C))"; +by (pc_tac prems 1); +result(); + +(*more general goal with same proof*) +val prems = goal CTT.thy + "[| A type; !!x. x:A ==> B(x) type; !!z. z: (SUM x:A. B(x)) ==> C(z) type|] \ +\ ==> ?a : (PROD z : (SUM x:A . B(x)) . C(z)) \ +\ --> (PROD x:A . PROD y:B(x) . C())"; +by (pc_tac prems 1); +result(); + +writeln"Martin-Lof (1984), page 48: axiom of sum-elimination (uncurry)"; +val prems = goal CTT.thy + "[| A type; B type; C type |] ==> ?a : (A --> (B-->C)) --> (A*B --> C)"; +by (pc_tac prems 1); +result(); + +(*more general goal with same proof*) +val prems = goal CTT.thy + "[| A type; !!x. x:A ==> B(x) type; !!z. z : (SUM x:A . B(x)) ==> C(z) type|] \ +\ ==> ?a : (PROD x:A . PROD y:B(x) . C()) \ +\ --> (PROD z : (SUM x:A . B(x)) . C(z))"; +by (pc_tac prems 1); +result(); + +writeln"Function application"; +val prems = goal CTT.thy + "[| A type; B type |] ==> ?a : ((A --> B) * A) --> B"; +by (pc_tac prems 1); +result(); + +writeln"Basic test of quantifier reasoning"; +val prems = goal CTT.thy + "[| A type; B type; !!x y.[| x:A; y:B |] ==> C(x,y) type |] ==> \ +\ ?a : (SUM y:B . PROD x:A . C(x,y)) \ +\ --> (PROD x:A . SUM y:B . C(x,y))"; +by (pc_tac prems 1); +result(); + +(*faulty proof attempt, stripping the quantifiers in wrong sequence +by (intr_tac[]); +by (pc_tac prems 1); ...fails!! *) + +writeln"Martin-Lof (1984) pages 36-7: the combinator S"; +val prems = goal CTT.thy + "[| A type; !!x. x:A ==> B(x) type; \ +\ !!x y.[| x:A; y:B(x) |] ==> C(x,y) type |] \ +\ ==> ?a : (PROD x:A. PROD y:B(x). C(x,y)) \ +\ --> (PROD f: (PROD x:A. B(x)). PROD x:A. C(x, f`x))"; +by (pc_tac prems 1); +result(); + +writeln"Martin-Lof (1984) page 58: the axiom of disjunction elimination"; +val prems = goal CTT.thy + "[| A type; B type; !!z. z: A+B ==> C(z) type|] ==> \ +\ ?a : (PROD x:A. C(inl(x))) --> (PROD y:B. C(inr(y))) \ +\ --> (PROD z: A+B. C(z))"; +by (pc_tac prems 1); +result(); + +(*towards AXIOM OF CHOICE*) +val prems = goal CTT.thy + "[| A type; B type; C type |] ==> ?a : (A --> B*C) --> (A-->B) * (A-->C)"; +by (pc_tac prems 1); +by (fold_tac basic_defs); (*puts in fst and snd*) +result(); + +(*Martin-Lof (1984) page 50*) +writeln"AXIOM OF CHOICE!!! Delicate use of elimination rules"; +val prems = goal CTT.thy + "[| A type; !!x. x:A ==> B(x) type; \ +\ !!x y.[| x:A; y:B(x) |] ==> C(x,y) type|] \ +\ ==> ?a : (PROD x:A. SUM y:B(x). C(x,y)) \ +\ --> (SUM f: (PROD x:A. B(x)). PROD x:A. C(x, f`x))"; +by (intr_tac prems); +by (add_mp_tac 2); +by (add_mp_tac 1); +by (etac SumE_fst 1); +by (rtac replace_type 1); +by (rtac subst_eqtyparg 1); +by (resolve_tac comp_rls 1); +by (rtac SumE_snd 4); +by (typechk_tac (SumE_fst::prems)); +result(); + +writeln"Axiom of choice. Proof without fst, snd. Harder still!"; +val prems = goal CTT.thy + "[| A type; !!x.x:A ==> B(x) type; \ +\ !!x y.[| x:A; y:B(x) |] ==> C(x,y) type|] \ +\ ==> ?a : (PROD x:A. SUM y:B(x). C(x,y)) \ +\ --> (SUM f: (PROD x:A. B(x)). PROD x:A. C(x, f`x))"; +by (intr_tac prems); +(*Must not use add_mp_tac as subst_prodE hides the construction.*) +by (resolve_tac [ProdE RS SumE] 1 THEN assume_tac 1); +by (TRYALL assume_tac); +by (rtac replace_type 1); +by (rtac subst_eqtyparg 1); +by (resolve_tac comp_rls 1); +by (etac (ProdE RS SumE) 4); +by (typechk_tac prems); +by (rtac replace_type 1); +by (rtac subst_eqtyparg 1); +by (resolve_tac comp_rls 1); +by (typechk_tac prems); +by (assume_tac 1); +by (fold_tac basic_defs); (*puts in fst and snd*) +result(); + +writeln"Example of sequent_style deduction"; +(*When splitting z:A*B, the assumption C(z) is affected; ?a becomes + lam u. split(u,%v w.split(v,%x y.lam z. >) ` w) *) +val prems = goal CTT.thy + "[| A type; B type; !!z. z:A*B ==> C(z) type |] ==> \ +\ ?a : (SUM z:A*B. C(z)) --> (SUM u:A. SUM v:B. C())"; +by (resolve_tac intr_rls 1); +by (biresolve_tac safe_brls 2); +(*Now must convert assumption C(z) into antecedent C() *) +by (res_inst_tac [ ("a","y") ] ProdE 2); +by (typechk_tac prems); +by (rtac SumE 1 THEN assume_tac 1); +by (intr_tac[]); +by (TRYALL assume_tac); +by (typechk_tac prems); +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/CTT/ex/equal.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/ex/equal.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,85 @@ +(* Title: CTT/ex/equal + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Equality reasoning by rewriting. +*) + +val prems = +goal CTT.thy "p : Sum(A,B) ==> split(p,pair) = p : Sum(A,B)"; +by (rtac EqE 1); +by (resolve_tac elim_rls 1 THEN resolve_tac prems 1); +by (rew_tac prems); +val split_eq = result(); + +val prems = +goal CTT.thy + "[| A type; B type; p : A+B |] ==> when(p,inl,inr) = p : A + B"; +by (rtac EqE 1); +by (resolve_tac elim_rls 1 THEN resolve_tac prems 1); +by (rew_tac prems); +val when_eq = result(); + + +(*in the "rec" formulation of addition, 0+n=n *) +val prems = +goal CTT.thy "p:N ==> rec(p,0, %y z.succ(y)) = p : N"; +by (rtac EqE 1); +by (resolve_tac elim_rls 1 THEN resolve_tac prems 1); +by (rew_tac prems); +result(); + + +(*the harder version, n+0=n: recursive, uses induction hypothesis*) +val prems = +goal CTT.thy "p:N ==> rec(p,0, %y z.succ(z)) = p : N"; +by (rtac EqE 1); +by (resolve_tac elim_rls 1 THEN resolve_tac prems 1); +by (hyp_rew_tac prems); +result(); + + +(*Associativity of addition*) +val prems = +goal CTT.thy + "[| a:N; b:N; c:N |] ==> rec(rec(a, b, %x y.succ(y)), c, %x y.succ(y)) = \ +\ rec(a, rec(b, c, %x y.succ(y)), %x y.succ(y)) : N"; +by (NE_tac "a" 1); +by (hyp_rew_tac prems); +result(); + + +(*Martin-Lof (1984) page 62: pairing is surjective*) +val prems = +goal CTT.thy + "p : Sum(A,B) ==> = p : Sum(A,B)"; +by (rtac EqE 1); +by (resolve_tac elim_rls 1 THEN resolve_tac prems 1); +by (DEPTH_SOLVE_1 (rew_tac prems)); (*!!!!!!!*) +result(); + + +val prems = +goal CTT.thy "[| a : A; b : B |] ==> \ +\ (lam u. split(u, %v w.)) ` = : SUM x:B.A"; +by (rew_tac prems); +result(); + + +(*a contrived, complicated simplication, requires sum-elimination also*) +val prems = +goal CTT.thy + "(lam f. lam x. f`(f`x)) ` (lam u. split(u, %v w.)) = \ +\ lam x. x : PROD x:(SUM y:N.N). (SUM y:N.N)"; +by (resolve_tac reduction_rls 1); +by (resolve_tac intrL_rls 3); +by (rtac EqE 4); +by (rtac SumE 4 THEN assume_tac 4); +(*order of unifiers is essential here*) +by (rew_tac prems); +result(); + +writeln"Reached end of file."; +(*28 August 1988: loaded this file in 34 seconds*) +(*2 September 1988: loaded this file in 48 seconds*) diff -r 000000000000 -r a5a9c433f639 src/CTT/ex/synth.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/ex/synth.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,108 @@ +(* Title: CTT/ex/synth + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge +*) + +writeln"Synthesis examples, using a crude form of narrowing"; + + +writeln"discovery of predecessor function"; +goal CTT.thy + "?a : SUM pred:?A . Eq(N, pred`0, 0) \ +\ * (PROD n:N. Eq(N, pred ` succ(n), n))"; +by (intr_tac[]); +by eqintr_tac; +by (resolve_tac reduction_rls 3); +by (resolve_tac comp_rls 5); +by (rew_tac[]); +result(); + +writeln"the function fst as an element of a function type"; +val prems = goal CTT.thy + "A type ==> ?a: SUM f:?B . PROD i:A. PROD j:A. Eq(A, f ` , i)"; +by (intr_tac prems); +by eqintr_tac; +by (resolve_tac reduction_rls 2); +by (resolve_tac comp_rls 4); +by (typechk_tac prems); +writeln"now put in A everywhere"; +by (REPEAT (resolve_tac prems 1)); +by (fold_tac basic_defs); +result(); + +writeln"An interesting use of the eliminator, when"; +(*The early implementation of unification caused non-rigid path in occur check + See following example.*) +goal CTT.thy + "?a : PROD i:N. Eq(?A, ?b(inl(i)), <0 , i>) \ +\ * Eq(?A, ?b(inr(i)), )"; +by (intr_tac[]); +by eqintr_tac; +by (resolve_tac comp_rls 1); +by (rew_tac[]); +uresult(); + +(*Here we allow the type to depend on i. + This prevents the cycle in the first unification (no longer needed). + Requires flex-flex to preserve the dependence. + Simpler still: make ?A into a constant type N*N.*) +goal CTT.thy + "?a : PROD i:N. Eq(?A(i), ?b(inl(i)), <0 , i>) \ +\ * Eq(?A(i), ?b(inr(i)), )"; + +writeln"A tricky combination of when and split"; +(*Now handled easily, but caused great problems once*) +goal CTT.thy + "?a : PROD i:N. PROD j:N. Eq(?A, ?b(inl()), i) \ +\ * Eq(?A, ?b(inr()), j)"; +by (intr_tac[]); +by eqintr_tac; +by (resolve_tac [ PlusC_inl RS trans_elem ] 1); +by (resolve_tac comp_rls 4); +by (resolve_tac reduction_rls 7); +by (resolve_tac comp_rls 10); +by (typechk_tac[]); (*2 secs*) +by (fold_tac basic_defs); +uresult(); + +(*similar but allows the type to depend on i and j*) +goal CTT.thy + "?a : PROD i:N. PROD j:N. Eq(?A(i,j), ?b(inl()), i) \ +\ * Eq(?A(i,j), ?b(inr()), j)"; + +(*similar but specifying the type N simplifies the unification problems*) +goal CTT.thy + "?a : PROD i:N. PROD j:N. Eq(N, ?b(inl()), i) \ +\ * Eq(N, ?b(inr()), j)"; + + +writeln"Deriving the addition operator"; +goal Arith.thy + "?c : PROD n:N. Eq(N, ?f(0,n), n) \ +\ * (PROD m:N. Eq(N, ?f(succ(m), n), succ(?f(m,n))))"; +by (intr_tac[]); +by eqintr_tac; +by (resolve_tac comp_rls 1); +by (rew_tac[]); +by (fold_tac arith_defs); +result(); + +writeln"The addition function -- using explicit lambdas"; +goal Arith.thy + "?c : SUM plus : ?A . \ +\ PROD x:N. Eq(N, plus`0`x, x) \ +\ * (PROD y:N. Eq(N, plus`succ(y)`x, succ(plus`y`x)))"; +by (intr_tac[]); +by eqintr_tac; +by (resolve_tac [TSimp.split_eqn] 3); +by (SELECT_GOAL (rew_tac[]) 4); +by (resolve_tac [TSimp.split_eqn] 3); +by (SELECT_GOAL (rew_tac[]) 4); +by (res_inst_tac [("p","y")] NC_succ 3); + (** by (resolve_tac comp_rls 3); caused excessive branching **) +by (rew_tac[]); +by (fold_tac arith_defs); +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/CTT/ex/typechk.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/ex/typechk.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,82 @@ +(* Title: CTT/ex/typechk + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Easy examples: type checking and type deduction +*) + +writeln"Single-step proofs: verifying that a type is well-formed"; + +goal CTT.thy "?A type"; +by (resolve_tac form_rls 1); +result(); +writeln"getting a second solution"; +back(); +by (resolve_tac form_rls 1); +by (resolve_tac form_rls 1); +result(); + +goal CTT.thy "PROD z:?A . N + ?B(z) type"; +by (resolve_tac form_rls 1); +by (resolve_tac form_rls 1); +by (resolve_tac form_rls 1); +by (resolve_tac form_rls 1); +by (resolve_tac form_rls 1); +uresult(); + + +writeln"Multi-step proofs: Type inference"; + +goal CTT.thy "PROD w:N. N + N type"; +by form_tac; +result(); + +goal CTT.thy "<0, succ(0)> : ?A"; +by (intr_tac[]); +result(); + +goal CTT.thy "PROD w:N . Eq(?A,w,w) type"; +by (typechk_tac[]); +result(); + +goal CTT.thy "PROD x:N . PROD y:N . Eq(?A,x,y) type"; +by (typechk_tac[]); +result(); + +writeln"typechecking an application of fst"; +goal CTT.thy "(lam u. split(u, %v w.v)) ` <0, succ(0)> : ?A"; +by (typechk_tac[]); +result(); + +writeln"typechecking the predecessor function"; +goal CTT.thy "lam n. rec(n, 0, %x y.x) : ?A"; +by (typechk_tac[]); +result(); + +writeln"typechecking the addition function"; +goal CTT.thy "lam n. lam m. rec(n, m, %x y.succ(y)) : ?A"; +by (typechk_tac[]); +result(); + +(*Proofs involving arbitrary types. + For concreteness, every type variable left over is forced to be N*) +val N_tac = TRYALL (rtac NF); + +goal CTT.thy "lam w. : ?A"; +by (typechk_tac[]); +by N_tac; +result(); + +goal CTT.thy "lam x. lam y. x : ?A"; +by (typechk_tac[]); +by N_tac; +result(); + +writeln"typechecking fst (as a function object) "; +goal CTT.thy "lam i. split(i, %j k.j) : ?A"; +by (typechk_tac[]); +by N_tac; +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/CTT/rew.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/rew.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,43 @@ +(* Title: CTT/rew + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Simplifier for CTT, using Typedsimp +*) + +(*Make list of ProdE RS ProdE ... RS ProdE RS EqE + for using assumptions as rewrite rules*) +fun peEs 0 = [] + | peEs n = EqE :: map (apl(ProdE, op RS)) (peEs (n-1)); + +(*Tactic used for proving conditions for the cond_rls*) +val prove_cond_tac = eresolve_tac (peEs 5); + + +structure TSimp_data: TSIMP_DATA = + struct + val refl = refl_elem + val sym = sym_elem + val trans = trans_elem + val refl_red = refl_red + val trans_red = trans_red + val red_if_equal = red_if_equal + val default_rls = comp_rls + val routine_tac = routine_tac routine_rls + end; + +structure TSimp = TSimpFun (TSimp_data); + +val standard_congr_rls = intrL2_rls @ elimL_rls; + +(*Make a rewriting tactic from a normalization tactic*) +fun make_rew_tac ntac = + TRY eqintr_tac THEN TRYALL (resolve_tac [TSimp.split_eqn]) THEN + ntac; + +fun rew_tac thms = make_rew_tac + (TSimp.norm_tac(standard_congr_rls, thms)); + +fun hyp_rew_tac thms = make_rew_tac + (TSimp.cond_norm_tac(prove_cond_tac, standard_congr_rls, thms)); diff -r 000000000000 -r a5a9c433f639 src/Cube/Cube.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Cube/Cube.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,60 @@ +(* Title: Cube/cube + ID: $Id$ + Author: Tobias Nipkow + Copyright 1990 University of Cambridge + +For cube.thy. The systems of the Lambda-cube that extend simple types +*) + +open Cube; + +val simple = [s_b,strip_s,strip_b,app,lam_ss,pi_ss]; + +val L2_thy = extend_theory Cube.thy "L2" ([],[],[],[],[],None) +[ + ("pi_bs", "[| A:[]; !!x. x:A ==> B(x):* |] ==> Prod(A,B):*"), + + ("lam_bs", "[| A:[]; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):* |] \ +\ ==> Abs(A,f) : Prod(A,B)") +]; + +val lam_bs = get_axiom L2_thy "lam_bs"; +val pi_bs = get_axiom L2_thy "pi_bs"; + +val L2 = simple @ [lam_bs,pi_bs]; + +val Lomega_thy = extend_theory Cube.thy "Lomega" ([],[],[],[],[],None) +[ + ("pi_bb", "[| A:[]; !!x. x:A ==> B(x):[] |] ==> Prod(A,B):[]"), + + ("lam_bb", "[| A:[]; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):[] |] \ +\ ==> Abs(A,f) : Prod(A,B)") +]; + +val lam_bb = get_axiom Lomega_thy "lam_bb"; +val pi_bb = get_axiom Lomega_thy "pi_bb"; +val Lomega = simple @ [lam_bb,pi_bb]; + +val LOmega_thy = merge_theories(L2_thy,Lomega_thy); +val LOmega = simple @ [lam_bs,pi_bs,lam_bb,pi_bb]; + +val LP_thy = extend_theory Cube.thy "LP" ([],[],[],[],[],None) +[ + ("pi_sb", "[| A:*; !!x. x:A ==> B(x):[] |] ==> Prod(A,B):[]"), + + ("lam_sb", "[| A:*; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):[] |] \ +\ ==> Abs(A,f) : Prod(A,B)") +]; + +val lam_sb = get_axiom LP_thy "lam_sb"; +val pi_sb = get_axiom LP_thy "pi_sb"; +val LP = simple @ [lam_sb,pi_sb]; + +val LP2_thy = merge_theories(L2_thy,LP_thy); +val LP2 = simple @ [lam_bs,pi_bs,lam_sb,pi_sb]; + +val LPomega_thy = merge_theories(LP_thy,Lomega_thy); +val LPomega = simple @ [lam_bb,pi_bb,lam_sb,pi_sb]; + +val CC_thy = merge_theories(L2_thy,LPomega_thy); +val CC = simple @ [lam_bs,pi_bs,lam_bb,pi_bb,lam_sb,pi_sb]; diff -r 000000000000 -r a5a9c433f639 src/Cube/Cube.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Cube/Cube.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,60 @@ +(* Title: Cube/cube.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 University of Cambridge + +Barendregt's Lambda-Cube +*) + +Cube = Pure + + +types + term, context, typing 0 + +arities + term :: logic + +consts + Abs, Prod :: "[term, term => term] => term" + Trueprop :: "[context, typing] => prop" ("(_/ |- _)") + Trueprop1 :: "typing => prop" ("(_)") + MT_context :: "context" ("") + "" :: "id => context" ("_ ") + "" :: "var => context" ("_ ") + Context :: "[typing, context] => context" ("_ _") + star :: "term" ("*") + box :: "term" ("[]") + "^" :: "[term, term] => term" (infixl 20) + Lam :: "[id, term, term] => term" ("(3Lam _:_./ _)" [0, 0, 0] 10) + Pi :: "[id, term, term] => term" ("(3Pi _:_./ _)" [0, 0] 10) + "->" :: "[term, term] => term" (infixr 10) + Has_type :: "[term, term] => typing" ("(_:/ _)" [0, 0] 5) + +translations + (prop) "x:X" == (prop) "|- x:X" + "Lam x:A. B" == "Abs(A, %x. B)" + "Pi x:A. B" => "Prod(A, %x. B)" + +rules + s_b "*: []" + + strip_s "[| A:*; a:A ==> G |- x:X |] ==> a:A G |- x:X" + strip_b "[| A:[]; a:A ==> G |- x:X |] ==> a:A G |- x:X" + + app "[| F:Prod(A, B); C:A |] ==> F^C: B(C)" + + pi_ss "[| A:*; !!x. x:A ==> B(x):* |] ==> Prod(A, B):*" + + lam_ss "[| A:*; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):* |] \ +\ ==> Abs(A, f) : Prod(A, B)" + + beta "Abs(A, f)^a == f(a)" + +end + + +ML + +val parse_translation = [("op ->", ndependent_tr "Prod")]; +val print_translation = [("Prod", dependent_tr' ("Pi", "op ->"))]; + diff -r 000000000000 -r a5a9c433f639 src/Cube/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Cube/Makefile Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,43 @@ +######################################################################### +# # +# Makefile for Isabelle (Cube) # +# # +######################################################################### + +#To make the system, cd to this directory and type +# make +#To make the system and test it on standard examples, type +# make test + +#Environment variable ISABELLECOMP specifies the compiler. +#Environment variable ISABELLEBIN specifies the destination directory. +#For Poly/ML, ISABELLEBIN must begin with a / +#WARNING: Poly/ML databases may fail if copied or moved! + +#Makes pure Isabelle (Pure) if this file is ABSENT -- but not +#if it is out of date, since this Makefile does not know its dependencies! + +BIN = $(ISABELLEBIN) +COMP = $(ISABELLECOMP) +FILES = ROOT.ML cube.thy cube.ML ex.ML + +$(BIN)/Cube: $(BIN)/Pure $(FILES) + case "$(COMP)" in \ + poly*) echo 'make_database"$(BIN)/Cube"; quit();' \ + | $(COMP) $(BIN)/Pure;\ + echo 'open PolyML; use"ROOT";' | $(COMP) $(BIN)/Cube ;;\ + sml*) echo 'use"ROOT.ML"; xML"$(BIN)/Cube" banner;' | $(BIN)/Pure ;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +$(BIN)/Pure: + cd ../Pure; $(MAKE) + +test: ex.ML $(BIN)/Cube + case "$(COMP)" in \ + poly*) echo 'use"ex.ML"; quit();' | $(COMP) $(BIN)/Cube ;;\ + sml*) echo 'use"ex.ML";' | $(BIN)/Cube;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +.PRECIOUS: $(BIN)/Pure $(BIN)/Cube diff -r 000000000000 -r a5a9c433f639 src/Cube/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Cube/README Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,21 @@ + Cube: Barendregt's Lambda-Cube + +This directory contains the Standard ML sources of the Isabelle system for +the Lambda-Cube. Important files include + +ROOT.ML -- loads all source files. Enter an ML image containing Pure +Isabelle and type: use "ROOT.ML"; + +Makefile -- compiles the files under Poly/ML or SML of New Jersey + +ex.ML -- examples file. To execute them, enter an ML image +containing Cube and type: use "ex.ML"; + +NB: the formalization is not completely sound! It does not enforce +distinctness of variable names in contexts! + +For more information about the Lambda-Cube, see + H. Barendregt + Introduction to Generalised Type Systems + J. Functional Programming + diff -r 000000000000 -r a5a9c433f639 src/Cube/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Cube/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,18 @@ +(* Title: Cube/ROOT + ID: $Id$ + Author: Tobias Nipkow + Copyright 1992 University of Cambridge + +The Lambda-Cube a la Barendregt +*) + +val banner = "Barendregt's Lambda-Cube"; +writeln banner; + +print_depth 1; +use_thy "cube"; + +use "../Pure/install_pp.ML"; +print_depth 8; + +val Cube_build_completed = (); (*indicate successful build*) diff -r 000000000000 -r a5a9c433f639 src/Cube/cube.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Cube/cube.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,60 @@ +(* Title: Cube/cube + ID: $Id$ + Author: Tobias Nipkow + Copyright 1990 University of Cambridge + +For cube.thy. The systems of the Lambda-cube that extend simple types +*) + +open Cube; + +val simple = [s_b,strip_s,strip_b,app,lam_ss,pi_ss]; + +val L2_thy = extend_theory Cube.thy "L2" ([],[],[],[],[],None) +[ + ("pi_bs", "[| A:[]; !!x. x:A ==> B(x):* |] ==> Prod(A,B):*"), + + ("lam_bs", "[| A:[]; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):* |] \ +\ ==> Abs(A,f) : Prod(A,B)") +]; + +val lam_bs = get_axiom L2_thy "lam_bs"; +val pi_bs = get_axiom L2_thy "pi_bs"; + +val L2 = simple @ [lam_bs,pi_bs]; + +val Lomega_thy = extend_theory Cube.thy "Lomega" ([],[],[],[],[],None) +[ + ("pi_bb", "[| A:[]; !!x. x:A ==> B(x):[] |] ==> Prod(A,B):[]"), + + ("lam_bb", "[| A:[]; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):[] |] \ +\ ==> Abs(A,f) : Prod(A,B)") +]; + +val lam_bb = get_axiom Lomega_thy "lam_bb"; +val pi_bb = get_axiom Lomega_thy "pi_bb"; +val Lomega = simple @ [lam_bb,pi_bb]; + +val LOmega_thy = merge_theories(L2_thy,Lomega_thy); +val LOmega = simple @ [lam_bs,pi_bs,lam_bb,pi_bb]; + +val LP_thy = extend_theory Cube.thy "LP" ([],[],[],[],[],None) +[ + ("pi_sb", "[| A:*; !!x. x:A ==> B(x):[] |] ==> Prod(A,B):[]"), + + ("lam_sb", "[| A:*; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):[] |] \ +\ ==> Abs(A,f) : Prod(A,B)") +]; + +val lam_sb = get_axiom LP_thy "lam_sb"; +val pi_sb = get_axiom LP_thy "pi_sb"; +val LP = simple @ [lam_sb,pi_sb]; + +val LP2_thy = merge_theories(L2_thy,LP_thy); +val LP2 = simple @ [lam_bs,pi_bs,lam_sb,pi_sb]; + +val LPomega_thy = merge_theories(LP_thy,Lomega_thy); +val LPomega = simple @ [lam_bb,pi_bb,lam_sb,pi_sb]; + +val CC_thy = merge_theories(L2_thy,LPomega_thy); +val CC = simple @ [lam_bs,pi_bs,lam_bb,pi_bb,lam_sb,pi_sb]; diff -r 000000000000 -r a5a9c433f639 src/Cube/cube.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Cube/cube.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,60 @@ +(* Title: Cube/cube.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 University of Cambridge + +Barendregt's Lambda-Cube +*) + +Cube = Pure + + +types + term, context, typing 0 + +arities + term :: logic + +consts + Abs, Prod :: "[term, term => term] => term" + Trueprop :: "[context, typing] => prop" ("(_/ |- _)") + Trueprop1 :: "typing => prop" ("(_)") + MT_context :: "context" ("") + "" :: "id => context" ("_ ") + "" :: "var => context" ("_ ") + Context :: "[typing, context] => context" ("_ _") + star :: "term" ("*") + box :: "term" ("[]") + "^" :: "[term, term] => term" (infixl 20) + Lam :: "[id, term, term] => term" ("(3Lam _:_./ _)" [0, 0, 0] 10) + Pi :: "[id, term, term] => term" ("(3Pi _:_./ _)" [0, 0] 10) + "->" :: "[term, term] => term" (infixr 10) + Has_type :: "[term, term] => typing" ("(_:/ _)" [0, 0] 5) + +translations + (prop) "x:X" == (prop) "|- x:X" + "Lam x:A. B" == "Abs(A, %x. B)" + "Pi x:A. B" => "Prod(A, %x. B)" + +rules + s_b "*: []" + + strip_s "[| A:*; a:A ==> G |- x:X |] ==> a:A G |- x:X" + strip_b "[| A:[]; a:A ==> G |- x:X |] ==> a:A G |- x:X" + + app "[| F:Prod(A, B); C:A |] ==> F^C: B(C)" + + pi_ss "[| A:*; !!x. x:A ==> B(x):* |] ==> Prod(A, B):*" + + lam_ss "[| A:*; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):* |] \ +\ ==> Abs(A, f) : Prod(A, B)" + + beta "Abs(A, f)^a == f(a)" + +end + + +ML + +val parse_translation = [("op ->", ndependent_tr "Prod")]; +val print_translation = [("Prod", dependent_tr' ("Pi", "op ->"))]; + diff -r 000000000000 -r a5a9c433f639 src/Cube/ex.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Cube/ex.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,240 @@ +(* Examples taken from + H. Barendregt. Introduction to Generalised Type Systems. + J. Functional Programming. +*) + +Cube_build_completed; (*Cause examples to fail if Cube did*) + +proof_timing := true; + +fun strip_asms_tac thms i = + REPEAT(resolve_tac[strip_b,strip_s]i THEN DEPTH_SOLVE_1(ares_tac thms i)); + +val imp_elim = prove_goal thy "[| f:A->B; a:A; f^a:B ==> PROP P |] ==> PROP P" + (fn asms => [REPEAT(resolve_tac (app::asms) 1)]); + +val pi_elim = prove_goal thy + "[| F:Prod(A,B); a:A; F^a:B(a) ==> PROP P |] ==> PROP P" + (fn asms => [REPEAT(resolve_tac (app::asms) 1)]); + +(* SIMPLE TYPES *) + +goal thy "A:* |- A->A : ?T"; +by (DEPTH_SOLVE (ares_tac simple 1)); +uresult(); + +goal thy "A:* |- Lam a:A.a : ?T"; +by (DEPTH_SOLVE (ares_tac simple 1)); +uresult(); + +goal thy "A:* B:* b:B |- Lam x:A.b : ?T"; +by (DEPTH_SOLVE (ares_tac simple 1)); +uresult(); + +goal thy "A:* b:A |- (Lam a:A.a)^b: ?T"; +by (DEPTH_SOLVE (ares_tac simple 1)); +uresult(); + +goal thy "A:* B:* c:A b:B |- (Lam x:A.b)^ c: ?T"; +by (DEPTH_SOLVE (ares_tac simple 1)); +uresult(); + +goal thy "A:* B:* |- Lam a:A.Lam b:B.a : ?T"; +by (DEPTH_SOLVE (ares_tac simple 1)); +uresult(); + +(* SECOND-ORDER TYPES *) + +goal L2_thy "|- Lam A:*. Lam a:A.a : ?T"; +by (DEPTH_SOLVE (ares_tac L2 1)); +uresult(); + +goal L2_thy "A:* |- (Lam B:*.Lam b:B.b)^A : ?T"; +by (DEPTH_SOLVE (ares_tac L2 1)); +uresult(); + +goal L2_thy "A:* b:A |- (Lam B:*.Lam b:B.b) ^ A ^ b: ?T"; +by (DEPTH_SOLVE (ares_tac L2 1)); +uresult(); + +goal L2_thy "|- Lam B:*.Lam a:(Pi A:*.A).a ^ ((Pi A:*.A)->B) ^ a: ?T"; +by (DEPTH_SOLVE (ares_tac L2 1)); +uresult(); + +(* Weakly higher-order proposiional logic *) + +goal Lomega_thy "|- Lam A:*.A->A : ?T"; +by (DEPTH_SOLVE (ares_tac Lomega 1)); +uresult(); + +goal Lomega_thy "B:* |- (Lam A:*.A->A) ^ B : ?T"; +by (DEPTH_SOLVE (ares_tac Lomega 1)); +uresult(); + +goal Lomega_thy "B:* b:B |- (Lam y:B.b): ?T"; +by (DEPTH_SOLVE (ares_tac Lomega 1)); +uresult(); + +goal Lomega_thy "A:* F:*->* |- F^(F^A): ?T"; +by (DEPTH_SOLVE (ares_tac Lomega 1)); +uresult(); + +goal Lomega_thy "A:* |- Lam F:*->*.F^(F^A): ?T"; +by (DEPTH_SOLVE (ares_tac Lomega 1)); +uresult(); + +(* LF *) + +goal LP_thy "A:* |- A -> * : ?T"; +by (DEPTH_SOLVE (ares_tac LP 1)); +uresult(); + +goal LP_thy "A:* P:A->* a:A |- P^a: ?T"; +by (DEPTH_SOLVE (ares_tac LP 1)); +uresult(); + +goal LP_thy "A:* P:A->A->* a:A |- Pi a:A.P^a^a: ?T"; +by (DEPTH_SOLVE (ares_tac LP 1)); +uresult(); + +goal LP_thy "A:* P:A->* Q:A->* |- Pi a:A.P^a -> Q^a: ?T"; +by (DEPTH_SOLVE (ares_tac LP 1)); +uresult(); + +goal LP_thy "A:* P:A->* |- Pi a:A.P^a -> P^a: ?T"; +by (DEPTH_SOLVE (ares_tac LP 1)); +uresult(); + +goal LP_thy "A:* P:A->* |- Lam a:A.Lam x:P^a.x: ?T"; +by (DEPTH_SOLVE (ares_tac LP 1)); +uresult(); + +goal LP_thy "A:* P:A->* Q:* |- (Pi a:A.P^a->Q) -> (Pi a:A.P^a) -> Q : ?T"; +by (DEPTH_SOLVE (ares_tac LP 1)); +uresult(); + +goal LP_thy "A:* P:A->* Q:* a0:A |- \ +\ Lam x:Pi a:A.P^a->Q. Lam y:Pi a:A.P^a. x^a0^(y^a0): ?T"; +by (DEPTH_SOLVE (ares_tac LP 1)); +uresult(); + +(* OMEGA-ORDER TYPES *) + +goal L2_thy "A:* B:* |- Pi C:*.(A->B->C)->C : ?T"; +by (DEPTH_SOLVE (ares_tac L2 1)); +uresult(); + +goal LOmega_thy "|- Lam A:*.Lam B:*.Pi C:*.(A->B->C)->C : ?T"; +by (DEPTH_SOLVE (ares_tac LOmega 1)); +uresult(); + +goal LOmega_thy "|- Lam A:*.Lam B:*.Lam x:A.Lam y:B.x : ?T"; +by (DEPTH_SOLVE (ares_tac LOmega 1)); +uresult(); + +goal LOmega_thy "A:* B:* |- ?p : (A->B) -> ((B->Pi P:*.P)->(A->Pi P:*.P))"; +by (strip_asms_tac LOmega 1); +by (rtac lam_ss 1); +by (DEPTH_SOLVE_1(ares_tac LOmega 1)); +by (DEPTH_SOLVE_1(ares_tac LOmega 2)); +by (rtac lam_ss 1); +by (DEPTH_SOLVE_1(ares_tac LOmega 1)); +by (DEPTH_SOLVE_1(ares_tac LOmega 2)); +by (rtac lam_ss 1); +by (assume_tac 1); +by (DEPTH_SOLVE_1(ares_tac LOmega 2)); +by (etac pi_elim 1); +by (assume_tac 1); +by (etac pi_elim 1); +by (assume_tac 1); +by (assume_tac 1); +uresult(); + +(* Second-order Predicate Logic *) + +goal LP2_thy "A:* P:A->* |- Lam a:A.P^a->(Pi A:*.A) : ?T"; +by (DEPTH_SOLVE (ares_tac LP2 1)); +uresult(); + +goal LP2_thy "A:* P:A->A->* |- \ +\ (Pi a:A.Pi b:A.P^a^b->P^b^a->Pi P:*.P) -> Pi a:A.P^a^a->Pi P:*.P : ?T"; +by (DEPTH_SOLVE (ares_tac LP2 1)); +uresult(); + +(* Antisymmetry implies irreflexivity: *) +goal LP2_thy "A:* P:A->A->* |- \ +\ ?p: (Pi a:A.Pi b:A.P^a^b->P^b^a->Pi P:*.P) -> Pi a:A.P^a^a->Pi P:*.P"; +by (strip_asms_tac LP2 1); +by (rtac lam_ss 1); +by (DEPTH_SOLVE_1(ares_tac LP2 1)); +by (DEPTH_SOLVE_1(ares_tac LP2 2)); +by (rtac lam_ss 1); +by (assume_tac 1); +by (DEPTH_SOLVE_1(ares_tac LP2 2)); +by (rtac lam_ss 1); +by (DEPTH_SOLVE_1(ares_tac LP2 1)); +by (DEPTH_SOLVE_1(ares_tac LP2 2)); +by (REPEAT(EVERY[etac pi_elim 1, assume_tac 1, TRY(assume_tac 1)])); +uresult(); + +(* LPomega *) + +goal LPomega_thy "A:* |- Lam P:A->A->*.Lam a:A.P^a^a : ?T"; +by (DEPTH_SOLVE (ares_tac LPomega 1)); +uresult(); + +goal LPomega_thy "|- Lam A:*.Lam P:A->A->*.Lam a:A.P^a^a : ?T"; +by (DEPTH_SOLVE (ares_tac LPomega 1)); +uresult(); + +(* CONSTRUCTIONS *) + +goal CC_thy "|- Lam A:*.Lam P:A->*.Lam a:A.P^a->Pi P:*.P: ?T"; +by (DEPTH_SOLVE (ares_tac CC 1)); +uresult(); + +goal CC_thy "|- Lam A:*.Lam P:A->*.Pi a:A.P^a: ?T"; +by (DEPTH_SOLVE (ares_tac CC 1)); +uresult(); + +goal CC_thy "A:* P:A->* a:A |- ?p : (Pi a:A.P^a)->P^a"; +by (strip_asms_tac CC 1); +by (rtac lam_ss 1); +by (DEPTH_SOLVE_1(ares_tac CC 1)); +by (DEPTH_SOLVE_1(ares_tac CC 2)); +by (EVERY[etac pi_elim 1, assume_tac 1, assume_tac 1]); +uresult(); + +(* Some random examples *) + +goal LP2_thy "A:* c:A f:A->A |- \ +\ Lam a:A. Pi P:A->*.P^c -> (Pi x:A. P^x->P^(f^x)) -> P^a : ?T"; +by (DEPTH_SOLVE(ares_tac LP2 1)); +uresult(); + +goal CC_thy "Lam A:*.Lam c:A.Lam f:A->A. \ +\ Lam a:A. Pi P:A->*.P^c -> (Pi x:A. P^x->P^(f^x)) -> P^a : ?T"; +by (DEPTH_SOLVE(ares_tac CC 1)); +uresult(); + +(* Symmetry of Leibnitz equality *) +goal LP2_thy "A:* a:A b:A |- ?p: (Pi P:A->*.P^a->P^b) -> (Pi P:A->*.P^b->P^a)"; +by (strip_asms_tac LP2 1); +by (rtac lam_ss 1); +by (DEPTH_SOLVE_1(ares_tac LP2 1)); +by (DEPTH_SOLVE_1(ares_tac LP2 2)); +by (eres_inst_tac [("a","Lam x:A.Pi Q:A->*.Q^x->Q^a")] pi_elim 1); +by (DEPTH_SOLVE_1(ares_tac LP2 1)); +by (rewtac beta); +by (etac imp_elim 1); +by (rtac lam_bs 1); +by (DEPTH_SOLVE_1(ares_tac LP2 1)); +by (DEPTH_SOLVE_1(ares_tac LP2 2)); +by (rtac lam_ss 1); +by (DEPTH_SOLVE_1(ares_tac LP2 1)); +by (DEPTH_SOLVE_1(ares_tac LP2 2)); +by (assume_tac 1); +by (assume_tac 1); +uresult(); + +maketest"END: file for Lambda-Cube examples"; diff -r 000000000000 -r a5a9c433f639 src/FOL/.fol.thy.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/.fol.thy.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,35 @@ +structure FOL = +struct + +local + val parse_ast_translation = [] + val parse_preproc = None + val parse_postproc = None + val parse_translation = [] + val print_translation = [] + val print_preproc = None + val print_postproc = None + val print_ast_translation = [] +in + +(**** begin of user section ****) + +(**** end of user section ****) + +val thy = extend_theory (IFOL.thy) + "FOL" + ([], + [], + [], + [], + [], + None) + [("classical", "(~P ==> P) ==> P")] + +val ax = get_axiom thy + +val classical = ax "classical" + + +end +end diff -r 000000000000 -r a5a9c433f639 src/FOL/.ifol.thy.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/.ifol.thy.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,96 @@ +structure IFOL = +struct + +local + val parse_ast_translation = [] + val parse_preproc = None + val parse_postproc = None + val parse_translation = [] + val print_translation = [] + val print_preproc = None + val print_postproc = None + val print_ast_translation = [] +in + +(**** begin of user section ****) + +(**** end of user section ****) + +val thy = extend_theory (Pure.thy) + "IFOL" + ([("term", ["logic"])], + ["term"], + [(["o"], 0)], + [(["o"], ([], "logic"))], + [(["True", "False"], "o")], + Some (NewSext { + mixfix = + [Mixfix("(_)", "o => prop", "Trueprop", [0], 5), + Infixl("=", "['a,'a] => o", 50), + Mixfix("~ _", "o => o", "Not", [40], 40), + Infixr("&", "[o,o] => o", 35), + Infixr("|", "[o,o] => o", 30), + Infixr("-->", "[o,o] => o", 25), + Infixr("<->", "[o,o] => o", 25), + Binder("ALL ", "('a => o) => o", "All", 0, 10), + Binder("EX ", "('a => o) => o", "Ex", 0, 10), + Binder("EX! ", "('a => o) => o", "Ex1", 0, 10)], + xrules = + [], + parse_ast_translation = parse_ast_translation, + parse_preproc = parse_preproc, + parse_postproc = parse_postproc, + parse_translation = parse_translation, + print_translation = print_translation, + print_preproc = print_preproc, + print_postproc = print_postproc, + print_ast_translation = print_ast_translation})) + [("refl", "a=a"), + ("subst", "[| a=b; P(a) |] ==> P(b)"), + ("conjI", "[| P; Q |] ==> P&Q"), + ("conjunct1", "P&Q ==> P"), + ("conjunct2", "P&Q ==> Q"), + ("disjI1", "P ==> P|Q"), + ("disjI2", "Q ==> P|Q"), + ("disjE", "[| P|Q; P ==> R; Q ==> R |] ==> R"), + ("impI", "(P ==> Q) ==> P-->Q"), + ("mp", "[| P-->Q; P |] ==> Q"), + ("FalseE", "False ==> P"), + ("True_def", "True == False-->False"), + ("not_def", "~P == P-->False"), + ("iff_def", "P<->Q == (P-->Q) & (Q-->P)"), + ("ex1_def", "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)"), + ("allI", "(!!x. P(x)) ==> (ALL x.P(x))"), + ("spec", "(ALL x.P(x)) ==> P(x)"), + ("exI", "P(x) ==> (EX x.P(x))"), + ("exE", "[| EX x.P(x); !!x. P(x) ==> R |] ==> R"), + ("eq_reflection", "(x=y) ==> (x==y)"), + ("iff_reflection", "(P<->Q) ==> (P==Q)")] + +val ax = get_axiom thy + +val refl = ax "refl" +val subst = ax "subst" +val conjI = ax "conjI" +val conjunct1 = ax "conjunct1" +val conjunct2 = ax "conjunct2" +val disjI1 = ax "disjI1" +val disjI2 = ax "disjI2" +val disjE = ax "disjE" +val impI = ax "impI" +val mp = ax "mp" +val FalseE = ax "FalseE" +val True_def = ax "True_def" +val not_def = ax "not_def" +val iff_def = ax "iff_def" +val ex1_def = ax "ex1_def" +val allI = ax "allI" +val spec = ax "spec" +val exI = ax "exI" +val exE = ax "exE" +val eq_reflection = ax "eq_reflection" +val iff_reflection = ax "iff_reflection" + + +end +end diff -r 000000000000 -r a5a9c433f639 src/FOL/FOL.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/FOL.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,94 @@ +(* Title: FOL/fol.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Tactics and lemmas for fol.thy (classical First-Order Logic) +*) + +open FOL; + +signature FOL_LEMMAS = + sig + val disjCI : thm + val excluded_middle : thm + val exCI : thm + val ex_classical : thm + val iffCE : thm + val impCE : thm + val notnotD : thm + val swap : thm + end; + + +structure FOL_Lemmas : FOL_LEMMAS = +struct + +(*** Classical introduction rules for | and EX ***) + +val disjCI = prove_goal FOL.thy + "(~Q ==> P) ==> P|Q" + (fn prems=> + [ (resolve_tac [classical] 1), + (REPEAT (ares_tac (prems@[disjI1,notI]) 1)), + (REPEAT (ares_tac (prems@[disjI2,notE]) 1)) ]); + +(*introduction rule involving only EX*) +val ex_classical = prove_goal FOL.thy + "( ~(EX x. P(x)) ==> P(a)) ==> EX x.P(x)" + (fn prems=> + [ (resolve_tac [classical] 1), + (eresolve_tac (prems RL [exI]) 1) ]); + +(*version of above, simplifying ~EX to ALL~ *) +val exCI = prove_goal FOL.thy + "(ALL x. ~P(x) ==> P(a)) ==> EX x.P(x)" + (fn [prem]=> + [ (resolve_tac [ex_classical] 1), + (resolve_tac [notI RS allI RS prem] 1), + (eresolve_tac [notE] 1), + (eresolve_tac [exI] 1) ]); + +val excluded_middle = prove_goal FOL.thy "~P | P" + (fn _=> [ rtac disjCI 1, assume_tac 1 ]); + + +(*** Special elimination rules *) + + +(*Classical implies (-->) elimination. *) +val impCE = prove_goal FOL.thy + "[| P-->Q; ~P ==> R; Q ==> R |] ==> R" + (fn major::prems=> + [ (resolve_tac [excluded_middle RS disjE] 1), + (DEPTH_SOLVE (ares_tac (prems@[major RS mp]) 1)) ]); + +(*Double negation law*) +val notnotD = prove_goal FOL.thy "~~P ==> P" + (fn [major]=> + [ (resolve_tac [classical] 1), (eresolve_tac [major RS notE] 1) ]); + + +(*** Tactics for implication and contradiction ***) + +(*Classical <-> elimination. Proof substitutes P=Q in + ~P ==> ~Q and P ==> Q *) +val iffCE = prove_goalw FOL.thy [iff_def] + "[| P<->Q; [| P; Q |] ==> R; [| ~P; ~Q |] ==> R |] ==> R" + (fn prems => + [ (resolve_tac [conjE] 1), + (REPEAT (DEPTH_SOLVE_1 + (etac impCE 1 ORELSE mp_tac 1 ORELSE ares_tac prems 1))) ]); + + +(*Should be used as swap since ~P becomes redundant*) +val swap = prove_goal FOL.thy + "~P ==> (~Q ==> P) ==> Q" + (fn major::prems=> + [ (resolve_tac [classical] 1), + (rtac (major RS notE) 1), + (REPEAT (ares_tac prems 1)) ]); + +end; + +open FOL_Lemmas; diff -r 000000000000 -r a5a9c433f639 src/FOL/FOL.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/FOL.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,4 @@ +FOL = IFOL + +rules +classical "(~P ==> P) ==> P" +end diff -r 000000000000 -r a5a9c433f639 src/FOL/IFOL.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/IFOL.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,408 @@ +(* Title: FOL/ifol.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Tactics and lemmas for ifol.thy (intuitionistic first-order logic) +*) + +open IFOL; + +signature IFOL_LEMMAS = + sig + val allE: thm + val all_cong: thm + val all_dupE: thm + val all_impE: thm + val box_equals: thm + val conjE: thm + val conj_cong: thm + val conj_impE: thm + val contrapos: thm + val disj_cong: thm + val disj_impE: thm + val eq_cong: thm + val eq_mp_tac: int -> tactic + val ex1I: thm + val ex1E: thm + val ex1_equalsE: thm + val ex1_cong: thm + val ex_cong: thm + val ex_impE: thm + val iffD1: thm + val iffD2: thm + val iffE: thm + val iffI: thm + val iff_cong: thm + val iff_impE: thm + val iff_refl: thm + val iff_sym: thm + val iff_trans: thm + val impE: thm + val imp_cong: thm + val imp_impE: thm + val mp_tac: int -> tactic + val notE: thm + val notI: thm + val not_cong: thm + val not_impE: thm + val not_sym: thm + val not_to_imp: thm + val pred1_cong: thm + val pred2_cong: thm + val pred3_cong: thm + val pred_congs: thm list + val rev_mp: thm + val simp_equals: thm + val ssubst: thm + val subst_context: thm + val subst_context2: thm + val subst_context3: thm + val sym: thm + val trans: thm + val TrueI: thm + end; + + +structure IFOL_Lemmas : IFOL_LEMMAS = +struct + +val TrueI = prove_goalw IFOL.thy [True_def] "True" + (fn _ => [ (REPEAT (ares_tac [impI] 1)) ]); + +(*** Sequent-style elimination rules for & --> and ALL ***) + +val conjE = prove_goal IFOL.thy + "[| P&Q; [| P; Q |] ==> R |] ==> R" + (fn prems=> + [ (REPEAT (resolve_tac prems 1 + ORELSE (resolve_tac [conjunct1, conjunct2] 1 THEN + resolve_tac prems 1))) ]); + +val impE = prove_goal IFOL.thy + "[| P-->Q; P; Q ==> R |] ==> R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); + +val allE = prove_goal IFOL.thy + "[| ALL x.P(x); P(x) ==> R |] ==> R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[spec]) 1)) ]); + +(*Duplicates the quantifier; for use with eresolve_tac*) +val all_dupE = prove_goal IFOL.thy + "[| ALL x.P(x); [| P(x); ALL x.P(x) |] ==> R \ +\ |] ==> R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[spec]) 1)) ]); + + +(*** Negation rules, which translate between ~P and P-->False ***) + +val notI = prove_goalw IFOL.thy [not_def] "(P ==> False) ==> ~P" + (fn prems=> [ (REPEAT (ares_tac (prems@[impI]) 1)) ]); + +val notE = prove_goalw IFOL.thy [not_def] "[| ~P; P |] ==> R" + (fn prems=> + [ (resolve_tac [mp RS FalseE] 1), + (REPEAT (resolve_tac prems 1)) ]); + +(*This is useful with the special implication rules for each kind of P. *) +val not_to_imp = prove_goal IFOL.thy + "[| ~P; (P-->False) ==> Q |] ==> Q" + (fn prems=> [ (REPEAT (ares_tac (prems@[impI,notE]) 1)) ]); + + +(* For substitution int an assumption P, reduce Q to P-->Q, substitute into + this implication, then apply impI to move P back into the assumptions. + To specify P use something like + eres_inst_tac [ ("P","ALL y. ?S(x,y)") ] rev_mp 1 *) +val rev_mp = prove_goal IFOL.thy "[| P; P --> Q |] ==> Q" + (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); + + +(*Contrapositive of an inference rule*) +val contrapos = prove_goal IFOL.thy "[| ~Q; P==>Q |] ==> ~P" + (fn [major,minor]=> + [ (rtac (major RS notE RS notI) 1), + (etac minor 1) ]); + + +(*** Modus Ponens Tactics ***) + +(*Finds P-->Q and P in the assumptions, replaces implication by Q *) +fun mp_tac i = eresolve_tac [notE,impE] i THEN assume_tac i; + +(*Like mp_tac but instantiates no variables*) +fun eq_mp_tac i = eresolve_tac [notE,impE] i THEN eq_assume_tac i; + + +(*** If-and-only-if ***) + +val iffI = prove_goalw IFOL.thy [iff_def] + "[| P ==> Q; Q ==> P |] ==> P<->Q" + (fn prems=> [ (REPEAT (ares_tac (prems@[conjI, impI]) 1)) ]); + + +(*Observe use of rewrite_rule to unfold "<->" in meta-assumptions (prems) *) +val iffE = prove_goalw IFOL.thy [iff_def] + "[| P <-> Q; [| P-->Q; Q-->P |] ==> R |] ==> R" + (fn prems => [ (resolve_tac [conjE] 1), (REPEAT (ares_tac prems 1)) ]); + +(* Destruct rules for <-> similar to Modus Ponens *) + +val iffD1 = prove_goalw IFOL.thy [iff_def] "[| P <-> Q; P |] ==> Q" + (fn prems => [ (rtac (conjunct1 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); + +val iffD2 = prove_goalw IFOL.thy [iff_def] "[| P <-> Q; Q |] ==> P" + (fn prems => [ (rtac (conjunct2 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); + +val iff_refl = prove_goal IFOL.thy "P <-> P" + (fn _ => [ (REPEAT (ares_tac [iffI] 1)) ]); + +val iff_sym = prove_goal IFOL.thy "Q <-> P ==> P <-> Q" + (fn [major] => + [ (rtac (major RS iffE) 1), + (rtac iffI 1), + (REPEAT (eresolve_tac [asm_rl,mp] 1)) ]); + +val iff_trans = prove_goal IFOL.thy + "!!P Q R. [| P <-> Q; Q<-> R |] ==> P <-> R" + (fn _ => + [ (rtac iffI 1), + (REPEAT (eresolve_tac [asm_rl,iffE] 1 ORELSE mp_tac 1)) ]); + + +(*** Unique existence. NOTE THAT the following 2 quantifications + EX!x such that [EX!y such that P(x,y)] (sequential) + EX!x,y such that P(x,y) (simultaneous) + do NOT mean the same thing. The parser treats EX!x y.P(x,y) as sequential. +***) + +val ex1I = prove_goalw IFOL.thy [ex1_def] + "[| P(a); !!x. P(x) ==> x=a |] ==> EX! x. P(x)" + (fn prems => [ (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ]); + +val ex1E = prove_goalw IFOL.thy [ex1_def] + "[| EX! x.P(x); !!x. [| P(x); ALL y. P(y) --> y=x |] ==> R |] ==> R" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (eresolve_tac [exE,conjE] 1 ORELSE ares_tac prems 1)) ]); + + +(*** <-> congruence rules for simplification ***) + +(*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*) +fun iff_tac prems i = + resolve_tac (prems RL [iffE]) i THEN + REPEAT1 (eresolve_tac [asm_rl,mp] i); + +val conj_cong = prove_goal IFOL.thy + "[| P <-> P'; P' ==> Q <-> Q' |] ==> (P&Q) <-> (P'&Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [iffI,conjI] 1 + ORELSE eresolve_tac [iffE,conjE,mp] 1 + ORELSE iff_tac prems 1)) ]); + +val disj_cong = prove_goal IFOL.thy + "[| P <-> P'; Q <-> Q' |] ==> (P|Q) <-> (P'|Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (eresolve_tac [iffE,disjE,disjI1,disjI2] 1 + ORELSE ares_tac [iffI] 1 + ORELSE mp_tac 1)) ]); + +val imp_cong = prove_goal IFOL.thy + "[| P <-> P'; P' ==> Q <-> Q' |] ==> (P-->Q) <-> (P'-->Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [iffI,impI] 1 + ORELSE eresolve_tac [iffE] 1 + ORELSE mp_tac 1 ORELSE iff_tac prems 1)) ]); + +val iff_cong = prove_goal IFOL.thy + "[| P <-> P'; Q <-> Q' |] ==> (P<->Q) <-> (P'<->Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (eresolve_tac [iffE] 1 + ORELSE ares_tac [iffI] 1 + ORELSE mp_tac 1)) ]); + +val not_cong = prove_goal IFOL.thy + "P <-> P' ==> ~P <-> ~P'" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [iffI,notI] 1 + ORELSE mp_tac 1 + ORELSE eresolve_tac [iffE,notE] 1)) ]); + +val all_cong = prove_goal IFOL.thy + "(!!x.P(x) <-> Q(x)) ==> (ALL x.P(x)) <-> (ALL x.Q(x))" + (fn prems => + [ (REPEAT (ares_tac [iffI,allI] 1 + ORELSE mp_tac 1 + ORELSE eresolve_tac [allE] 1 ORELSE iff_tac prems 1)) ]); + +val ex_cong = prove_goal IFOL.thy + "(!!x.P(x) <-> Q(x)) ==> (EX x.P(x)) <-> (EX x.Q(x))" + (fn prems => + [ (REPEAT (eresolve_tac [exE] 1 ORELSE ares_tac [iffI,exI] 1 + ORELSE mp_tac 1 + ORELSE iff_tac prems 1)) ]); + +val ex1_cong = prove_goal IFOL.thy + "(!!x.P(x) <-> Q(x)) ==> (EX! x.P(x)) <-> (EX! x.Q(x))" + (fn prems => + [ (REPEAT (eresolve_tac [ex1E, spec RS mp] 1 ORELSE ares_tac [iffI,ex1I] 1 + ORELSE mp_tac 1 + ORELSE iff_tac prems 1)) ]); + +(*** Equality rules ***) + +val sym = prove_goal IFOL.thy "a=b ==> b=a" + (fn [major] => [ (rtac (major RS subst) 1), (rtac refl 1) ]); + +val trans = prove_goal IFOL.thy "[| a=b; b=c |] ==> a=c" + (fn [prem1,prem2] => [ (rtac (prem2 RS subst) 1), (rtac prem1 1) ]); + +(** ~ b=a ==> ~ a=b **) +val [not_sym] = compose(sym,2,contrapos); + +(*calling "standard" reduces maxidx to 0*) +val ssubst = standard (sym RS subst); + +(*A special case of ex1E that would otherwise need quantifier expansion*) +val ex1_equalsE = prove_goal IFOL.thy + "[| EX! x.P(x); P(a); P(b) |] ==> a=b" + (fn prems => + [ (cut_facts_tac prems 1), + (etac ex1E 1), + (rtac trans 1), + (rtac sym 2), + (REPEAT (eresolve_tac [asm_rl, spec RS mp] 1)) ]); + +(** Polymorphic congruence rules **) + +val subst_context = prove_goal IFOL.thy + "[| a=b |] ==> t(a)=t(b)" + (fn prems=> + [ (resolve_tac (prems RL [ssubst]) 1), + (resolve_tac [refl] 1) ]); + +val subst_context2 = prove_goal IFOL.thy + "[| a=b; c=d |] ==> t(a,c)=t(b,d)" + (fn prems=> + [ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); + +val subst_context3 = prove_goal IFOL.thy + "[| a=b; c=d; e=f |] ==> t(a,c,e)=t(b,d,f)" + (fn prems=> + [ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); + +(*Useful with eresolve_tac for proving equalties from known equalities. + a = b + | | + c = d *) +val box_equals = prove_goal IFOL.thy + "[| a=b; a=c; b=d |] ==> c=d" + (fn prems=> + [ (resolve_tac [trans] 1), + (resolve_tac [trans] 1), + (resolve_tac [sym] 1), + (REPEAT (resolve_tac prems 1)) ]); + +(*Dual of box_equals: for proving equalities backwards*) +val simp_equals = prove_goal IFOL.thy + "[| a=c; b=d; c=d |] ==> a=b" + (fn prems=> + [ (resolve_tac [trans] 1), + (resolve_tac [trans] 1), + (REPEAT (resolve_tac (prems @ (prems RL [sym])) 1)) ]); + +(** Congruence rules for predicate letters **) + +val pred1_cong = prove_goal IFOL.thy + "a=a' ==> P(a) <-> P(a')" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); + +val pred2_cong = prove_goal IFOL.thy + "[| a=a'; b=b' |] ==> P(a,b) <-> P(a',b')" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); + +val pred3_cong = prove_goal IFOL.thy + "[| a=a'; b=b'; c=c' |] ==> P(a,b,c) <-> P(a',b',c')" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); + +(*special cases for free variables P, Q, R, S -- up to 3 arguments*) + +val pred_congs = + flat (map (fn c => + map (fn th => read_instantiate [("P",c)] th) + [pred1_cong,pred2_cong,pred3_cong]) + (explode"PQRS")); + +(*special case for the equality predicate!*) +val eq_cong = read_instantiate [("P","op =")] pred2_cong; + + +(*** Simplifications of assumed implications. + Roy Dyckhoff has proved that conj_impE, disj_impE, and imp_impE + used with mp_tac (restricted to atomic formulae) is COMPLETE for + intuitionistic propositional logic. See + R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic + (preprint, University of St Andrews, 1991) ***) + +val conj_impE = prove_goal IFOL.thy + "[| (P&Q)-->S; P-->(Q-->S) ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (ares_tac ([conjI, impI, major RS mp]@prems) 1)) ]); + +val disj_impE = prove_goal IFOL.thy + "[| (P|Q)-->S; [| P-->S; Q-->S |] ==> R |] ==> R" + (fn major::prems=> + [ (DEPTH_SOLVE (ares_tac ([disjI1, disjI2, impI, major RS mp]@prems) 1)) ]); + +(*Simplifies the implication. Classical version is stronger. + Still UNSAFE since Q must be provable -- backtracking needed. *) +val imp_impE = prove_goal IFOL.thy + "[| (P-->Q)-->S; [| P; Q-->S |] ==> Q; S ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (ares_tac ([impI, major RS mp]@prems) 1)) ]); + +(*Simplifies the implication. Classical version is stronger. + Still UNSAFE since ~P must be provable -- backtracking needed. *) +val not_impE = prove_goal IFOL.thy + "[| ~P --> S; P ==> False; S ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (ares_tac ([notI, impI, major RS mp]@prems) 1)) ]); + +(*Simplifies the implication. UNSAFE. *) +val iff_impE = prove_goal IFOL.thy + "[| (P<->Q)-->S; [| P; Q-->S |] ==> Q; [| Q; P-->S |] ==> P; \ +\ S ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (ares_tac ([iffI, impI, major RS mp]@prems) 1)) ]); + +(*What if (ALL x.~~P(x)) --> ~~(ALL x.P(x)) is an assumption? UNSAFE*) +val all_impE = prove_goal IFOL.thy + "[| (ALL x.P(x))-->S; !!x.P(x); S ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (ares_tac ([allI, impI, major RS mp]@prems) 1)) ]); + +(*Unsafe: (EX x.P(x))-->S is equivalent to ALL x.P(x)-->S. *) +val ex_impE = prove_goal IFOL.thy + "[| (EX x.P(x))-->S; P(x)-->S ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (ares_tac ([exI, impI, major RS mp]@prems) 1)) ]); + +end; + +open IFOL_Lemmas; + diff -r 000000000000 -r a5a9c433f639 src/FOL/IFOL.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/IFOL.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,70 @@ +IFOL = Pure + + +classes term < logic + +default term + +types o 0 + +arities o :: logic + +consts + Trueprop :: "o => prop" ("(_)" [0] 5) + True,False :: "o" + (*Connectives*) + "=" :: "['a,'a] => o" (infixl 50) + Not :: "o => o" ("~ _" [40] 40) + "&" :: "[o,o] => o" (infixr 35) + "|" :: "[o,o] => o" (infixr 30) + "-->" :: "[o,o] => o" (infixr 25) + "<->" :: "[o,o] => o" (infixr 25) + (*Quantifiers*) + All :: "('a => o) => o" (binder "ALL " 10) + Ex :: "('a => o) => o" (binder "EX " 10) + Ex1 :: "('a => o) => o" (binder "EX! " 10) + +rules + + (*Equality*) + +refl "a=a" +subst "[| a=b; P(a) |] ==> P(b)" + + (*Propositional logic*) + +conjI "[| P; Q |] ==> P&Q" +conjunct1 "P&Q ==> P" +conjunct2 "P&Q ==> Q" + +disjI1 "P ==> P|Q" +disjI2 "Q ==> P|Q" +disjE "[| P|Q; P ==> R; Q ==> R |] ==> R" + +impI "(P ==> Q) ==> P-->Q" +mp "[| P-->Q; P |] ==> Q" + +FalseE "False ==> P" + + (*Definitions*) + +True_def "True == False-->False" +not_def "~P == P-->False" +iff_def "P<->Q == (P-->Q) & (Q-->P)" + + (*Unique existence*) +ex1_def "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)" + + (*Quantifiers*) + +allI "(!!x. P(x)) ==> (ALL x.P(x))" +spec "(ALL x.P(x)) ==> P(x)" + +exI "P(x) ==> (EX x.P(x))" +exE "[| EX x.P(x); !!x. P(x) ==> R |] ==> R" + + (* Reflection *) + +eq_reflection "(x=y) ==> (x==y)" +iff_reflection "(P<->Q) ==> (P==Q)" + +end diff -r 000000000000 -r a5a9c433f639 src/FOL/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/Makefile Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,43 @@ +######################################################################### +# # +# Makefile for Isabelle (FOL) # +# # +######################################################################### + +#To make the system, cd to this directory and type +# make -f Makefile +#To make the system and test it on standard examples, type +# make -f Makefile test + +#Environment variable ISABELLECOMP specifies the compiler. +#Environment variable ISABELLEBIN specifies the destination directory. +#For Poly/ML, ISABELLEBIN must begin with a / + +#Makes pure Isabelle (Pure) if this file is ABSENT -- but not +#if it is out of date, since this Makefile does not know its dependencies! + +BIN = $(ISABELLEBIN) +COMP = $(ISABELLECOMP) +FILES = ROOT.ML ifol.thy ifol.ML fol.thy fol.ML int-prover.ML simpdata.ML \ + ../Provers/classical.ML ../Provers/simplifier.ML ../Provers/ind.ML + +$(BIN)/FOL: $(BIN)/Pure $(FILES) + case "$(COMP)" in \ + poly*) echo 'make_database"$(BIN)/FOL"; quit();' \ + | $(COMP) $(BIN)/Pure;\ + echo 'open PolyML; use"ROOT";' | $(COMP) $(BIN)/FOL;;\ + sml*) echo 'use"ROOT.ML"; xML"$(BIN)/FOL" banner;' | $(BIN)/Pure;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +$(BIN)/Pure: + cd ../Pure; $(MAKE) + +test: ex/ROOT.ML $(BIN)/FOL + case "$(COMP)" in \ + poly*) echo 'use"ex/ROOT.ML"; quit();' | $(COMP) $(BIN)/FOL ;;\ + sml*) echo 'use"ex/ROOT.ML";' | $(BIN)/FOL;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +.PRECIOUS: $(BIN)/Pure $(BIN)/FOL diff -r 000000000000 -r a5a9c433f639 src/FOL/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/README Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,19 @@ + FOL: First-Order Logic with Natural Deduction + +This directory contains the Standard ML sources of the Isabelle system for +First-Order Logic (constructive and classical versions). For a classical +sequent calculus, see LK. Important files include + +ROOT.ML -- loads all source files. Enter an ML image containing Pure +Isabelle and type: use "ROOT.ML"; + +Makefile -- compiles the files under Poly/ML or SML of New Jersey + +ex -- subdirectory containing examples. To execute them, enter an ML image +containing FOL and type: use "ex/ROOT.ML"; + +Useful references on First-Order Logic: + + Antony Galton, Logic for Information Technology (Wiley, 1990) + + Michael Dummett, Elements of Intuitionism (Oxford, 1977) \ No newline at end of file diff -r 000000000000 -r a5a9c433f639 src/FOL/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,79 @@ +(* Title: FOL/ROOT + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Adds First-Order Logic to a database containing pure Isabelle. +Should be executed in the subdirectory FOL. +*) + +val banner = "First-Order Logic with Natural Deduction"; + +writeln banner; + +print_depth 1; +use_thy "ifol"; +use_thy "fol"; + +use "../Provers/hypsubst.ML"; +use "../Provers/classical.ML"; +use "../Provers/simplifier.ML"; +use "../Provers/splitter.ML"; +use "../Provers/ind.ML"; + +(*** Applying HypsubstFun to generate hyp_subst_tac ***) + +structure Hypsubst_Data = + struct + (*Take apart an equality judgement; otherwise raise Match!*) + fun dest_eq (Const("Trueprop",_) $ (Const("op =",_) $ t $ u)) = (t,u); + + val imp_intr = impI + + (*etac rev_cut_eq moves an equality to be the last premise. *) + val rev_cut_eq = prove_goal IFOL.thy "[| a=b; a=b ==> R |] ==> R" + (fn prems => [ REPEAT(resolve_tac prems 1) ]); + + val rev_mp = rev_mp + val subst = subst + val sym = sym + end; + +structure Hypsubst = HypsubstFun(Hypsubst_Data); +open Hypsubst; + +use "int-prover.ML"; + +(*** Applying ClassicalFun to create a classical prover ***) +structure Classical_Data = + struct + val sizef = size_of_thm + val mp = mp + val not_elim = notE + val swap = swap + val hyp_subst_tacs=[hyp_subst_tac] + end; + +structure Cla = ClassicalFun(Classical_Data); +open Cla; + +(*Propositional rules + -- iffCE might seem better, but in the examples in ex/cla + run about 7% slower than with iffE*) +val prop_cs = empty_cs addSIs [refl,TrueI,conjI,disjCI,impI,notI,iffI] + addSEs [conjE,disjE,impCE,FalseE,iffE]; + +(*Quantifier rules*) +val FOL_cs = prop_cs addSIs [allI] addIs [exI,ex1I] + addSEs [exE,ex1E] addEs [allE]; + +val FOL_dup_cs = prop_cs addSIs [allI] addIs [exCI,ex1I] + addSEs [exE,ex1E] addEs [all_dupE]; + +use "simpdata.ML"; + +use "../Pure/install_pp.ML"; +print_depth 8; + +val FOL_build_completed = (); (*indicate successful build*) + diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/.if.thy.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/.if.thy.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,35 @@ +structure If = +struct + +local + val parse_ast_translation = [] + val parse_preproc = None + val parse_postproc = None + val parse_translation = [] + val print_translation = [] + val print_preproc = None + val print_postproc = None + val print_ast_translation = [] +in + +(**** begin of user section ****) + +(**** end of user section ****) + +val thy = extend_theory (FOL.thy) + "If" + ([], + [], + [], + [], + [(["if"], "[o,o,o]=>o")], + None) + [("if_def", "if(P,Q,R) == P&Q | ~P&R")] + +val ax = get_axiom thy + +val if_def = ax "if_def" + + +end +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/.list.thy.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/.list.thy.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,81 @@ +structure List = +struct + +local + val parse_ast_translation = [] + val parse_preproc = None + val parse_postproc = None + val parse_translation = [] + val print_translation = [] + val print_preproc = None + val print_postproc = None + val print_ast_translation = [] +in + +(**** begin of user section ****) + +(**** end of user section ****) + +val thy = extend_theory (Nat2.thy) + "List" + ([], + [], + [(["list"], 1)], + [(["list"], ([["term"]], "term"))], + [(["hd"], "'a list => 'a"), + (["tl"], "'a list => 'a list"), + (["forall"], "['a list, 'a => o] => o"), + (["len"], "'a list => nat"), + (["at"], "['a list, nat] => 'a")], + Some (NewSext { + mixfix = + [Delimfix("[]", "'a list", "[]"), + Infixr(".", "['a, 'a list] => 'a list", 80), + Infixr("++", "['a list, 'a list] => 'a list", 70)], + xrules = + [], + parse_ast_translation = parse_ast_translation, + parse_preproc = parse_preproc, + parse_postproc = parse_postproc, + parse_translation = parse_translation, + print_translation = print_translation, + print_preproc = print_preproc, + print_postproc = print_postproc, + print_ast_translation = print_ast_translation})) + [("list_ind", "[| P([]); ALL x l. P(l)-->P(x.l) |] ==> All(P)"), + ("forall_cong", "[| l = l'; !!x. P(x)<->P'(x) |] ==> forall(l,P) <-> forall(l',P')"), + ("list_distinct1", "~[] = x.l"), + ("list_distinct2", "~x.l = []"), + ("list_free", "x.l = x'.l' <-> x=x' & l=l'"), + ("app_nil", "[]++l = l"), + ("app_cons", "(x.l)++l' = x.(l++l')"), + ("tl_eq", "tl(m.q) = q"), + ("hd_eq", "hd(m.q) = m"), + ("forall_nil", "forall([],P)"), + ("forall_cons", "forall(x.l,P) <-> P(x) & forall(l,P)"), + ("len_nil", "len([]) = 0"), + ("len_cons", "len(m.q) = succ(len(q))"), + ("at_0", "at(m.q,0) = m"), + ("at_succ", "at(m.q,succ(n)) = at(q,n)")] + +val ax = get_axiom thy + +val list_ind = ax "list_ind" +val forall_cong = ax "forall_cong" +val list_distinct1 = ax "list_distinct1" +val list_distinct2 = ax "list_distinct2" +val list_free = ax "list_free" +val app_nil = ax "app_nil" +val app_cons = ax "app_cons" +val tl_eq = ax "tl_eq" +val hd_eq = ax "hd_eq" +val forall_nil = ax "forall_nil" +val forall_cons = ax "forall_cons" +val len_nil = ax "len_nil" +val len_cons = ax "len_cons" +val at_0 = ax "at_0" +val at_succ = ax "at_succ" + + +end +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/.nat.thy.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/.nat.thy.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,59 @@ +structure Nat = +struct + +local + val parse_ast_translation = [] + val parse_preproc = None + val parse_postproc = None + val parse_translation = [] + val print_translation = [] + val print_preproc = None + val print_postproc = None + val print_ast_translation = [] +in + +(**** begin of user section ****) + +(**** end of user section ****) + +val thy = extend_theory (FOL.thy) + "Nat" + ([], + [], + [(["nat"], 0)], + [(["nat"], ([], "term"))], + [(["Suc"], "nat=>nat"), + (["rec"], "[nat, 'a, [nat,'a]=>'a] => 'a")], + Some (NewSext { + mixfix = + [Delimfix("0", "nat", "0"), + Infixl("+", "[nat, nat] => nat", 60)], + xrules = + [], + parse_ast_translation = parse_ast_translation, + parse_preproc = parse_preproc, + parse_postproc = parse_postproc, + parse_translation = parse_translation, + print_translation = print_translation, + print_preproc = print_preproc, + print_postproc = print_postproc, + print_ast_translation = print_ast_translation})) + [("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))"), + ("add_def", "m+n == rec(m, n, %x y. Suc(y))")] + +val ax = get_axiom thy + +val induct = ax "induct" +val Suc_inject = ax "Suc_inject" +val Suc_neq_0 = ax "Suc_neq_0" +val rec_0 = ax "rec_0" +val rec_Suc = ax "rec_Suc" +val add_def = ax "add_def" + + +end +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/.nat2.thy.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/.nat2.thy.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,76 @@ +structure Nat2 = +struct + +local + val parse_ast_translation = [] + val parse_preproc = None + val parse_postproc = None + val parse_translation = [] + val print_translation = [] + val print_preproc = None + val print_postproc = None + val print_ast_translation = [] +in + +(**** begin of user section ****) + +(**** end of user section ****) + +val thy = extend_theory (FOL.thy) + "Nat2" + ([], + [], + [(["nat"], 0)], + [(["nat"], ([], "term"))], + [(["succ", "pred"], "nat => nat")], + Some (NewSext { + mixfix = + [Delimfix("0", "nat", "0"), + Infixr("+", "[nat,nat] => nat", 90), + Infixr("<", "[nat,nat] => o", 70), + Infixr("<=", "[nat,nat] => o", 70)], + xrules = + [], + parse_ast_translation = parse_ast_translation, + parse_preproc = parse_preproc, + parse_postproc = parse_postproc, + parse_translation = parse_translation, + print_translation = print_translation, + print_preproc = print_preproc, + print_postproc = print_postproc, + print_ast_translation = print_ast_translation})) + [("pred_0", "pred(0) = 0"), + ("pred_succ", "pred(succ(m)) = m"), + ("plus_0", "0+n = n"), + ("plus_succ", "succ(m)+n = succ(m+n)"), + ("nat_distinct1", "~ 0=succ(n)"), + ("nat_distinct2", "~ succ(m)=0"), + ("succ_inject", "succ(m)=succ(n) <-> m=n"), + ("leq_0", "0 <= n"), + ("leq_succ_succ", "succ(m)<=succ(n) <-> m<=n"), + ("leq_succ_0", "~ succ(m) <= 0"), + ("lt_0_succ", "0 < succ(n)"), + ("lt_succ_succ", "succ(m) mP(succ(n)) |] ==> All(P)")] + +val ax = get_axiom thy + +val pred_0 = ax "pred_0" +val pred_succ = ax "pred_succ" +val plus_0 = ax "plus_0" +val plus_succ = ax "plus_succ" +val nat_distinct1 = ax "nat_distinct1" +val nat_distinct2 = ax "nat_distinct2" +val succ_inject = ax "succ_inject" +val leq_0 = ax "leq_0" +val leq_succ_succ = ax "leq_succ_succ" +val leq_succ_0 = ax "leq_succ_0" +val lt_0_succ = ax "lt_0_succ" +val lt_succ_succ = ax "lt_succ_succ" +val lt_0 = ax "lt_0" +val nat_ind = ax "nat_ind" + + +end +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/.prolog.thy.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/.prolog.thy.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,55 @@ +structure Prolog = +struct + +local + val parse_ast_translation = [] + val parse_preproc = None + val parse_postproc = None + val parse_translation = [] + val print_translation = [] + val print_preproc = None + val print_postproc = None + val print_ast_translation = [] +in + +(**** begin of user section ****) + +(**** end of user section ****) + +val thy = extend_theory (FOL.thy) + "Prolog" + ([], + [], + [(["list"], 1)], + [(["list"], ([["term"]], "term"))], + [(["Nil"], "'a list"), + (["app"], "['a list, 'a list, 'a list] => o"), + (["rev"], "['a list, 'a list] => o")], + Some (NewSext { + mixfix = + [Infixr(":", "['a, 'a list]=> 'a list", 60)], + xrules = + [], + parse_ast_translation = parse_ast_translation, + parse_preproc = parse_preproc, + parse_postproc = parse_postproc, + parse_translation = parse_translation, + print_translation = print_translation, + print_preproc = print_preproc, + print_postproc = print_postproc, + print_ast_translation = print_ast_translation})) + [("appNil", "app(Nil,ys,ys)"), + ("appCons", "app(xs,ys,zs) ==> app(x:xs, ys, x:zs)"), + ("revNil", "rev(Nil,Nil)"), + ("revCons", "[| rev(xs,ys); app(ys, x:Nil, zs) |] ==> rev(x:xs, zs)")] + +val ax = get_axiom thy + +val appNil = ax "appNil" +val appCons = ax "appCons" +val revNil = ax "revNil" +val revCons = ax "revCons" + + +end +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/If.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/If.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,64 @@ +(* Title: FOL/ex/if + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For ex/if.thy. First-Order Logic: the 'if' example +*) + +open If; +open Cla; (*in case structure Int is open!*) + +val prems = goalw If.thy [if_def] + "[| P ==> Q; ~P ==> R |] ==> if(P,Q,R)"; +by (fast_tac (FOL_cs addIs prems) 1); +val ifI = result(); + +val major::prems = goalw If.thy [if_def] + "[| if(P,Q,R); [| P; Q |] ==> S; [| ~P; R |] ==> S |] ==> S"; +by (cut_facts_tac [major] 1); +by (fast_tac (FOL_cs addIs prems) 1); +val ifE = result(); + + +goal If.thy + "if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))"; +by (resolve_tac [iffI] 1); +by (eresolve_tac [ifE] 1); +by (eresolve_tac [ifE] 1); +by (resolve_tac [ifI] 1); +by (resolve_tac [ifI] 1); + +choplev 0; +val if_cs = FOL_cs addSIs [ifI] addSEs[ifE]; +by (fast_tac if_cs 1); +val if_commute = result(); + + +goal If.thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))"; +by (fast_tac if_cs 1); +val nested_ifs = result(); + +choplev 0; +by (rewrite_goals_tac [if_def]); +by (fast_tac FOL_cs 1); +result(); + + +(*An invalid formula. High-level rules permit a simpler diagnosis*) +goal If.thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))"; +by (fast_tac if_cs 1) handle ERROR => writeln"Failed, as expected"; +(*Check that subgoals remain: proof failed.*) +getgoal 1; +by (REPEAT (step_tac if_cs 1)); + +choplev 0; +by (rewrite_goals_tac [if_def]); +by (fast_tac FOL_cs 1) handle ERROR => writeln"Failed, as expected"; +(*Check that subgoals remain: proof failed.*) +getgoal 1; +by (REPEAT (step_tac FOL_cs 1)); + + + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/If.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/If.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,5 @@ +If = FOL + +consts if :: "[o,o,o]=>o" +rules + if_def "if(P,Q,R) == P&Q | ~P&R" +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/List.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/List.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,67 @@ +(* Title: FOL/ex/list + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +For ex/list.thy. Examples of simplification and induction on lists +*) + +open List; + +val prems = goal List.thy "[| P([]); !!x l. P(x.l) |] ==> All(P)"; +by (rtac list_ind 1); +by (REPEAT (resolve_tac (prems@[allI,impI]) 1)); +val list_exh = result(); + +val list_rew_thms = [list_distinct1,list_distinct2,app_nil,app_cons, + hd_eq,tl_eq,forall_nil,forall_cons,list_free, + len_nil,len_cons,at_0,at_succ]; + +val list_ss = nat_ss addsimps list_rew_thms; + +goal List.thy "~l=[] --> hd(l).tl(l) = l"; +by(IND_TAC list_exh (simp_tac list_ss) "l" 1); +result(); + +goal List.thy "(l1++l2)++l3 = l1++(l2++l3)"; +by(IND_TAC list_ind (simp_tac list_ss) "l1" 1); +val append_assoc = result(); + +goal List.thy "l++[] = l"; +by(IND_TAC list_ind (simp_tac list_ss) "l" 1); +val app_nil_right = result(); + +goal List.thy "l1++l2=[] <-> l1=[] & l2=[]"; +by(IND_TAC list_exh (simp_tac list_ss) "l1" 1); +val app_eq_nil_iff = result(); + +goal List.thy "forall(l++l',P) <-> forall(l,P) & forall(l',P)"; +by(IND_TAC list_ind (simp_tac list_ss) "l" 1); +val forall_app = result(); + +goal List.thy "forall(l,%x.P(x)&Q(x)) <-> forall(l,P) & forall(l,Q)"; +by(IND_TAC list_ind (simp_tac list_ss) "l" 1); +by(fast_tac FOL_cs 1); +val forall_conj = result(); + +goal List.thy "~l=[] --> forall(l,P) <-> P(hd(l)) & forall(tl(l),P)"; +by(IND_TAC list_ind (simp_tac list_ss) "l" 1); +val forall_ne = result(); + +(*** Lists with natural numbers ***) + +goal List.thy "len(l1++l2) = len(l1)+len(l2)"; +by (IND_TAC list_ind (simp_tac list_ss) "l1" 1); +val len_app = result(); + +goal List.thy "i at(l1++l2,i) = at(l1,i)"; +by (IND_TAC list_ind (simp_tac list_ss) "l1" 1); +by (REPEAT (rtac allI 1)); +by (rtac impI 1); +by (ALL_IND_TAC nat_exh (asm_simp_tac list_ss) 1); +val at_app1 = result(); + +goal List.thy "at(l1++(x.l2), len(l1)) = x"; +by (IND_TAC list_ind (simp_tac list_ss) "l1" 1); +val at_app_hd2 = result(); + diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/List.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/List.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,49 @@ +(* Title: FOL/ex/list + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +Examples of simplification and induction on lists +*) + +List = Nat2 + + +types list 1 +arities list :: (term)term + +consts + hd :: "'a list => 'a" + tl :: "'a list => 'a list" + forall :: "['a list, 'a => o] => o" + len :: "'a list => nat" + at :: "['a list, nat] => 'a" + "[]" :: "'a list" ("[]") + "." :: "['a, 'a list] => 'a list" (infixr 80) + "++" :: "['a list, 'a list] => 'a list" (infixr 70) + +rules + list_ind "[| P([]); ALL x l. P(l)-->P(x.l) |] ==> All(P)" + + forall_cong + "[| l = l'; !!x. P(x)<->P'(x) |] ==> forall(l,P) <-> forall(l',P')" + + list_distinct1 "~[] = x.l" + list_distinct2 "~x.l = []" + + list_free "x.l = x'.l' <-> x=x' & l=l'" + + app_nil "[]++l = l" + app_cons "(x.l)++l' = x.(l++l')" + tl_eq "tl(m.q) = q" + hd_eq "hd(m.q) = m" + + forall_nil "forall([],P)" + forall_cons "forall(x.l,P) <-> P(x) & forall(l,P)" + + len_nil "len([]) = 0" + len_cons "len(m.q) = succ(len(q))" + + at_0 "at(m.q,0) = m" + at_succ "at(m.q,succ(n)) = at(q,n)" + +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/Nat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/Nat.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,69 @@ +(* Title: FOL/ex/nat.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Examples for the manual "Introduction to Isabelle" + +Proofs about the natural numbers + +INCOMPATIBLE with nat2.ML, Nipkow's examples + +To generate similar output to manual, execute these commands: + Pretty.setmargin 72; print_depth 0; +*) + +open Nat; + +goal Nat.thy "~ (Suc(k) = k)"; +by (res_inst_tac [("n","k")] induct 1); +by (resolve_tac [notI] 1); +by (eresolve_tac [Suc_neq_0] 1); +by (resolve_tac [notI] 1); +by (eresolve_tac [notE] 1); +by (eresolve_tac [Suc_inject] 1); +val Suc_n_not_n = result(); + + +goal Nat.thy "(k+m)+n = k+(m+n)"; +prths ([induct] RL [topthm()]); (*prints all 14 next states!*) +by (resolve_tac [induct] 1); +back(); +back(); +back(); +back(); +back(); +back(); + +goalw Nat.thy [add_def] "0+n = n"; +by (resolve_tac [rec_0] 1); +val add_0 = result(); + +goalw Nat.thy [add_def] "Suc(m)+n = Suc(m+n)"; +by (resolve_tac [rec_Suc] 1); +val add_Suc = result(); + +val add_ss = FOL_ss addsimps [add_0, add_Suc]; + +goal Nat.thy "(k+m)+n = k+(m+n)"; +by (res_inst_tac [("n","k")] induct 1); +by (simp_tac add_ss 1); +by (asm_simp_tac add_ss 1); +val add_assoc = result(); + +goal Nat.thy "m+0 = m"; +by (res_inst_tac [("n","m")] induct 1); +by (simp_tac add_ss 1); +by (asm_simp_tac add_ss 1); +val add_0_right = result(); + +goal Nat.thy "m+Suc(n) = Suc(m+n)"; +by (res_inst_tac [("n","m")] induct 1); +by (ALLGOALS (asm_simp_tac add_ss)); +val add_Suc_right = result(); + +val [prem] = goal Nat.thy "(!!n. f(Suc(n)) = Suc(f(n))) ==> f(i+j) = i+f(j)"; +by (res_inst_tac [("n","i")] induct 1); +by (simp_tac add_ss 1); +by (asm_simp_tac (add_ss addsimps [prem]) 1); +result(); diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/Nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/Nat.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,26 @@ +(* Title: FOL/ex/nat.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Examples for the manual "Introduction to Isabelle" + +Theory of the natural numbers: Peano's axioms, primitive recursion + +INCOMPATIBLE with nat2.thy, Nipkow's example +*) + +Nat = FOL + +types nat 0 +arities nat :: term +consts "0" :: "nat" ("0") + Suc :: "nat=>nat" + rec :: "[nat, 'a, [nat,'a]=>'a] => 'a" + "+" :: "[nat, nat] => nat" (infixl 60) +rules 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))" + add_def "m+n == rec(m, n, %x y. Suc(y))" +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/Nat2.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/Nat2.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,163 @@ +(* Title: FOL/ex/nat2.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +For ex/nat.thy. +Examples of simplification and induction on the natural numbers +*) + +open Nat2; + +val nat_rews = [pred_0, pred_succ, plus_0, plus_succ, + nat_distinct1, nat_distinct2, succ_inject, + leq_0,leq_succ_succ,leq_succ_0, + lt_0_succ,lt_succ_succ,lt_0]; + +val nat_ss = FOL_ss addsimps nat_rews; + +val prems = goal Nat2.thy + "[| P(0); !!x. P(succ(x)) |] ==> All(P)"; +by (rtac nat_ind 1); +by (REPEAT (resolve_tac (prems@[allI,impI]) 1)); +val nat_exh = result(); + +goal Nat2.thy "~ n=succ(n)"; +by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1); +result(); + +goal Nat2.thy "~ succ(n)=n"; +by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1); +result(); + +goal Nat2.thy "~ succ(succ(n))=n"; +by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1); +result(); + +goal Nat2.thy "~ n=succ(succ(n))"; +by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1); +result(); + +goal Nat2.thy "m+0 = m"; +by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1); +val plus_0_right = result(); + +goal Nat2.thy "m+succ(n) = succ(m+n)"; +by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1); +val plus_succ_right = result(); + +goal Nat2.thy "~n=0 --> m+pred(n) = pred(m+n)"; +by (IND_TAC nat_ind (simp_tac (nat_ss addsimps [plus_succ_right])) "n" 1); +result(); + +goal Nat2.thy "~n=0 --> succ(pred(n))=n"; +by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1); +result(); + +goal Nat2.thy "m+n=0 <-> m=0 & n=0"; +by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1); +result(); + +goal Nat2.thy "m <= n --> m <= succ(n)"; +by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1); +by (rtac (impI RS allI) 1); +by (ALL_IND_TAC nat_ind (simp_tac nat_ss) 1); +by (fast_tac FOL_cs 1); +val le_imp_le_succ = result() RS mp; + +goal Nat2.thy "n m < succ(n)"; +by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1); +by (rtac (impI RS allI) 1); +by (ALL_IND_TAC nat_ind (simp_tac nat_ss) 1); +by (fast_tac FOL_cs 1); +result(); + +goal Nat2.thy "m <= n --> m <= n+k"; +by (IND_TAC nat_ind + (simp_tac (nat_ss addsimps [plus_0_right, plus_succ_right, le_imp_le_succ])) + "k" 1); +val le_plus = result(); + +goal Nat2.thy "succ(m) <= n --> m <= n"; +by (res_inst_tac [("x","n")]spec 1); +by (ALL_IND_TAC nat_exh (simp_tac (nat_ss addsimps [le_imp_le_succ])) 1); +val succ_le = result(); + +goal Nat2.thy "~m n<=m"; +by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1); +by (rtac (impI RS allI) 1); +by (ALL_IND_TAC nat_ind (asm_simp_tac nat_ss) 1); +val not_less = result(); + +goal Nat2.thy "n<=m --> ~m ~n<=m"; +by (cut_facts_tac [not_less] 1 THEN fast_tac FOL_cs 1); +val not_le = result(); + +goal Nat2.thy "m+k<=n --> m<=n"; +by (IND_TAC nat_ind (K all_tac) "k" 1); +by (simp_tac (nat_ss addsimps [plus_0_right]) 1); +by (rtac (impI RS allI) 1); +by (simp_tac (nat_ss addsimps [plus_succ_right]) 1); +by (REPEAT (resolve_tac [allI,impI] 1)); +by (cut_facts_tac [succ_le] 1); +by (fast_tac FOL_cs 1); +val plus_le = result(); + +val prems = goal Nat2.thy "[| ~m=0; m <= n |] ==> ~n=0"; +by (cut_facts_tac prems 1); +by (REPEAT (etac rev_mp 1)); +by (IND_TAC nat_exh (simp_tac nat_ss) "m" 1); +by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1); +val not0 = result(); + +goal Nat2.thy "a<=a' & b<=b' --> a+b<=a'+b'"; +by (IND_TAC nat_ind (simp_tac (nat_ss addsimps [plus_0_right,le_plus])) "b" 1); +by (resolve_tac [impI RS allI] 1); +by (resolve_tac [allI RS allI] 1); +by (ALL_IND_TAC nat_exh (asm_simp_tac (nat_ss addsimps [plus_succ_right])) 1); +val plus_le_plus = result(); + +goal Nat2.thy "i<=j --> j<=k --> i<=k"; +by (IND_TAC nat_ind (K all_tac) "i" 1); +by (simp_tac nat_ss 1); +by (resolve_tac [impI RS allI] 1); +by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1); +by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1); +by (fast_tac FOL_cs 1); +val le_trans = result(); + +goal Nat2.thy "i < j --> j <=k --> i < k"; +by (IND_TAC nat_ind (K all_tac) "j" 1); +by (simp_tac nat_ss 1); +by (resolve_tac [impI RS allI] 1); +by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1); +by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1); +by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1); +by (fast_tac FOL_cs 1); +val less_le_trans = result(); + +goal Nat2.thy "succ(i) <= j <-> i < j"; +by (IND_TAC nat_ind (simp_tac nat_ss) "j" 1); +by (resolve_tac [impI RS allI] 1); +by (ALL_IND_TAC nat_exh (asm_simp_tac nat_ss) 1); +val succ_le = result(); + +goal Nat2.thy "i i=j | i nat" + "0" :: "nat" ("0") + "+" :: "[nat,nat] => nat" (infixr 90) + "<","<=" :: "[nat,nat] => o" (infixr 70) + +rules + pred_0 "pred(0) = 0" + pred_succ "pred(succ(m)) = m" + + plus_0 "0+n = n" + plus_succ "succ(m)+n = succ(m+n)" + + nat_distinct1 "~ 0=succ(n)" + nat_distinct2 "~ succ(m)=0" + succ_inject "succ(m)=succ(n) <-> m=n" + + leq_0 "0 <= n" + leq_succ_succ "succ(m)<=succ(n) <-> m<=n" + leq_succ_0 "~ succ(m) <= 0" + + lt_0_succ "0 < succ(n)" + lt_succ_succ "succ(m) mP(succ(n)) |] ==> All(P)" +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/Prolog.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/Prolog.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,74 @@ +(* Title: FOL/ex/prolog.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For ex/prolog.thy. First-Order Logic: PROLOG examples +*) + +open Prolog; + +goal Prolog.thy "app(a:b:c:Nil, d:e:Nil, ?x)"; +by (resolve_tac [appNil,appCons] 1); +by (resolve_tac [appNil,appCons] 1); +by (resolve_tac [appNil,appCons] 1); +by (resolve_tac [appNil,appCons] 1); +prth (result()); + +goal Prolog.thy "app(?x, c:d:Nil, a:b:c:d:Nil)"; +by (REPEAT (resolve_tac [appNil,appCons] 1)); +result(); + + +goal Prolog.thy "app(?x, ?y, a:b:c:d:Nil)"; +by (REPEAT (resolve_tac [appNil,appCons] 1)); +back(); +back(); +back(); +back(); +result(); + + +(*app([x1,...,xn], y, ?z) requires (n+1) inferences*) +(*rev([x1,...,xn], ?y) requires (n+1)(n+2)/2 inferences*) + +goal Prolog.thy "rev(a:b:c:d:Nil, ?x)"; +val rules = [appNil,appCons,revNil,revCons]; +by (REPEAT (resolve_tac rules 1)); +result(); + +goal Prolog.thy "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:Nil, ?w)"; +by (REPEAT (resolve_tac rules 1)); +result(); + +goal Prolog.thy "rev(?x, a:b:c:Nil)"; +by (REPEAT (resolve_tac rules 1)); (*does not solve it directly!*) +back(); +back(); + +(*backtracking version*) +val prolog_tac = DEPTH_FIRST (has_fewer_prems 1) (resolve_tac rules 1); + +choplev 0; +by prolog_tac; +result(); + +goal Prolog.thy "rev(a:?x:c:?y:Nil, d:?z:b:?u)"; +by prolog_tac; +result(); + +(*rev([a..p], ?w) requires 153 inferences *) +goal Prolog.thy "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil, ?w)"; +by (DEPTH_SOLVE (resolve_tac ([refl,conjI]@rules) 1)); +(*Poly/ML: 4 secs >> 38 lips*) +result(); + +(*?x has 16, ?y has 32; rev(?y,?w) requires 561 (rather large) inferences; + total inferences = 2 + 1 + 17 + 561 = 581*) +goal Prolog.thy + "a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil = ?x & app(?x,?x,?y) & rev(?y,?w)"; +by (DEPTH_SOLVE (resolve_tac ([refl,conjI]@rules) 1)); +(*Poly/ML: 29 secs >> 20 lips*) +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/Prolog.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/Prolog.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,22 @@ +(* Title: FOL/ex/prolog.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +First-Order Logic: PROLOG examples + +Inherits from FOL the class term, the type o, and the coercion Trueprop +*) + +Prolog = FOL + +types list 1 +arities list :: (term)term +consts Nil :: "'a list" + ":" :: "['a, 'a list]=> 'a list" (infixr 60) + app :: "['a list, 'a list, 'a list] => o" + rev :: "['a list, 'a list] => o" +rules appNil "app(Nil,ys,ys)" + appCons "app(xs,ys,zs) ==> app(x:xs, ys, x:zs)" + revNil "rev(Nil,Nil)" + revCons "[| rev(xs,ys); app(ys, x:Nil, zs) |] ==> rev(x:xs, zs)" +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,39 @@ +(* Title: FOL/ex/ROOT + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Executes all examples for First-Order Logic. +*) + +writeln"Root file for FOL examples"; + +FOL_build_completed; (*Cause examples to fail if FOL did*) + +proof_timing := true; + +time_use "ex/intro.ML"; +time_use_thy "ex/nat"; +time_use "ex/foundn.ML"; +time_use_thy "ex/prolog"; + +writeln"\n** Intuitionistic examples **\n"; +time_use "ex/int.ML"; + +val thy = IFOL.thy and tac = Int.fast_tac 1; +time_use "ex/prop.ML"; +time_use "ex/quant.ML"; + +writeln"\n** Classical examples **\n"; +time_use "ex/cla.ML"; +time_use_thy "ex/if"; + +val thy = FOL.thy and tac = Cla.fast_tac FOL_cs 1; +time_use "ex/prop.ML"; +time_use "ex/quant.ML"; + +writeln"\n** Simplification examples **\n"; +time_use_thy "ex/nat2"; +time_use_thy "ex/list"; + +maketest"END: Root file for FOL examples"; diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/cla.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/cla.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,484 @@ +(* Title: FOL/ex/cla + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Classical First-Order Logic +*) + +writeln"File FOL/ex/cla."; + +open Cla; (*in case structure Int is open!*) + +goal FOL.thy "(P --> Q | R) --> (P-->Q) | (P-->R)"; +by (fast_tac FOL_cs 1); +result(); + +(*If and only if*) + +goal FOL.thy "(P<->Q) <-> (Q<->P)"; +by (fast_tac FOL_cs 1); +result(); + +goal FOL.thy "~ (P <-> ~P)"; +by (fast_tac FOL_cs 1); +result(); + + +(*Sample problems from + F. J. Pelletier, + Seventy-Five Problems for Testing Automatic Theorem Provers, + J. Automated Reasoning 2 (1986), 191-216. + Errata, JAR 4 (1988), 236-236. + +The hardest problems -- judging by experience with several theorem provers, +including matrix ones -- are 34 and 43. +*) + +writeln"Pelletier's examples"; +(*1*) +goal FOL.thy "(P-->Q) <-> (~Q --> ~P)"; +by (fast_tac FOL_cs 1); +result(); + +(*2*) +goal FOL.thy "~ ~ P <-> P"; +by (fast_tac FOL_cs 1); +result(); + +(*3*) +goal FOL.thy "~(P-->Q) --> (Q-->P)"; +by (fast_tac FOL_cs 1); +result(); + +(*4*) +goal FOL.thy "(~P-->Q) <-> (~Q --> P)"; +by (fast_tac FOL_cs 1); +result(); + +(*5*) +goal FOL.thy "((P|Q)-->(P|R)) --> (P|(Q-->R))"; +by (fast_tac FOL_cs 1); +result(); + +(*6*) +goal FOL.thy "P | ~ P"; +by (fast_tac FOL_cs 1); +result(); + +(*7*) +goal FOL.thy "P | ~ ~ ~ P"; +by (fast_tac FOL_cs 1); +result(); + +(*8. Peirce's law*) +goal FOL.thy "((P-->Q) --> P) --> P"; +by (fast_tac FOL_cs 1); +result(); + +(*9*) +goal FOL.thy "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"; +by (fast_tac FOL_cs 1); +result(); + +(*10*) +goal FOL.thy "(Q-->R) & (R-->P&Q) & (P-->Q|R) --> (P<->Q)"; +by (fast_tac FOL_cs 1); +result(); + +(*11. Proved in each direction (incorrectly, says Pelletier!!) *) +goal FOL.thy "P<->P"; +by (fast_tac FOL_cs 1); +result(); + +(*12. "Dijkstra's law"*) +goal FOL.thy "((P <-> Q) <-> R) <-> (P <-> (Q <-> R))"; +by (fast_tac FOL_cs 1); +result(); + +(*13. Distributive law*) +goal FOL.thy "P | (Q & R) <-> (P | Q) & (P | R)"; +by (fast_tac FOL_cs 1); +result(); + +(*14*) +goal FOL.thy "(P <-> Q) <-> ((Q | ~P) & (~Q|P))"; +by (fast_tac FOL_cs 1); +result(); + +(*15*) +goal FOL.thy "(P --> Q) <-> (~P | Q)"; +by (fast_tac FOL_cs 1); +result(); + +(*16*) +goal FOL.thy "(P-->Q) | (Q-->P)"; +by (fast_tac FOL_cs 1); +result(); + +(*17*) +goal FOL.thy "((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S))"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Classical Logic: examples with quantifiers"; + +goal FOL.thy "(ALL x. P(x) & Q(x)) <-> (ALL x. P(x)) & (ALL x. Q(x))"; +by (fast_tac FOL_cs 1); +result(); + +goal FOL.thy "(EX x. P-->Q(x)) <-> (P --> (EX x.Q(x)))"; +by (fast_tac FOL_cs 1); +result(); + +goal FOL.thy "(EX x.P(x)-->Q) <-> (ALL x.P(x)) --> Q"; +by (fast_tac FOL_cs 1); +result(); + +goal FOL.thy "(ALL x.P(x)) | Q <-> (ALL x. P(x) | Q)"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problems requiring quantifier duplication"; + +(*Needs multiple instantiation of ALL.*) +goal FOL.thy "(ALL x. P(x)-->P(f(x))) & P(d)-->P(f(f(f(d))))"; +by (best_tac FOL_dup_cs 1); +result(); + +(*Needs double instantiation of the quantifier*) +goal FOL.thy "EX x. P(x) --> P(a) & P(b)"; +by (best_tac FOL_dup_cs 1); +result(); + +goal FOL.thy "EX z. P(z) --> (ALL x. P(x))"; +by (best_tac FOL_dup_cs 1); +result(); + +(*from Vladimir Lifschitz, What Is the Inverse Method?, JAR 5 (1989), 1--23*) +goal FOL.thy "EX x x'. ALL y. EX z z'. \ +\ (~P(y,y) | P(x,x) | ~S(z,x)) & \ +\ (S(x,y) | ~S(y,z) | Q(z',z')) & \ +\ (Q(x',y) | ~Q(y,z') | S(x',x'))"; + +writeln"Hard examples with quantifiers"; + +writeln"Problem 18"; +goal FOL.thy "EX y. ALL x. P(y)-->P(x)"; +by (best_tac FOL_dup_cs 1); +result(); + +writeln"Problem 19"; +goal FOL.thy "EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))"; +by (best_tac FOL_dup_cs 1); +result(); + +writeln"Problem 20"; +goal FOL.thy "(ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w))) \ +\ --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"; +by (fast_tac FOL_cs 1); +goal FOL.thy "(ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w))) \ +\ --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 21"; +goal FOL.thy "(EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> (EX x. P<->Q(x))"; +by (best_tac FOL_dup_cs 1); +result(); + +writeln"Problem 22"; +goal FOL.thy "(ALL x. P <-> Q(x)) --> (P <-> (ALL x. Q(x)))"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 23"; +goal FOL.thy "(ALL x. P | Q(x)) <-> (P | (ALL x. Q(x)))"; +by (best_tac FOL_cs 1); +result(); + +writeln"Problem 24"; +goal FOL.thy "~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) & \ +\ ~(EX x.P(x)) --> (EX x.Q(x)) & (ALL x. Q(x)|R(x) --> S(x)) \ +\ --> (EX x. P(x)&R(x))"; +by (fast_tac FOL_cs 1); +(*slow*) +result(); + +writeln"Problem 25"; +goal FOL.thy "(EX x. P(x)) & \ +\ (ALL x. L(x) --> ~ (M(x) & R(x))) & \ +\ (ALL x. P(x) --> (M(x) & L(x))) & \ +\ ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x))) \ +\ --> (EX x. Q(x)&P(x))"; +by (best_tac FOL_cs 1); +(*slow*) +result(); + +writeln"Problem 26"; +goal FOL.thy "((EX x. p(x)) <-> (EX x. q(x))) & \ +\ (ALL x. ALL y. p(x) & q(y) --> (r(x) <-> s(y))) \ +\ --> ((ALL x. p(x)-->r(x)) <-> (ALL x. q(x)-->s(x)))"; +by (fast_tac FOL_cs 1); +(*slow*) +result(); + +writeln"Problem 27"; +goal FOL.thy "(EX x. P(x) & ~Q(x)) & \ +\ (ALL x. P(x) --> R(x)) & \ +\ (ALL x. M(x) & L(x) --> P(x)) & \ +\ ((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x))) \ +\ --> (ALL x. M(x) --> ~L(x))"; +by (fast_tac FOL_cs 1); +(*slow*) +result(); + +writeln"Problem 28. AMENDED"; +goal FOL.thy "(ALL x. P(x) --> (ALL x. Q(x))) & \ +\ ((ALL x. Q(x)|R(x)) --> (EX x. Q(x)&S(x))) & \ +\ ((EX x.S(x)) --> (ALL x. L(x) --> M(x))) \ +\ --> (ALL x. P(x) & L(x) --> M(x))"; +by (fast_tac FOL_cs 1); +(*slow*) +result(); + +writeln"Problem 29. Essentially the same as Principia Mathematica *11.71"; +goal FOL.thy "(EX x. P(x)) & (EX y. Q(y)) \ +\ --> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y)) <-> \ +\ (ALL x y. P(x) & Q(y) --> R(x) & S(y)))"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 30"; +goal FOL.thy "(ALL x. P(x) | Q(x) --> ~ R(x)) & \ +\ (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x)) \ +\ --> (ALL x. S(x))"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 31"; +goal FOL.thy "~(EX x.P(x) & (Q(x) | R(x))) & \ +\ (EX x. L(x) & P(x)) & \ +\ (ALL x. ~ R(x) --> M(x)) \ +\ --> (EX x. L(x) & M(x))"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 32"; +goal FOL.thy "(ALL x. P(x) & (Q(x)|R(x))-->S(x)) & \ +\ (ALL x. S(x) & R(x) --> L(x)) & \ +\ (ALL x. M(x) --> R(x)) \ +\ --> (ALL x. P(x) & M(x) --> L(x))"; +by (best_tac FOL_cs 1); +result(); + +writeln"Problem 33"; +goal FOL.thy "(ALL x. P(a) & (P(x)-->P(b))-->P(c)) <-> \ +\ (ALL x. (~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c)))"; +by (best_tac FOL_cs 1); +result(); + +writeln"Problem 34 AMENDED (TWICE!!) NOT PROVED AUTOMATICALLY"; +(*Andrews's challenge*) +goal FOL.thy "((EX x. ALL y. p(x) <-> p(y)) <-> \ +\ ((EX x. q(x)) <-> (ALL y. p(y)))) <-> \ +\ ((EX x. ALL y. q(x) <-> q(y)) <-> \ +\ ((EX x. p(x)) <-> (ALL y. q(y))))"; +by (safe_tac FOL_cs); (*22 secs*) +by (TRYALL (fast_tac FOL_cs)); (*128 secs*) +by (TRYALL (best_tac FOL_dup_cs)); (*77 secs*) +result(); + +writeln"Problem 35"; +goal FOL.thy "EX x y. P(x,y) --> (ALL u v. P(u,v))"; +by (best_tac FOL_dup_cs 1); +(*6.1 secs*) +result(); + +writeln"Problem 36"; +goal FOL.thy + "(ALL x. EX y. J(x,y)) & \ +\ (ALL x. EX y. G(x,y)) & \ +\ (ALL x y. J(x,y) | G(x,y) --> (ALL z. J(y,z) | G(y,z) --> H(x,z))) \ +\ --> (ALL x. EX y. H(x,y))"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 37"; +goal FOL.thy "(ALL z. EX w. ALL x. EX y. \ +\ (P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (EX u.Q(u,w)))) & \ +\ (ALL x z. ~P(x,z) --> (EX y. Q(y,z))) & \ +\ ((EX x y. Q(x,y)) --> (ALL x. R(x,x))) \ +\ --> (ALL x. EX y. R(x,y))"; +by (fast_tac FOL_cs 1); +(*slow...Poly/ML: 9.7 secs*) +result(); + +writeln"Problem 38. NOT PROVED"; +goal FOL.thy + "(ALL x. p(a) & (p(x) --> (EX y. p(y) & r(x,y))) --> \ +\ (EX z. EX w. p(z) & r(x,w) & r(w,z))) <-> \ +\ (ALL x. (~p(a) | p(x) | (EX z. EX w. p(z) & r(x,w) & r(w,z))) & \ +\ (~p(a) | ~(EX y. p(y) & r(x,y)) | \ +\ (EX z. EX w. p(z) & r(x,w) & r(w,z))))"; + +writeln"Problem 39"; +goal FOL.thy "~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 40. AMENDED"; +goal FOL.thy "(EX y. ALL x. F(x,y) <-> F(x,x)) --> \ +\ ~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 41"; +goal FOL.thy "(ALL z. EX y. ALL x. f(x,y) <-> f(x,z) & ~ f(x,x)) \ +\ --> ~ (EX z. ALL x. f(x,z))"; +by (best_tac FOL_cs 1); +result(); + +writeln"Problem 42 NOT PROVED"; +goal FOL.thy "~ (EX y. ALL x. p(x,y) <-> ~ (EX z. p(x,z) & p(z,x)))"; + +writeln"Problem 43 NOT PROVED AUTOMATICALLY"; +goal FOL.thy "(ALL x. ALL y. q(x,y) <-> (ALL z. p(z,x) <-> p(z,y))) \ +\ --> (ALL x. (ALL y. q(x,y) <-> q(y,x)))"; + + +writeln"Problem 44"; +goal FOL.thy "(ALL x. f(x) --> \ +\ (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y)))) & \ +\ (EX x. j(x) & (ALL y. g(y) --> h(x,y))) \ +\ --> (EX x. j(x) & ~f(x))"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 45"; +goal FOL.thy "(ALL x. f(x) & (ALL y. g(y) & h(x,y) --> j(x,y)) \ +\ --> (ALL y. g(y) & h(x,y) --> k(y))) & \ +\ ~ (EX y. l(y) & k(y)) & \ +\ (EX x. f(x) & (ALL y. h(x,y) --> l(y)) \ +\ & (ALL y. g(y) & h(x,y) --> j(x,y))) \ +\ --> (EX x. f(x) & ~ (EX y. g(y) & h(x,y)))"; +by (best_tac FOL_cs 1); +(*41.5 secs*) +result(); + + +writeln"Problems (mainly) involving equality or functions"; + +writeln"Problem 48"; +goal FOL.thy "(a=b | c=d) & (a=c | b=d) --> a=d | b=c"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 49 NOT PROVED AUTOMATICALLY"; +(*Hard because it involves substitution for Vars; + the type constraint ensures that x,y,z have the same type as a,b,u. *) +goal FOL.thy "(EX x y::'a. ALL z. z=x | z=y) & P(a) & P(b) & (~a=b) \ +\ --> (ALL u::'a.P(u))"; +by (safe_tac FOL_cs); +by (res_inst_tac [("x","a")] allE 1); +ba 1; +by (res_inst_tac [("x","b")] allE 1); +ba 1; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 50"; +(*What has this to do with equality?*) +goal FOL.thy "(ALL x. P(a,x) | (ALL y.P(x,y))) --> (EX x. ALL y.P(x,y))"; +by (best_tac FOL_dup_cs 1); +result(); + +writeln"Problem 51"; +goal FOL.thy + "(EX z w. ALL x y. P(x,y) <-> (x=z & y=w)) --> \ +\ (EX z. ALL x. EX w. (ALL y. P(x,y) <-> y=w) <-> x=z)"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 52"; +(*Almost the same as 51. *) +goal FOL.thy + "(EX z w. ALL x y. P(x,y) <-> (x=z & y=w)) --> \ +\ (EX w. ALL y. EX z. (ALL x. P(x,y) <-> x=z) <-> y=w)"; +by (best_tac FOL_cs 1); +result(); + +writeln"Problem 55"; + +(*Original, equational version by Len Schubert, via Pelletier *** NOT PROVED +goal FOL.thy + "(EX x. lives(x) & killed(x,agatha)) & \ +\ lives(agatha) & lives(butler) & lives(charles) & \ +\ (ALL x. lives(x) --> x=agatha | x=butler | x=charles) & \ +\ (ALL x y. killed(x,y) --> hates(x,y)) & \ +\ (ALL x y. killed(x,y) --> ~richer(x,y)) & \ +\ (ALL x. hates(agatha,x) --> ~hates(charles,x)) & \ +\ (ALL x. ~ x=butler --> hates(agatha,x)) & \ +\ (ALL x. ~richer(x,agatha) --> hates(butler,x)) & \ +\ (ALL x. hates(agatha,x) --> hates(butler,x)) & \ +\ (ALL x. EX y. ~hates(x,y)) & \ +\ ~ agatha=butler --> \ +\ killed(?who,agatha)"; +by (safe_tac FOL_cs); +by (dres_inst_tac [("x1","x")] (spec RS mp) 1); +ba 1; +be (spec RS exE) 1; +by (REPEAT (etac allE 1)); +by (fast_tac FOL_cs 1); +result(); +****) + +(*Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988). + fast_tac DISCOVERS who killed Agatha. *) +goal FOL.thy "lives(agatha) & lives(butler) & lives(charles) & \ +\ (killed(agatha,agatha) | killed(butler,agatha) | killed(charles,agatha)) & \ +\ (ALL x y. killed(x,y) --> hates(x,y) & ~richer(x,y)) & \ +\ (ALL x. hates(agatha,x) --> ~hates(charles,x)) & \ +\ (hates(agatha,agatha) & hates(agatha,charles)) & \ +\ (ALL x. lives(x) & ~richer(x,agatha) --> hates(butler,x)) & \ +\ (ALL x. hates(agatha,x) --> hates(butler,x)) & \ +\ (ALL x. ~hates(x,agatha) | ~hates(x,butler) | ~hates(x,charles)) --> \ +\ killed(?who,agatha)"; +by (fast_tac FOL_cs 1); +result(); + + +writeln"Problem 56"; +goal FOL.thy + "(ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 57"; +goal FOL.thy + "P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) & \ +\ (ALL x y z. P(x,y) & P(y,z) --> P(x,z)) --> P(f(a,b), f(a,c))"; +by (fast_tac FOL_cs 1); +result(); + +writeln"Problem 58 NOT PROVED AUTOMATICALLY"; +goal FOL.thy "(ALL x y. f(x)=g(y)) --> (ALL x y. f(f(x))=f(g(y)))"; +val f_cong = read_instantiate [("t","f")] subst_context; +by (fast_tac (FOL_cs addIs [f_cong]) 1); +result(); + +writeln"Problem 59"; +goal FOL.thy "(ALL x. P(x) <-> ~P(f(x))) --> (EX x. P(x) & ~P(f(x)))"; +by (best_tac FOL_dup_cs 1); +result(); + +writeln"Problem 60"; +goal FOL.thy + "ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"; +by (fast_tac FOL_cs 1); +result(); + + +writeln"Reached end of file."; + +(*Thu Jul 23 1992: loaded in 467.1s using iffE*) diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/foundn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/foundn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,134 @@ +(* Title: FOL/ex/foundn + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Intuitionistic FOL: Examples from The Foundation of a Generic Theorem Prover +*) + +writeln"File FOL/ex/foundn."; + +goal IFOL.thy "A&B --> (C-->A&C)"; +by (resolve_tac [impI] 1); +by (resolve_tac [impI] 1); +by (resolve_tac [conjI] 1); +by (assume_tac 2); +by (resolve_tac [conjunct1] 1); +by (assume_tac 1); +result(); + +(*A form of conj-elimination*) +val prems = +goal IFOL.thy "A&B ==> ([| A; B |] ==> C) ==> C"; +by (resolve_tac prems 1); +by (resolve_tac [conjunct1] 1); +by (resolve_tac prems 1); +by (resolve_tac [conjunct2] 1); +by (resolve_tac prems 1); +result(); + + +val prems = +goal IFOL.thy "(!!A. ~ ~A ==> A) ==> B | ~B"; +by (resolve_tac prems 1); +by (resolve_tac [notI] 1); +by (res_inst_tac [ ("P", "~B") ] notE 1); +by (resolve_tac [notI] 2); +by (res_inst_tac [ ("P", "B | ~B") ] notE 2); +by (assume_tac 2); +by (resolve_tac [disjI1] 2); +by (assume_tac 2); +by (resolve_tac [notI] 1); +by (res_inst_tac [ ("P", "B | ~B") ] notE 1); +by (assume_tac 1); +by (resolve_tac [disjI2] 1); +by (assume_tac 1); +result(); + + +val prems = +goal IFOL.thy "(!!A. ~ ~A ==> A) ==> B | ~B"; +by (resolve_tac prems 1); +by (resolve_tac [notI] 1); +by (resolve_tac [notE] 1); +by (resolve_tac [notI] 2); +by (eresolve_tac [notE] 2); +by (eresolve_tac [disjI1] 2); +by (resolve_tac [notI] 1); +by (eresolve_tac [notE] 1); +by (eresolve_tac [disjI2] 1); +result(); + + +val prems = +goal IFOL.thy "[| A | ~A; ~ ~A |] ==> A"; +by (resolve_tac [disjE] 1); +by (resolve_tac prems 1); +by (assume_tac 1); +by (resolve_tac [FalseE] 1); +by (res_inst_tac [ ("P", "~A") ] notE 1); +by (resolve_tac prems 1); +by (assume_tac 1); +result(); + + +writeln"Examples with quantifiers"; + +val prems = +goal IFOL.thy "ALL z. G(z) ==> ALL z. G(z)|H(z)"; +by (resolve_tac [allI] 1); +by (resolve_tac [disjI1] 1); +by (resolve_tac (prems RL [spec]) 1); + (*can use instead + by (resolve_tac [spec] 1); by (resolve_tac prems 1); *) +result(); + + +goal IFOL.thy "ALL x. EX y. x=y"; +by (resolve_tac [allI] 1); +by (resolve_tac [exI] 1); +by (resolve_tac [refl] 1); +result(); + + +goal IFOL.thy "EX y. ALL x. x=y"; +by (resolve_tac [exI] 1); +by (resolve_tac [allI] 1); +by (resolve_tac [refl] 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + + +(*Parallel lifting example. *) +goal IFOL.thy "EX u.ALL x.EX v.ALL y.EX w. P(u,x,v,y,w)"; +by (resolve_tac [exI, allI] 1); +by (resolve_tac [exI, allI] 1); +by (resolve_tac [exI, allI] 1); +by (resolve_tac [exI, allI] 1); +by (resolve_tac [exI, allI] 1); + + +val prems = +goal IFOL.thy "(EX z.F(z)) & B ==> (EX z. F(z) & B)"; +by (resolve_tac [conjE] 1); +by (resolve_tac prems 1); +by (resolve_tac [exE] 1); +by (assume_tac 1); +by (resolve_tac [exI] 1); +by (resolve_tac [conjI] 1); +by (assume_tac 1); +by (assume_tac 1); +result(); + + +(*A bigger demonstration of quantifiers -- not in the paper*) +goal IFOL.thy "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"; +by (resolve_tac [impI] 1); +by (resolve_tac [allI] 1); +by (resolve_tac [exE] 1 THEN assume_tac 1); +by (resolve_tac [exI] 1); +by (resolve_tac [allE] 1 THEN assume_tac 1); +by (assume_tac 1); +result(); + + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/if.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/if.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,64 @@ +(* Title: FOL/ex/if + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For ex/if.thy. First-Order Logic: the 'if' example +*) + +open If; +open Cla; (*in case structure Int is open!*) + +val prems = goalw If.thy [if_def] + "[| P ==> Q; ~P ==> R |] ==> if(P,Q,R)"; +by (fast_tac (FOL_cs addIs prems) 1); +val ifI = result(); + +val major::prems = goalw If.thy [if_def] + "[| if(P,Q,R); [| P; Q |] ==> S; [| ~P; R |] ==> S |] ==> S"; +by (cut_facts_tac [major] 1); +by (fast_tac (FOL_cs addIs prems) 1); +val ifE = result(); + + +goal If.thy + "if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))"; +by (resolve_tac [iffI] 1); +by (eresolve_tac [ifE] 1); +by (eresolve_tac [ifE] 1); +by (resolve_tac [ifI] 1); +by (resolve_tac [ifI] 1); + +choplev 0; +val if_cs = FOL_cs addSIs [ifI] addSEs[ifE]; +by (fast_tac if_cs 1); +val if_commute = result(); + + +goal If.thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))"; +by (fast_tac if_cs 1); +val nested_ifs = result(); + +choplev 0; +by (rewrite_goals_tac [if_def]); +by (fast_tac FOL_cs 1); +result(); + + +(*An invalid formula. High-level rules permit a simpler diagnosis*) +goal If.thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))"; +by (fast_tac if_cs 1) handle ERROR => writeln"Failed, as expected"; +(*Check that subgoals remain: proof failed.*) +getgoal 1; +by (REPEAT (step_tac if_cs 1)); + +choplev 0; +by (rewrite_goals_tac [if_def]); +by (fast_tac FOL_cs 1) handle ERROR => writeln"Failed, as expected"; +(*Check that subgoals remain: proof failed.*) +getgoal 1; +by (REPEAT (step_tac FOL_cs 1)); + + + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/if.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/if.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,5 @@ +If = FOL + +consts if :: "[o,o,o]=>o" +rules + if_def "if(P,Q,R) == P&Q | ~P&R" +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/int.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/int.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,421 @@ +(* Title: FOL/ex/int + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Intuitionistic First-Order Logic + +Single-step commands: +by (Int.step_tac 1); +by (biresolve_tac safe_brls 1); +by (biresolve_tac haz_brls 1); +by (assume_tac 1); +by (Int.safe_tac 1); +by (Int.mp_tac 1); +by (Int.fast_tac 1); +*) + +writeln"File FOL/ex/int."; + +(*Note: for PROPOSITIONAL formulae... + ~A is classically provable iff it is intuitionistically provable. + Therefore A is classically provable iff ~~A is intuitionistically provable. + +Let Q be the conjuction of the propositions A|~A, one for each atom A in +P. If P is provable classically, then clearly P&Q is provable +intuitionistically, so ~~(P&Q) is also provable intuitionistically. +The latter is intuitionistically equivalent to ~~P&~~Q, hence to ~~P, +since ~~Q is intuitionistically provable. Finally, if P is a negation then +~~P is intuitionstically equivalent to P. [Andy Pitts] +*) + +goal IFOL.thy "~~(P&Q) <-> ~~P & ~~Q"; +by (Int.fast_tac 1); +result(); + +goal IFOL.thy "~~~P <-> ~P"; +by (Int.fast_tac 1); +result(); + +goal IFOL.thy "~~((P --> Q | R) --> (P-->Q) | (P-->R))"; +by (Int.fast_tac 1); +result(); + +goal IFOL.thy "(P<->Q) <-> (Q<->P)"; +by (Int.fast_tac 1); +result(); + + +writeln"Lemmas for the propositional double-negation translation"; + +goal IFOL.thy "P --> ~~P"; +by (Int.fast_tac 1); +result(); + +goal IFOL.thy "~~(~~P --> P)"; +by (Int.fast_tac 1); +result(); + +goal IFOL.thy "~~P & ~~(P --> Q) --> ~~Q"; +by (Int.fast_tac 1); +result(); + + +writeln"The following are classically but not constructively valid."; + +(*The attempt to prove them terminates quickly!*) +goal IFOL.thy "((P-->Q) --> P) --> P"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +(*Check that subgoals remain: proof failed.*) +getgoal 1; + +goal IFOL.thy "(P&Q-->R) --> (P-->R) | (Q-->R)"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + + +writeln"Intuitionistic FOL: propositional problems based on Pelletier."; + +writeln"Problem ~~1"; +goal IFOL.thy "~~((P-->Q) <-> (~Q --> ~P))"; +by (Int.fast_tac 1); +result(); +(*5 secs*) + + +writeln"Problem ~~2"; +goal IFOL.thy "~~(~~P <-> P)"; +by (Int.fast_tac 1); +result(); +(*1 secs*) + + +writeln"Problem 3"; +goal IFOL.thy "~(P-->Q) --> (Q-->P)"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~4"; +goal IFOL.thy "~~((~P-->Q) <-> (~Q --> P))"; +by (Int.fast_tac 1); +result(); +(*9 secs*) + +writeln"Problem ~~5"; +goal IFOL.thy "~~((P|Q-->P|R) --> P|(Q-->R))"; +by (Int.fast_tac 1); +result(); +(*10 secs*) + + +writeln"Problem ~~6"; +goal IFOL.thy "~~(P | ~P)"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~7"; +goal IFOL.thy "~~(P | ~~~P)"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~8. Peirce's law"; +goal IFOL.thy "~~(((P-->Q) --> P) --> P)"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 9"; +goal IFOL.thy "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"; +by (Int.fast_tac 1); +result(); +(*9 secs*) + + +writeln"Problem 10"; +goal IFOL.thy "(Q-->R) --> (R-->P&Q) --> (P-->(Q|R)) --> (P<->Q)"; +by (Int.fast_tac 1); +result(); + +writeln"11. Proved in each direction (incorrectly, says Pelletier!!) "; +goal IFOL.thy "P<->P"; +by (Int.fast_tac 1); + +writeln"Problem ~~12. Dijkstra's law "; +goal IFOL.thy "~~(((P <-> Q) <-> R) <-> (P <-> (Q <-> R)))"; +by (Int.fast_tac 1); +result(); + +goal IFOL.thy "((P <-> Q) <-> R) --> ~~(P <-> (Q <-> R))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 13. Distributive law"; +goal IFOL.thy "P | (Q & R) <-> (P | Q) & (P | R)"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~14"; +goal IFOL.thy "~~((P <-> Q) <-> ((Q | ~P) & (~Q|P)))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~15"; +goal IFOL.thy "~~((P --> Q) <-> (~P | Q))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~16"; +goal IFOL.thy "~~((P-->Q) | (Q-->P))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~17"; +goal IFOL.thy + "~~(((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S)))"; +by (Int.fast_tac 1); +result(); + +(*Dijkstra's "Golden Rule"*) +goal IFOL.thy "(P&Q) <-> P <-> Q <-> (P|Q)"; +by (Int.fast_tac 1); +result(); + + +writeln"U****Examples with quantifiers****"; + + +writeln"The converse is classical in the following implications..."; + +goal IFOL.thy "(EX x.P(x)-->Q) --> (ALL x.P(x)) --> Q"; +by (Int.fast_tac 1); +result(); + +goal IFOL.thy "((ALL x.P(x))-->Q) --> ~ (ALL x. P(x) & ~Q)"; +by (Int.fast_tac 1); +result(); + +goal IFOL.thy "((ALL x. ~P(x))-->Q) --> ~ (ALL x. ~ (P(x)|Q))"; +by (Int.fast_tac 1); +result(); + +goal IFOL.thy "(ALL x.P(x)) | Q --> (ALL x. P(x) | Q)"; +by (Int.fast_tac 1); +result(); + +goal IFOL.thy "(EX x. P --> Q(x)) --> (P --> (EX x. Q(x)))"; +by (Int.fast_tac 1); +result(); + + + + +writeln"The following are not constructively valid!"; +(*The attempt to prove them terminates quickly!*) + +goal IFOL.thy "((ALL x.P(x))-->Q) --> (EX x.P(x)-->Q)"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + +goal IFOL.thy "(P --> (EX x.Q(x))) --> (EX x. P-->Q(x))"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + +goal IFOL.thy "(ALL x. P(x) | Q) --> ((ALL x.P(x)) | Q)"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + +goal IFOL.thy "(ALL x. ~~P(x)) --> ~~(ALL x. P(x))"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + +(*Classically but not intuitionistically valid. Proved by a bug in 1986!*) +goal IFOL.thy "EX x. Q(x) --> (ALL x. Q(x))"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + + +writeln"Hard examples with quantifiers"; + +(*The ones that have not been proved are not known to be valid! + Some will require quantifier duplication -- not currently available*) + +writeln"Problem ~~18"; +goal IFOL.thy "~~(EX y. ALL x. P(y)-->P(x))"; +(*NOT PROVED*) + +writeln"Problem ~~19"; +goal IFOL.thy "~~(EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x)))"; +(*NOT PROVED*) + +writeln"Problem 20"; +goal IFOL.thy "(ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w))) \ +\ --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 21"; +goal IFOL.thy "(EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> ~~(EX x. P<->Q(x))"; +(*NOT PROVED*) + +writeln"Problem 22"; +goal IFOL.thy "(ALL x. P <-> Q(x)) --> (P <-> (ALL x. Q(x)))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~23"; +goal IFOL.thy "~~ ((ALL x. P | Q(x)) <-> (P | (ALL x. Q(x))))"; +by (Int.best_tac 1); +result(); + +writeln"Problem 24"; +goal IFOL.thy "~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) & \ +\ ~(EX x.P(x)) --> (EX x.Q(x)) & (ALL x. Q(x)|R(x) --> S(x)) \ +\ --> (EX x. P(x)&R(x))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 25"; +goal IFOL.thy "(EX x. P(x)) & \ +\ (ALL x. L(x) --> ~ (M(x) & R(x))) & \ +\ (ALL x. P(x) --> (M(x) & L(x))) & \ +\ ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x))) \ +\ --> (EX x. Q(x)&P(x))"; +by (Int.best_tac 1); +result(); + +writeln"Problem ~~26"; +goal IFOL.thy "(~~(EX x. p(x)) <-> ~~(EX x. q(x))) & \ +\ (ALL x. ALL y. p(x) & q(y) --> (r(x) <-> s(y))) \ +\ --> ((ALL x. p(x)-->r(x)) <-> (ALL x. q(x)-->s(x)))"; +(*NOT PROVED*) + +writeln"Problem 27"; +goal IFOL.thy "(EX x. P(x) & ~Q(x)) & \ +\ (ALL x. P(x) --> R(x)) & \ +\ (ALL x. M(x) & L(x) --> P(x)) & \ +\ ((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x))) \ +\ --> (ALL x. M(x) --> ~L(x))"; +by (Int.fast_tac 1); (*44 secs*) +result(); + +writeln"Problem ~~28. AMENDED"; +goal IFOL.thy "(ALL x. P(x) --> (ALL x. Q(x))) & \ +\ (~~(ALL x. Q(x)|R(x)) --> (EX x. Q(x)&S(x))) & \ +\ (~~(EX x.S(x)) --> (ALL x. L(x) --> M(x))) \ +\ --> (ALL x. P(x) & L(x) --> M(x))"; +by (Int.fast_tac 1); (*101 secs*) +result(); + +writeln"Problem 29. Essentially the same as Principia Mathematica *11.71"; +goal IFOL.thy "(EX x. P(x)) & (EX y. Q(y)) \ +\ --> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y)) <-> \ +\ (ALL x y. P(x) & Q(y) --> R(x) & S(y)))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~30"; +goal IFOL.thy "(ALL x. (P(x) | Q(x)) --> ~ R(x)) & \ +\ (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x)) \ +\ --> (ALL x. ~~S(x))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 31"; +goal IFOL.thy "~(EX x.P(x) & (Q(x) | R(x))) & \ +\ (EX x. L(x) & P(x)) & \ +\ (ALL x. ~ R(x) --> M(x)) \ +\ --> (EX x. L(x) & M(x))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 32"; +goal IFOL.thy "(ALL x. P(x) & (Q(x)|R(x))-->S(x)) & \ +\ (ALL x. S(x) & R(x) --> L(x)) & \ +\ (ALL x. M(x) --> R(x)) \ +\ --> (ALL x. P(x) & M(x) --> L(x))"; +by (Int.best_tac 1); +result(); + +writeln"Problem ~~33"; +goal IFOL.thy "(ALL x. ~~(P(a) & (P(x)-->P(b))-->P(c))) <-> \ +\ (ALL x. ~~((~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c))))"; +by (Int.best_tac 1); +result(); + + +writeln"Problem 36"; +goal IFOL.thy + "(ALL x. EX y. J(x,y)) & \ +\ (ALL x. EX y. G(x,y)) & \ +\ (ALL x y. J(x,y) | G(x,y) --> (ALL z. J(y,z) | G(y,z) --> H(x,z))) \ +\ --> (ALL x. EX y. H(x,y))"; +by (Int.fast_tac 1); (*35 secs*) +result(); + +writeln"Problem 37"; +goal IFOL.thy + "(ALL z. EX w. ALL x. EX y. \ +\ ~~(P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (EX u.Q(u,w)))) & \ +\ (ALL x z. ~P(x,z) --> (EX y. Q(y,z))) & \ +\ (~~(EX x y. Q(x,y)) --> (ALL x. R(x,x))) \ +\ --> ~~(ALL x. EX y. R(x,y))"; +(*NOT PROVED*) + +writeln"Problem 39"; +goal IFOL.thy "~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 40. AMENDED"; +goal IFOL.thy "(EX y. ALL x. F(x,y) <-> F(x,x)) --> \ +\ ~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 44"; +goal IFOL.thy "(ALL x. f(x) --> \ +\ (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y)))) & \ +\ (EX x. j(x) & (ALL y. g(y) --> h(x,y))) \ +\ --> (EX x. j(x) & ~f(x))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 48"; +goal IFOL.thy "(a=b | c=d) & (a=c | b=d) --> a=d | b=c"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 51"; +goal IFOL.thy + "(EX z w. ALL x y. P(x,y) <-> (x=z & y=w)) --> \ +\ (EX z. ALL x. EX w. (ALL y. P(x,y) <-> y=w) <-> x=z)"; +by (Int.best_tac 1); (*60 seconds*) +result(); + +writeln"Problem 52"; +(*Almost the same as 51. *) +goal IFOL.thy + "(EX z w. ALL x y. P(x,y) <-> (x=z & y=w)) --> \ +\ (EX w. ALL y. EX z. (ALL x. P(x,y) <-> x=z) <-> y=w)"; +by (Int.best_tac 1); (*60 seconds*) +result(); + +writeln"Problem 56"; +goal IFOL.thy + "(ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 57"; +goal IFOL.thy + "P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) & \ +\ (ALL x y z. P(x,y) & P(y,z) --> P(x,z)) --> P(f(a,b), f(a,c))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 60"; +goal IFOL.thy + "ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"; +by (Int.fast_tac 1); +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/intro.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/intro.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,92 @@ +(* Title: FOL/ex/intro + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Examples for the manual "Introduction to Isabelle" + +Derives some inference rules, illustrating the use of definitions + +To generate similar output to manual, execute these commands: + Pretty.setmargin 72; print_depth 0; +*) + + +(**** Some simple backward proofs ****) + +goal FOL.thy "P|P --> P"; +by (resolve_tac [impI] 1); +by (resolve_tac [disjE] 1); +by (assume_tac 3); +by (assume_tac 2); +by (assume_tac 1); +val mythm = result(); + +goal FOL.thy "(P & Q) | R --> (P | R)"; +by (resolve_tac [impI] 1); +by (eresolve_tac [disjE] 1); +by (dresolve_tac [conjunct1] 1); +by (resolve_tac [disjI1] 1); +by (resolve_tac [disjI2] 2); +by (REPEAT (assume_tac 1)); +result(); + +(*Correct version, delaying use of "spec" until last*) +goal FOL.thy "(ALL x y.P(x,y)) --> (ALL z w.P(w,z))"; +by (resolve_tac [impI] 1); +by (resolve_tac [allI] 1); +by (resolve_tac [allI] 1); +by (dresolve_tac [spec] 1); +by (dresolve_tac [spec] 1); +by (assume_tac 1); +result(); + + +(**** Demonstration of fast_tac ****) + +goal FOL.thy "(EX y. ALL x. J(y,x) <-> ~J(x,x)) \ +\ --> ~ (ALL x. EX y. ALL z. J(z,y) <-> ~ J(z,x))"; +by (fast_tac FOL_cs 1); +result(); + +goal FOL.thy "ALL x. P(x,f(x)) <-> \ +\ (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"; +by (fast_tac FOL_cs 1); +result(); + + +(**** Derivation of conjunction elimination rule ****) + +val [major,minor] = goal FOL.thy "[| P&Q; [| P; Q |] ==> R |] ==> R"; +by (resolve_tac [minor] 1); +by (resolve_tac [major RS conjunct1] 1); +by (resolve_tac [major RS conjunct2] 1); +prth (topthm()); +result(); + + +(**** Derived rules involving definitions ****) + +(** Derivation of negation introduction **) + +val prems = goal FOL.thy "(P ==> False) ==> ~P"; +by (rewrite_goals_tac [not_def]); +by (resolve_tac [impI] 1); +by (resolve_tac prems 1); +by (assume_tac 1); +result(); + +val [major,minor] = goal FOL.thy "[| ~P; P |] ==> R"; +by (resolve_tac [FalseE] 1); +by (resolve_tac [mp] 1); +by (resolve_tac [rewrite_rule [not_def] major] 1); +by (resolve_tac [minor] 1); +result(); + +(*Alternative proof of above result*) +val [major,minor] = goalw FOL.thy [not_def] + "[| ~P; P |] ==> R"; +by (resolve_tac [minor RS (major RS mp RS FalseE)] 1); +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/list.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/list.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,67 @@ +(* Title: FOL/ex/list + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +For ex/list.thy. Examples of simplification and induction on lists +*) + +open List; + +val prems = goal List.thy "[| P([]); !!x l. P(x.l) |] ==> All(P)"; +by (rtac list_ind 1); +by (REPEAT (resolve_tac (prems@[allI,impI]) 1)); +val list_exh = result(); + +val list_rew_thms = [list_distinct1,list_distinct2,app_nil,app_cons, + hd_eq,tl_eq,forall_nil,forall_cons,list_free, + len_nil,len_cons,at_0,at_succ]; + +val list_ss = nat_ss addsimps list_rew_thms; + +goal List.thy "~l=[] --> hd(l).tl(l) = l"; +by(IND_TAC list_exh (simp_tac list_ss) "l" 1); +result(); + +goal List.thy "(l1++l2)++l3 = l1++(l2++l3)"; +by(IND_TAC list_ind (simp_tac list_ss) "l1" 1); +val append_assoc = result(); + +goal List.thy "l++[] = l"; +by(IND_TAC list_ind (simp_tac list_ss) "l" 1); +val app_nil_right = result(); + +goal List.thy "l1++l2=[] <-> l1=[] & l2=[]"; +by(IND_TAC list_exh (simp_tac list_ss) "l1" 1); +val app_eq_nil_iff = result(); + +goal List.thy "forall(l++l',P) <-> forall(l,P) & forall(l',P)"; +by(IND_TAC list_ind (simp_tac list_ss) "l" 1); +val forall_app = result(); + +goal List.thy "forall(l,%x.P(x)&Q(x)) <-> forall(l,P) & forall(l,Q)"; +by(IND_TAC list_ind (simp_tac list_ss) "l" 1); +by(fast_tac FOL_cs 1); +val forall_conj = result(); + +goal List.thy "~l=[] --> forall(l,P) <-> P(hd(l)) & forall(tl(l),P)"; +by(IND_TAC list_ind (simp_tac list_ss) "l" 1); +val forall_ne = result(); + +(*** Lists with natural numbers ***) + +goal List.thy "len(l1++l2) = len(l1)+len(l2)"; +by (IND_TAC list_ind (simp_tac list_ss) "l1" 1); +val len_app = result(); + +goal List.thy "i at(l1++l2,i) = at(l1,i)"; +by (IND_TAC list_ind (simp_tac list_ss) "l1" 1); +by (REPEAT (rtac allI 1)); +by (rtac impI 1); +by (ALL_IND_TAC nat_exh (asm_simp_tac list_ss) 1); +val at_app1 = result(); + +goal List.thy "at(l1++(x.l2), len(l1)) = x"; +by (IND_TAC list_ind (simp_tac list_ss) "l1" 1); +val at_app_hd2 = result(); + diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/list.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/list.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,49 @@ +(* Title: FOL/ex/list + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +Examples of simplification and induction on lists +*) + +List = Nat2 + + +types list 1 +arities list :: (term)term + +consts + hd :: "'a list => 'a" + tl :: "'a list => 'a list" + forall :: "['a list, 'a => o] => o" + len :: "'a list => nat" + at :: "['a list, nat] => 'a" + "[]" :: "'a list" ("[]") + "." :: "['a, 'a list] => 'a list" (infixr 80) + "++" :: "['a list, 'a list] => 'a list" (infixr 70) + +rules + list_ind "[| P([]); ALL x l. P(l)-->P(x.l) |] ==> All(P)" + + forall_cong + "[| l = l'; !!x. P(x)<->P'(x) |] ==> forall(l,P) <-> forall(l',P')" + + list_distinct1 "~[] = x.l" + list_distinct2 "~x.l = []" + + list_free "x.l = x'.l' <-> x=x' & l=l'" + + app_nil "[]++l = l" + app_cons "(x.l)++l' = x.(l++l')" + tl_eq "tl(m.q) = q" + hd_eq "hd(m.q) = m" + + forall_nil "forall([],P)" + forall_cons "forall(x.l,P) <-> P(x) & forall(l,P)" + + len_nil "len([]) = 0" + len_cons "len(m.q) = succ(len(q))" + + at_0 "at(m.q,0) = m" + at_succ "at(m.q,succ(n)) = at(q,n)" + +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/nat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/nat.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,69 @@ +(* Title: FOL/ex/nat.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Examples for the manual "Introduction to Isabelle" + +Proofs about the natural numbers + +INCOMPATIBLE with nat2.ML, Nipkow's examples + +To generate similar output to manual, execute these commands: + Pretty.setmargin 72; print_depth 0; +*) + +open Nat; + +goal Nat.thy "~ (Suc(k) = k)"; +by (res_inst_tac [("n","k")] induct 1); +by (resolve_tac [notI] 1); +by (eresolve_tac [Suc_neq_0] 1); +by (resolve_tac [notI] 1); +by (eresolve_tac [notE] 1); +by (eresolve_tac [Suc_inject] 1); +val Suc_n_not_n = result(); + + +goal Nat.thy "(k+m)+n = k+(m+n)"; +prths ([induct] RL [topthm()]); (*prints all 14 next states!*) +by (resolve_tac [induct] 1); +back(); +back(); +back(); +back(); +back(); +back(); + +goalw Nat.thy [add_def] "0+n = n"; +by (resolve_tac [rec_0] 1); +val add_0 = result(); + +goalw Nat.thy [add_def] "Suc(m)+n = Suc(m+n)"; +by (resolve_tac [rec_Suc] 1); +val add_Suc = result(); + +val add_ss = FOL_ss addsimps [add_0, add_Suc]; + +goal Nat.thy "(k+m)+n = k+(m+n)"; +by (res_inst_tac [("n","k")] induct 1); +by (simp_tac add_ss 1); +by (asm_simp_tac add_ss 1); +val add_assoc = result(); + +goal Nat.thy "m+0 = m"; +by (res_inst_tac [("n","m")] induct 1); +by (simp_tac add_ss 1); +by (asm_simp_tac add_ss 1); +val add_0_right = result(); + +goal Nat.thy "m+Suc(n) = Suc(m+n)"; +by (res_inst_tac [("n","m")] induct 1); +by (ALLGOALS (asm_simp_tac add_ss)); +val add_Suc_right = result(); + +val [prem] = goal Nat.thy "(!!n. f(Suc(n)) = Suc(f(n))) ==> f(i+j) = i+f(j)"; +by (res_inst_tac [("n","i")] induct 1); +by (simp_tac add_ss 1); +by (asm_simp_tac (add_ss addsimps [prem]) 1); +result(); diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/nat.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,26 @@ +(* Title: FOL/ex/nat.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Examples for the manual "Introduction to Isabelle" + +Theory of the natural numbers: Peano's axioms, primitive recursion + +INCOMPATIBLE with nat2.thy, Nipkow's example +*) + +Nat = FOL + +types nat 0 +arities nat :: term +consts "0" :: "nat" ("0") + Suc :: "nat=>nat" + rec :: "[nat, 'a, [nat,'a]=>'a] => 'a" + "+" :: "[nat, nat] => nat" (infixl 60) +rules 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))" + add_def "m+n == rec(m, n, %x y. Suc(y))" +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/nat2.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/nat2.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,163 @@ +(* Title: FOL/ex/nat2.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +For ex/nat.thy. +Examples of simplification and induction on the natural numbers +*) + +open Nat2; + +val nat_rews = [pred_0, pred_succ, plus_0, plus_succ, + nat_distinct1, nat_distinct2, succ_inject, + leq_0,leq_succ_succ,leq_succ_0, + lt_0_succ,lt_succ_succ,lt_0]; + +val nat_ss = FOL_ss addsimps nat_rews; + +val prems = goal Nat2.thy + "[| P(0); !!x. P(succ(x)) |] ==> All(P)"; +by (rtac nat_ind 1); +by (REPEAT (resolve_tac (prems@[allI,impI]) 1)); +val nat_exh = result(); + +goal Nat2.thy "~ n=succ(n)"; +by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1); +result(); + +goal Nat2.thy "~ succ(n)=n"; +by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1); +result(); + +goal Nat2.thy "~ succ(succ(n))=n"; +by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1); +result(); + +goal Nat2.thy "~ n=succ(succ(n))"; +by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1); +result(); + +goal Nat2.thy "m+0 = m"; +by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1); +val plus_0_right = result(); + +goal Nat2.thy "m+succ(n) = succ(m+n)"; +by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1); +val plus_succ_right = result(); + +goal Nat2.thy "~n=0 --> m+pred(n) = pred(m+n)"; +by (IND_TAC nat_ind (simp_tac (nat_ss addsimps [plus_succ_right])) "n" 1); +result(); + +goal Nat2.thy "~n=0 --> succ(pred(n))=n"; +by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1); +result(); + +goal Nat2.thy "m+n=0 <-> m=0 & n=0"; +by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1); +result(); + +goal Nat2.thy "m <= n --> m <= succ(n)"; +by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1); +by (rtac (impI RS allI) 1); +by (ALL_IND_TAC nat_ind (simp_tac nat_ss) 1); +by (fast_tac FOL_cs 1); +val le_imp_le_succ = result() RS mp; + +goal Nat2.thy "n m < succ(n)"; +by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1); +by (rtac (impI RS allI) 1); +by (ALL_IND_TAC nat_ind (simp_tac nat_ss) 1); +by (fast_tac FOL_cs 1); +result(); + +goal Nat2.thy "m <= n --> m <= n+k"; +by (IND_TAC nat_ind + (simp_tac (nat_ss addsimps [plus_0_right, plus_succ_right, le_imp_le_succ])) + "k" 1); +val le_plus = result(); + +goal Nat2.thy "succ(m) <= n --> m <= n"; +by (res_inst_tac [("x","n")]spec 1); +by (ALL_IND_TAC nat_exh (simp_tac (nat_ss addsimps [le_imp_le_succ])) 1); +val succ_le = result(); + +goal Nat2.thy "~m n<=m"; +by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1); +by (rtac (impI RS allI) 1); +by (ALL_IND_TAC nat_ind (asm_simp_tac nat_ss) 1); +val not_less = result(); + +goal Nat2.thy "n<=m --> ~m ~n<=m"; +by (cut_facts_tac [not_less] 1 THEN fast_tac FOL_cs 1); +val not_le = result(); + +goal Nat2.thy "m+k<=n --> m<=n"; +by (IND_TAC nat_ind (K all_tac) "k" 1); +by (simp_tac (nat_ss addsimps [plus_0_right]) 1); +by (rtac (impI RS allI) 1); +by (simp_tac (nat_ss addsimps [plus_succ_right]) 1); +by (REPEAT (resolve_tac [allI,impI] 1)); +by (cut_facts_tac [succ_le] 1); +by (fast_tac FOL_cs 1); +val plus_le = result(); + +val prems = goal Nat2.thy "[| ~m=0; m <= n |] ==> ~n=0"; +by (cut_facts_tac prems 1); +by (REPEAT (etac rev_mp 1)); +by (IND_TAC nat_exh (simp_tac nat_ss) "m" 1); +by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1); +val not0 = result(); + +goal Nat2.thy "a<=a' & b<=b' --> a+b<=a'+b'"; +by (IND_TAC nat_ind (simp_tac (nat_ss addsimps [plus_0_right,le_plus])) "b" 1); +by (resolve_tac [impI RS allI] 1); +by (resolve_tac [allI RS allI] 1); +by (ALL_IND_TAC nat_exh (asm_simp_tac (nat_ss addsimps [plus_succ_right])) 1); +val plus_le_plus = result(); + +goal Nat2.thy "i<=j --> j<=k --> i<=k"; +by (IND_TAC nat_ind (K all_tac) "i" 1); +by (simp_tac nat_ss 1); +by (resolve_tac [impI RS allI] 1); +by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1); +by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1); +by (fast_tac FOL_cs 1); +val le_trans = result(); + +goal Nat2.thy "i < j --> j <=k --> i < k"; +by (IND_TAC nat_ind (K all_tac) "j" 1); +by (simp_tac nat_ss 1); +by (resolve_tac [impI RS allI] 1); +by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1); +by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1); +by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1); +by (fast_tac FOL_cs 1); +val less_le_trans = result(); + +goal Nat2.thy "succ(i) <= j <-> i < j"; +by (IND_TAC nat_ind (simp_tac nat_ss) "j" 1); +by (resolve_tac [impI RS allI] 1); +by (ALL_IND_TAC nat_exh (asm_simp_tac nat_ss) 1); +val succ_le = result(); + +goal Nat2.thy "i i=j | i nat" + "0" :: "nat" ("0") + "+" :: "[nat,nat] => nat" (infixr 90) + "<","<=" :: "[nat,nat] => o" (infixr 70) + +rules + pred_0 "pred(0) = 0" + pred_succ "pred(succ(m)) = m" + + plus_0 "0+n = n" + plus_succ "succ(m)+n = succ(m+n)" + + nat_distinct1 "~ 0=succ(n)" + nat_distinct2 "~ succ(m)=0" + succ_inject "succ(m)=succ(n) <-> m=n" + + leq_0 "0 <= n" + leq_succ_succ "succ(m)<=succ(n) <-> m<=n" + leq_succ_0 "~ succ(m) <= 0" + + lt_0_succ "0 < succ(n)" + lt_succ_succ "succ(m) mP(succ(n)) |] ==> All(P)" +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/prolog.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/prolog.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,74 @@ +(* Title: FOL/ex/prolog.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For ex/prolog.thy. First-Order Logic: PROLOG examples +*) + +open Prolog; + +goal Prolog.thy "app(a:b:c:Nil, d:e:Nil, ?x)"; +by (resolve_tac [appNil,appCons] 1); +by (resolve_tac [appNil,appCons] 1); +by (resolve_tac [appNil,appCons] 1); +by (resolve_tac [appNil,appCons] 1); +prth (result()); + +goal Prolog.thy "app(?x, c:d:Nil, a:b:c:d:Nil)"; +by (REPEAT (resolve_tac [appNil,appCons] 1)); +result(); + + +goal Prolog.thy "app(?x, ?y, a:b:c:d:Nil)"; +by (REPEAT (resolve_tac [appNil,appCons] 1)); +back(); +back(); +back(); +back(); +result(); + + +(*app([x1,...,xn], y, ?z) requires (n+1) inferences*) +(*rev([x1,...,xn], ?y) requires (n+1)(n+2)/2 inferences*) + +goal Prolog.thy "rev(a:b:c:d:Nil, ?x)"; +val rules = [appNil,appCons,revNil,revCons]; +by (REPEAT (resolve_tac rules 1)); +result(); + +goal Prolog.thy "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:Nil, ?w)"; +by (REPEAT (resolve_tac rules 1)); +result(); + +goal Prolog.thy "rev(?x, a:b:c:Nil)"; +by (REPEAT (resolve_tac rules 1)); (*does not solve it directly!*) +back(); +back(); + +(*backtracking version*) +val prolog_tac = DEPTH_FIRST (has_fewer_prems 1) (resolve_tac rules 1); + +choplev 0; +by prolog_tac; +result(); + +goal Prolog.thy "rev(a:?x:c:?y:Nil, d:?z:b:?u)"; +by prolog_tac; +result(); + +(*rev([a..p], ?w) requires 153 inferences *) +goal Prolog.thy "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil, ?w)"; +by (DEPTH_SOLVE (resolve_tac ([refl,conjI]@rules) 1)); +(*Poly/ML: 4 secs >> 38 lips*) +result(); + +(*?x has 16, ?y has 32; rev(?y,?w) requires 561 (rather large) inferences; + total inferences = 2 + 1 + 17 + 561 = 581*) +goal Prolog.thy + "a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil = ?x & app(?x,?x,?y) & rev(?y,?w)"; +by (DEPTH_SOLVE (resolve_tac ([refl,conjI]@rules) 1)); +(*Poly/ML: 29 secs >> 20 lips*) +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/prolog.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/prolog.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,22 @@ +(* Title: FOL/ex/prolog.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +First-Order Logic: PROLOG examples + +Inherits from FOL the class term, the type o, and the coercion Trueprop +*) + +Prolog = FOL + +types list 1 +arities list :: (term)term +consts Nil :: "'a list" + ":" :: "['a, 'a list]=> 'a list" (infixr 60) + app :: "['a list, 'a list, 'a list] => o" + rev :: "['a list, 'a list] => o" +rules appNil "app(Nil,ys,ys)" + appCons "app(xs,ys,zs) ==> app(x:xs, ys, x:zs)" + revNil "rev(Nil,Nil)" + revCons "[| rev(xs,ys); app(ys, x:Nil, zs) |] ==> rev(x:xs, zs)" +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/prop.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/prop.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,153 @@ +(* Title: FOL/ex/prop + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +First-Order Logic: propositional examples (intuitionistic and classical) +Needs declarations of the theory "thy" and the tactic "tac" +*) + +writeln"File FOL/ex/prop."; + + +writeln"commutative laws of & and | "; +goal thy "P & Q --> Q & P"; +by tac; +result(); + +goal thy "P | Q --> Q | P"; +by tac; +result(); + + +writeln"associative laws of & and | "; +goal thy "(P & Q) & R --> P & (Q & R)"; +by tac; +result(); + +goal thy "(P | Q) | R --> P | (Q | R)"; +by tac; +result(); + + + +writeln"distributive laws of & and | "; +goal thy "(P & Q) | R --> (P | R) & (Q | R)"; +by tac; +result(); + +goal thy "(P | R) & (Q | R) --> (P & Q) | R"; +by tac; +result(); + +goal thy "(P | Q) & R --> (P & R) | (Q & R)"; +by tac; +result(); + + +goal thy "(P & R) | (Q & R) --> (P | Q) & R"; +by tac; +result(); + + +writeln"Laws involving implication"; + +goal thy "(P-->R) & (Q-->R) <-> (P|Q --> R)"; +by tac; +result(); + + +goal thy "(P & Q --> R) <-> (P--> (Q-->R))"; +by tac; +result(); + + +goal thy "((P-->R)-->R) --> ((Q-->R)-->R) --> (P&Q-->R) --> R"; +by tac; +result(); + +goal thy "~(P-->R) --> ~(Q-->R) --> ~(P&Q-->R)"; +by tac; +result(); + +goal thy "(P --> Q & R) <-> (P-->Q) & (P-->R)"; +by tac; +result(); + + +writeln"Propositions-as-types"; + +(*The combinator K*) +goal thy "P --> (Q --> P)"; +by tac; +result(); + +(*The combinator S*) +goal thy "(P-->Q-->R) --> (P-->Q) --> (P-->R)"; +by tac; +result(); + + +(*Converse is classical*) +goal thy "(P-->Q) | (P-->R) --> (P --> Q | R)"; +by tac; +result(); + +goal thy "(P-->Q) --> (~Q --> ~P)"; +by tac; +result(); + + +writeln"Schwichtenberg's examples (via T. Nipkow)"; + +(* stab-imp *) +goal thy "(((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"; +by tac; +result(); + +(* stab-to-peirce *) +goal thy "(((P --> R) --> R) --> P) --> (((Q --> R) --> R) --> Q) \ +\ --> ((P --> Q) --> P) --> P"; +by tac; +result(); + +(* peirce-imp1 *) +goal thy "(((Q --> R) --> Q) --> Q) \ +\ --> (((P --> Q) --> R) --> P --> Q) --> P --> Q"; +by tac; +result(); + +(* peirce-imp2 *) +goal thy "(((P --> R) --> P) --> P) --> ((P --> Q --> R) --> P) --> P"; +by tac; +result(); + +(* mints *) +goal thy "((((P --> Q) --> P) --> P) --> Q) --> Q"; +by tac; +result(); + +(* mints-solovev *) +goal thy "(P --> (Q --> R) --> Q) --> ((P --> Q) --> R) --> R"; +by tac; +result(); + +(* tatsuta *) +goal thy "(((P7 --> P1) --> P10) --> P4 --> P5) \ +\ --> (((P8 --> P2) --> P9) --> P3 --> P10) \ +\ --> (P1 --> P8) --> P6 --> P7 \ +\ --> (((P3 --> P2) --> P9) --> P4) \ +\ --> (P1 --> P3) --> (((P6 --> P1) --> P2) --> P9) --> P5"; +by tac; +result(); + +(* tatsuta1 *) +goal thy "(((P8 --> P2) --> P9) --> P3 --> P10) \ +\ --> (((P3 --> P2) --> P9) --> P4) \ +\ --> (((P6 --> P1) --> P2) --> P9) \ +\ --> (((P7 --> P1) --> P10) --> P4 --> P5) \ +\ --> (P1 --> P3) --> (P1 --> P8) --> P6 --> P7 --> P5"; +by tac; +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOL/ex/quant.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ex/quant.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,129 @@ +(* Title: FOL/ex/quant + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +First-Order Logic: quantifier examples (intuitionistic and classical) +Needs declarations of the theory "thy" and the tactic "tac" +*) + +writeln"File FOL/ex/quant."; + +goal thy "(ALL x y.P(x,y)) --> (ALL y x.P(x,y))"; +by tac; +result(); + + +goal thy "(EX x y.P(x,y)) --> (EX y x.P(x,y))"; +by tac; +result(); + + +(*Converse is false*) +goal thy "(ALL x.P(x)) | (ALL x.Q(x)) --> (ALL x. P(x) | Q(x))"; +by tac; +result(); + +goal thy "(ALL x. P-->Q(x)) <-> (P--> (ALL x.Q(x)))"; +by tac; +result(); + + +goal thy "(ALL x.P(x)-->Q) <-> ((EX x.P(x)) --> Q)"; +by tac; +result(); + + +writeln"Some harder ones"; + +goal thy "(EX x. P(x) | Q(x)) <-> (EX x.P(x)) | (EX x.Q(x))"; +by tac; +result(); +(*6 secs*) + +(*Converse is false*) +goal thy "(EX x. P(x)&Q(x)) --> (EX x.P(x)) & (EX x.Q(x))"; +by tac; +result(); + + +writeln"Basic test of quantifier reasoning"; +(*TRUE*) +goal thy "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"; +by tac; +result(); + + +goal thy "(ALL x. Q(x)) --> (EX x. Q(x))"; +by tac; +result(); + + +writeln"The following should fail, as they are false!"; + +goal thy "(ALL x. EX y. Q(x,y)) --> (EX y. ALL x. Q(x,y))"; +by tac handle ERROR => writeln"Failed, as expected"; +(*Check that subgoals remain: proof failed.*) +getgoal 1; + +goal thy "(EX x. Q(x)) --> (ALL x. Q(x))"; +by tac handle ERROR => writeln"Failed, as expected"; +getgoal 1; + +goal thy "P(?a) --> (ALL x.P(x))"; +by tac handle ERROR => writeln"Failed, as expected"; +(*Check that subgoals remain: proof failed.*) +getgoal 1; + +goal thy + "(P(?a) --> (ALL x.Q(x))) --> (ALL x. P(x) --> Q(x))"; +by tac handle ERROR => writeln"Failed, as expected"; +getgoal 1; + + +writeln"Back to things that are provable..."; + +goal thy "(ALL x.P(x)-->Q(x)) & (EX x.P(x)) --> (EX x.Q(x))"; +by tac; +result(); + + +(*An example of why exI should be delayed as long as possible*) +goal thy "(P --> (EX x.Q(x))) & P --> (EX x.Q(x))"; +by tac; +result(); + +goal thy "(ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"; +by tac; +(*Verify that no subgoals remain.*) +uresult(); + + +goal thy "(ALL x. Q(x)) --> (EX x. Q(x))"; +by tac; +result(); + + +writeln"Some slow ones"; + + +(*Principia Mathematica *11.53 *) +goal thy "(ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))"; +by tac; +result(); +(*6 secs*) + +(*Principia Mathematica *11.55 *) +goal thy "(EX x y. P(x) & Q(x,y)) <-> (EX x. P(x) & (EX y. Q(x,y)))"; +by tac; +result(); +(*9 secs*) + +(*Principia Mathematica *11.61 *) +goal thy "(EX y. ALL x. P(x) --> Q(x,y)) --> (ALL x. P(x) --> (EX y. Q(x,y)))"; +by tac; +result(); +(*3 secs*) + +writeln"Reached end of file."; + diff -r 000000000000 -r a5a9c433f639 src/FOL/fol.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/fol.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,94 @@ +(* Title: FOL/fol.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Tactics and lemmas for fol.thy (classical First-Order Logic) +*) + +open FOL; + +signature FOL_LEMMAS = + sig + val disjCI : thm + val excluded_middle : thm + val exCI : thm + val ex_classical : thm + val iffCE : thm + val impCE : thm + val notnotD : thm + val swap : thm + end; + + +structure FOL_Lemmas : FOL_LEMMAS = +struct + +(*** Classical introduction rules for | and EX ***) + +val disjCI = prove_goal FOL.thy + "(~Q ==> P) ==> P|Q" + (fn prems=> + [ (resolve_tac [classical] 1), + (REPEAT (ares_tac (prems@[disjI1,notI]) 1)), + (REPEAT (ares_tac (prems@[disjI2,notE]) 1)) ]); + +(*introduction rule involving only EX*) +val ex_classical = prove_goal FOL.thy + "( ~(EX x. P(x)) ==> P(a)) ==> EX x.P(x)" + (fn prems=> + [ (resolve_tac [classical] 1), + (eresolve_tac (prems RL [exI]) 1) ]); + +(*version of above, simplifying ~EX to ALL~ *) +val exCI = prove_goal FOL.thy + "(ALL x. ~P(x) ==> P(a)) ==> EX x.P(x)" + (fn [prem]=> + [ (resolve_tac [ex_classical] 1), + (resolve_tac [notI RS allI RS prem] 1), + (eresolve_tac [notE] 1), + (eresolve_tac [exI] 1) ]); + +val excluded_middle = prove_goal FOL.thy "~P | P" + (fn _=> [ rtac disjCI 1, assume_tac 1 ]); + + +(*** Special elimination rules *) + + +(*Classical implies (-->) elimination. *) +val impCE = prove_goal FOL.thy + "[| P-->Q; ~P ==> R; Q ==> R |] ==> R" + (fn major::prems=> + [ (resolve_tac [excluded_middle RS disjE] 1), + (DEPTH_SOLVE (ares_tac (prems@[major RS mp]) 1)) ]); + +(*Double negation law*) +val notnotD = prove_goal FOL.thy "~~P ==> P" + (fn [major]=> + [ (resolve_tac [classical] 1), (eresolve_tac [major RS notE] 1) ]); + + +(*** Tactics for implication and contradiction ***) + +(*Classical <-> elimination. Proof substitutes P=Q in + ~P ==> ~Q and P ==> Q *) +val iffCE = prove_goalw FOL.thy [iff_def] + "[| P<->Q; [| P; Q |] ==> R; [| ~P; ~Q |] ==> R |] ==> R" + (fn prems => + [ (resolve_tac [conjE] 1), + (REPEAT (DEPTH_SOLVE_1 + (etac impCE 1 ORELSE mp_tac 1 ORELSE ares_tac prems 1))) ]); + + +(*Should be used as swap since ~P becomes redundant*) +val swap = prove_goal FOL.thy + "~P ==> (~Q ==> P) ==> Q" + (fn major::prems=> + [ (resolve_tac [classical] 1), + (rtac (major RS notE) 1), + (REPEAT (ares_tac prems 1)) ]); + +end; + +open FOL_Lemmas; diff -r 000000000000 -r a5a9c433f639 src/FOL/fol.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/fol.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,4 @@ +FOL = IFOL + +rules +classical "(~P ==> P) ==> P" +end diff -r 000000000000 -r a5a9c433f639 src/FOL/ifol.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ifol.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,408 @@ +(* Title: FOL/ifol.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Tactics and lemmas for ifol.thy (intuitionistic first-order logic) +*) + +open IFOL; + +signature IFOL_LEMMAS = + sig + val allE: thm + val all_cong: thm + val all_dupE: thm + val all_impE: thm + val box_equals: thm + val conjE: thm + val conj_cong: thm + val conj_impE: thm + val contrapos: thm + val disj_cong: thm + val disj_impE: thm + val eq_cong: thm + val eq_mp_tac: int -> tactic + val ex1I: thm + val ex1E: thm + val ex1_equalsE: thm + val ex1_cong: thm + val ex_cong: thm + val ex_impE: thm + val iffD1: thm + val iffD2: thm + val iffE: thm + val iffI: thm + val iff_cong: thm + val iff_impE: thm + val iff_refl: thm + val iff_sym: thm + val iff_trans: thm + val impE: thm + val imp_cong: thm + val imp_impE: thm + val mp_tac: int -> tactic + val notE: thm + val notI: thm + val not_cong: thm + val not_impE: thm + val not_sym: thm + val not_to_imp: thm + val pred1_cong: thm + val pred2_cong: thm + val pred3_cong: thm + val pred_congs: thm list + val rev_mp: thm + val simp_equals: thm + val ssubst: thm + val subst_context: thm + val subst_context2: thm + val subst_context3: thm + val sym: thm + val trans: thm + val TrueI: thm + end; + + +structure IFOL_Lemmas : IFOL_LEMMAS = +struct + +val TrueI = prove_goalw IFOL.thy [True_def] "True" + (fn _ => [ (REPEAT (ares_tac [impI] 1)) ]); + +(*** Sequent-style elimination rules for & --> and ALL ***) + +val conjE = prove_goal IFOL.thy + "[| P&Q; [| P; Q |] ==> R |] ==> R" + (fn prems=> + [ (REPEAT (resolve_tac prems 1 + ORELSE (resolve_tac [conjunct1, conjunct2] 1 THEN + resolve_tac prems 1))) ]); + +val impE = prove_goal IFOL.thy + "[| P-->Q; P; Q ==> R |] ==> R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); + +val allE = prove_goal IFOL.thy + "[| ALL x.P(x); P(x) ==> R |] ==> R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[spec]) 1)) ]); + +(*Duplicates the quantifier; for use with eresolve_tac*) +val all_dupE = prove_goal IFOL.thy + "[| ALL x.P(x); [| P(x); ALL x.P(x) |] ==> R \ +\ |] ==> R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[spec]) 1)) ]); + + +(*** Negation rules, which translate between ~P and P-->False ***) + +val notI = prove_goalw IFOL.thy [not_def] "(P ==> False) ==> ~P" + (fn prems=> [ (REPEAT (ares_tac (prems@[impI]) 1)) ]); + +val notE = prove_goalw IFOL.thy [not_def] "[| ~P; P |] ==> R" + (fn prems=> + [ (resolve_tac [mp RS FalseE] 1), + (REPEAT (resolve_tac prems 1)) ]); + +(*This is useful with the special implication rules for each kind of P. *) +val not_to_imp = prove_goal IFOL.thy + "[| ~P; (P-->False) ==> Q |] ==> Q" + (fn prems=> [ (REPEAT (ares_tac (prems@[impI,notE]) 1)) ]); + + +(* For substitution int an assumption P, reduce Q to P-->Q, substitute into + this implication, then apply impI to move P back into the assumptions. + To specify P use something like + eres_inst_tac [ ("P","ALL y. ?S(x,y)") ] rev_mp 1 *) +val rev_mp = prove_goal IFOL.thy "[| P; P --> Q |] ==> Q" + (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); + + +(*Contrapositive of an inference rule*) +val contrapos = prove_goal IFOL.thy "[| ~Q; P==>Q |] ==> ~P" + (fn [major,minor]=> + [ (rtac (major RS notE RS notI) 1), + (etac minor 1) ]); + + +(*** Modus Ponens Tactics ***) + +(*Finds P-->Q and P in the assumptions, replaces implication by Q *) +fun mp_tac i = eresolve_tac [notE,impE] i THEN assume_tac i; + +(*Like mp_tac but instantiates no variables*) +fun eq_mp_tac i = eresolve_tac [notE,impE] i THEN eq_assume_tac i; + + +(*** If-and-only-if ***) + +val iffI = prove_goalw IFOL.thy [iff_def] + "[| P ==> Q; Q ==> P |] ==> P<->Q" + (fn prems=> [ (REPEAT (ares_tac (prems@[conjI, impI]) 1)) ]); + + +(*Observe use of rewrite_rule to unfold "<->" in meta-assumptions (prems) *) +val iffE = prove_goalw IFOL.thy [iff_def] + "[| P <-> Q; [| P-->Q; Q-->P |] ==> R |] ==> R" + (fn prems => [ (resolve_tac [conjE] 1), (REPEAT (ares_tac prems 1)) ]); + +(* Destruct rules for <-> similar to Modus Ponens *) + +val iffD1 = prove_goalw IFOL.thy [iff_def] "[| P <-> Q; P |] ==> Q" + (fn prems => [ (rtac (conjunct1 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); + +val iffD2 = prove_goalw IFOL.thy [iff_def] "[| P <-> Q; Q |] ==> P" + (fn prems => [ (rtac (conjunct2 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); + +val iff_refl = prove_goal IFOL.thy "P <-> P" + (fn _ => [ (REPEAT (ares_tac [iffI] 1)) ]); + +val iff_sym = prove_goal IFOL.thy "Q <-> P ==> P <-> Q" + (fn [major] => + [ (rtac (major RS iffE) 1), + (rtac iffI 1), + (REPEAT (eresolve_tac [asm_rl,mp] 1)) ]); + +val iff_trans = prove_goal IFOL.thy + "!!P Q R. [| P <-> Q; Q<-> R |] ==> P <-> R" + (fn _ => + [ (rtac iffI 1), + (REPEAT (eresolve_tac [asm_rl,iffE] 1 ORELSE mp_tac 1)) ]); + + +(*** Unique existence. NOTE THAT the following 2 quantifications + EX!x such that [EX!y such that P(x,y)] (sequential) + EX!x,y such that P(x,y) (simultaneous) + do NOT mean the same thing. The parser treats EX!x y.P(x,y) as sequential. +***) + +val ex1I = prove_goalw IFOL.thy [ex1_def] + "[| P(a); !!x. P(x) ==> x=a |] ==> EX! x. P(x)" + (fn prems => [ (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ]); + +val ex1E = prove_goalw IFOL.thy [ex1_def] + "[| EX! x.P(x); !!x. [| P(x); ALL y. P(y) --> y=x |] ==> R |] ==> R" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (eresolve_tac [exE,conjE] 1 ORELSE ares_tac prems 1)) ]); + + +(*** <-> congruence rules for simplification ***) + +(*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*) +fun iff_tac prems i = + resolve_tac (prems RL [iffE]) i THEN + REPEAT1 (eresolve_tac [asm_rl,mp] i); + +val conj_cong = prove_goal IFOL.thy + "[| P <-> P'; P' ==> Q <-> Q' |] ==> (P&Q) <-> (P'&Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [iffI,conjI] 1 + ORELSE eresolve_tac [iffE,conjE,mp] 1 + ORELSE iff_tac prems 1)) ]); + +val disj_cong = prove_goal IFOL.thy + "[| P <-> P'; Q <-> Q' |] ==> (P|Q) <-> (P'|Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (eresolve_tac [iffE,disjE,disjI1,disjI2] 1 + ORELSE ares_tac [iffI] 1 + ORELSE mp_tac 1)) ]); + +val imp_cong = prove_goal IFOL.thy + "[| P <-> P'; P' ==> Q <-> Q' |] ==> (P-->Q) <-> (P'-->Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [iffI,impI] 1 + ORELSE eresolve_tac [iffE] 1 + ORELSE mp_tac 1 ORELSE iff_tac prems 1)) ]); + +val iff_cong = prove_goal IFOL.thy + "[| P <-> P'; Q <-> Q' |] ==> (P<->Q) <-> (P'<->Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (eresolve_tac [iffE] 1 + ORELSE ares_tac [iffI] 1 + ORELSE mp_tac 1)) ]); + +val not_cong = prove_goal IFOL.thy + "P <-> P' ==> ~P <-> ~P'" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [iffI,notI] 1 + ORELSE mp_tac 1 + ORELSE eresolve_tac [iffE,notE] 1)) ]); + +val all_cong = prove_goal IFOL.thy + "(!!x.P(x) <-> Q(x)) ==> (ALL x.P(x)) <-> (ALL x.Q(x))" + (fn prems => + [ (REPEAT (ares_tac [iffI,allI] 1 + ORELSE mp_tac 1 + ORELSE eresolve_tac [allE] 1 ORELSE iff_tac prems 1)) ]); + +val ex_cong = prove_goal IFOL.thy + "(!!x.P(x) <-> Q(x)) ==> (EX x.P(x)) <-> (EX x.Q(x))" + (fn prems => + [ (REPEAT (eresolve_tac [exE] 1 ORELSE ares_tac [iffI,exI] 1 + ORELSE mp_tac 1 + ORELSE iff_tac prems 1)) ]); + +val ex1_cong = prove_goal IFOL.thy + "(!!x.P(x) <-> Q(x)) ==> (EX! x.P(x)) <-> (EX! x.Q(x))" + (fn prems => + [ (REPEAT (eresolve_tac [ex1E, spec RS mp] 1 ORELSE ares_tac [iffI,ex1I] 1 + ORELSE mp_tac 1 + ORELSE iff_tac prems 1)) ]); + +(*** Equality rules ***) + +val sym = prove_goal IFOL.thy "a=b ==> b=a" + (fn [major] => [ (rtac (major RS subst) 1), (rtac refl 1) ]); + +val trans = prove_goal IFOL.thy "[| a=b; b=c |] ==> a=c" + (fn [prem1,prem2] => [ (rtac (prem2 RS subst) 1), (rtac prem1 1) ]); + +(** ~ b=a ==> ~ a=b **) +val [not_sym] = compose(sym,2,contrapos); + +(*calling "standard" reduces maxidx to 0*) +val ssubst = standard (sym RS subst); + +(*A special case of ex1E that would otherwise need quantifier expansion*) +val ex1_equalsE = prove_goal IFOL.thy + "[| EX! x.P(x); P(a); P(b) |] ==> a=b" + (fn prems => + [ (cut_facts_tac prems 1), + (etac ex1E 1), + (rtac trans 1), + (rtac sym 2), + (REPEAT (eresolve_tac [asm_rl, spec RS mp] 1)) ]); + +(** Polymorphic congruence rules **) + +val subst_context = prove_goal IFOL.thy + "[| a=b |] ==> t(a)=t(b)" + (fn prems=> + [ (resolve_tac (prems RL [ssubst]) 1), + (resolve_tac [refl] 1) ]); + +val subst_context2 = prove_goal IFOL.thy + "[| a=b; c=d |] ==> t(a,c)=t(b,d)" + (fn prems=> + [ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); + +val subst_context3 = prove_goal IFOL.thy + "[| a=b; c=d; e=f |] ==> t(a,c,e)=t(b,d,f)" + (fn prems=> + [ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); + +(*Useful with eresolve_tac for proving equalties from known equalities. + a = b + | | + c = d *) +val box_equals = prove_goal IFOL.thy + "[| a=b; a=c; b=d |] ==> c=d" + (fn prems=> + [ (resolve_tac [trans] 1), + (resolve_tac [trans] 1), + (resolve_tac [sym] 1), + (REPEAT (resolve_tac prems 1)) ]); + +(*Dual of box_equals: for proving equalities backwards*) +val simp_equals = prove_goal IFOL.thy + "[| a=c; b=d; c=d |] ==> a=b" + (fn prems=> + [ (resolve_tac [trans] 1), + (resolve_tac [trans] 1), + (REPEAT (resolve_tac (prems @ (prems RL [sym])) 1)) ]); + +(** Congruence rules for predicate letters **) + +val pred1_cong = prove_goal IFOL.thy + "a=a' ==> P(a) <-> P(a')" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); + +val pred2_cong = prove_goal IFOL.thy + "[| a=a'; b=b' |] ==> P(a,b) <-> P(a',b')" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); + +val pred3_cong = prove_goal IFOL.thy + "[| a=a'; b=b'; c=c' |] ==> P(a,b,c) <-> P(a',b',c')" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); + +(*special cases for free variables P, Q, R, S -- up to 3 arguments*) + +val pred_congs = + flat (map (fn c => + map (fn th => read_instantiate [("P",c)] th) + [pred1_cong,pred2_cong,pred3_cong]) + (explode"PQRS")); + +(*special case for the equality predicate!*) +val eq_cong = read_instantiate [("P","op =")] pred2_cong; + + +(*** Simplifications of assumed implications. + Roy Dyckhoff has proved that conj_impE, disj_impE, and imp_impE + used with mp_tac (restricted to atomic formulae) is COMPLETE for + intuitionistic propositional logic. See + R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic + (preprint, University of St Andrews, 1991) ***) + +val conj_impE = prove_goal IFOL.thy + "[| (P&Q)-->S; P-->(Q-->S) ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (ares_tac ([conjI, impI, major RS mp]@prems) 1)) ]); + +val disj_impE = prove_goal IFOL.thy + "[| (P|Q)-->S; [| P-->S; Q-->S |] ==> R |] ==> R" + (fn major::prems=> + [ (DEPTH_SOLVE (ares_tac ([disjI1, disjI2, impI, major RS mp]@prems) 1)) ]); + +(*Simplifies the implication. Classical version is stronger. + Still UNSAFE since Q must be provable -- backtracking needed. *) +val imp_impE = prove_goal IFOL.thy + "[| (P-->Q)-->S; [| P; Q-->S |] ==> Q; S ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (ares_tac ([impI, major RS mp]@prems) 1)) ]); + +(*Simplifies the implication. Classical version is stronger. + Still UNSAFE since ~P must be provable -- backtracking needed. *) +val not_impE = prove_goal IFOL.thy + "[| ~P --> S; P ==> False; S ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (ares_tac ([notI, impI, major RS mp]@prems) 1)) ]); + +(*Simplifies the implication. UNSAFE. *) +val iff_impE = prove_goal IFOL.thy + "[| (P<->Q)-->S; [| P; Q-->S |] ==> Q; [| Q; P-->S |] ==> P; \ +\ S ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (ares_tac ([iffI, impI, major RS mp]@prems) 1)) ]); + +(*What if (ALL x.~~P(x)) --> ~~(ALL x.P(x)) is an assumption? UNSAFE*) +val all_impE = prove_goal IFOL.thy + "[| (ALL x.P(x))-->S; !!x.P(x); S ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (ares_tac ([allI, impI, major RS mp]@prems) 1)) ]); + +(*Unsafe: (EX x.P(x))-->S is equivalent to ALL x.P(x)-->S. *) +val ex_impE = prove_goal IFOL.thy + "[| (EX x.P(x))-->S; P(x)-->S ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (ares_tac ([exI, impI, major RS mp]@prems) 1)) ]); + +end; + +open IFOL_Lemmas; + diff -r 000000000000 -r a5a9c433f639 src/FOL/ifol.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/ifol.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,70 @@ +IFOL = Pure + + +classes term < logic + +default term + +types o 0 + +arities o :: logic + +consts + Trueprop :: "o => prop" ("(_)" [0] 5) + True,False :: "o" + (*Connectives*) + "=" :: "['a,'a] => o" (infixl 50) + Not :: "o => o" ("~ _" [40] 40) + "&" :: "[o,o] => o" (infixr 35) + "|" :: "[o,o] => o" (infixr 30) + "-->" :: "[o,o] => o" (infixr 25) + "<->" :: "[o,o] => o" (infixr 25) + (*Quantifiers*) + All :: "('a => o) => o" (binder "ALL " 10) + Ex :: "('a => o) => o" (binder "EX " 10) + Ex1 :: "('a => o) => o" (binder "EX! " 10) + +rules + + (*Equality*) + +refl "a=a" +subst "[| a=b; P(a) |] ==> P(b)" + + (*Propositional logic*) + +conjI "[| P; Q |] ==> P&Q" +conjunct1 "P&Q ==> P" +conjunct2 "P&Q ==> Q" + +disjI1 "P ==> P|Q" +disjI2 "Q ==> P|Q" +disjE "[| P|Q; P ==> R; Q ==> R |] ==> R" + +impI "(P ==> Q) ==> P-->Q" +mp "[| P-->Q; P |] ==> Q" + +FalseE "False ==> P" + + (*Definitions*) + +True_def "True == False-->False" +not_def "~P == P-->False" +iff_def "P<->Q == (P-->Q) & (Q-->P)" + + (*Unique existence*) +ex1_def "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)" + + (*Quantifiers*) + +allI "(!!x. P(x)) ==> (ALL x.P(x))" +spec "(ALL x.P(x)) ==> P(x)" + +exI "P(x) ==> (EX x.P(x))" +exE "[| EX x.P(x); !!x. P(x) ==> R |] ==> R" + + (* Reflection *) + +eq_reflection "(x=y) ==> (x==y)" +iff_reflection "(P<->Q) ==> (P==Q)" + +end diff -r 000000000000 -r a5a9c433f639 src/FOL/int-prover.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/int-prover.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,81 @@ +(* Title: FOL/int-prover + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +A naive prover for intuitionistic logic + +BEWARE OF NAME CLASHES WITH CLASSICAL TACTICS -- use Int.fast_tac ... + +Completeness (for propositional logic) is proved in + +Roy Dyckhoff. +Contraction-Free Sequent Calculi for Intuitionistic Logic. +J. Symbolic Logic (in press) +*) + +signature INT_PROVER = + sig + val best_tac: int -> tactic + val fast_tac: int -> tactic + val inst_step_tac: int -> tactic + val safe_step_tac: int -> tactic + val safe_brls: (bool * thm) list + val safe_tac: tactic + val step_tac: int -> tactic + val haz_brls: (bool * thm) list + end; + + +structure Int : INT_PROVER = +struct + +(*Negation is treated as a primitive symbol, with rules notI (introduction), + not_to_imp (converts the assumption ~P to P-->False), and not_impE + (handles double negations). Could instead rewrite by not_def as the first + step of an intuitionistic proof. +*) +val safe_brls = sort lessb + [ (true,FalseE), (false,TrueI), (false,refl), + (false,impI), (false,notI), (false,allI), + (true,conjE), (true,exE), + (false,conjI), (true,conj_impE), + (true,disj_impE), (true,ex_impE), + (true,disjE), (false,iffI), (true,iffE), (true,not_to_imp) ]; + +val haz_brls = + [ (false,disjI1), (false,disjI2), (false,exI), + (true,allE), (true,not_impE), (true,imp_impE), (true,iff_impE), + (true,all_impE), (true,impE) ]; + +(*0 subgoals vs 1 or more: the p in safep is for positive*) +val (safe0_brls, safep_brls) = + partition (apl(0,op=) o subgoals_of_brl) safe_brls; + +(*Attack subgoals using safe inferences -- matching, not resolution*) +val safe_step_tac = FIRST' [eq_assume_tac, + eq_mp_tac, + bimatch_tac safe0_brls, + hyp_subst_tac, + bimatch_tac safep_brls] ; + +(*Repeatedly attack subgoals using safe inferences -- it's deterministic!*) +val safe_tac = DETERM (REPEAT_FIRST safe_step_tac); + +(*These steps could instantiate variables and are therefore unsafe.*) +val inst_step_tac = + assume_tac APPEND' mp_tac APPEND' + biresolve_tac (safe0_brls @ safep_brls); + +(*One safe or unsafe step. *) +fun step_tac i = FIRST [safe_tac, inst_step_tac i, biresolve_tac haz_brls i]; + +(*Dumb but fast*) +val fast_tac = SELECT_GOAL (DEPTH_SOLVE (step_tac 1)); + +(*Slower but smarter than fast_tac*) +val best_tac = + SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) (step_tac 1)); + +end; + diff -r 000000000000 -r a5a9c433f639 src/FOL/intprover.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/intprover.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,81 @@ +(* Title: FOL/int-prover + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +A naive prover for intuitionistic logic + +BEWARE OF NAME CLASHES WITH CLASSICAL TACTICS -- use Int.fast_tac ... + +Completeness (for propositional logic) is proved in + +Roy Dyckhoff. +Contraction-Free Sequent Calculi for Intuitionistic Logic. +J. Symbolic Logic (in press) +*) + +signature INT_PROVER = + sig + val best_tac: int -> tactic + val fast_tac: int -> tactic + val inst_step_tac: int -> tactic + val safe_step_tac: int -> tactic + val safe_brls: (bool * thm) list + val safe_tac: tactic + val step_tac: int -> tactic + val haz_brls: (bool * thm) list + end; + + +structure Int : INT_PROVER = +struct + +(*Negation is treated as a primitive symbol, with rules notI (introduction), + not_to_imp (converts the assumption ~P to P-->False), and not_impE + (handles double negations). Could instead rewrite by not_def as the first + step of an intuitionistic proof. +*) +val safe_brls = sort lessb + [ (true,FalseE), (false,TrueI), (false,refl), + (false,impI), (false,notI), (false,allI), + (true,conjE), (true,exE), + (false,conjI), (true,conj_impE), + (true,disj_impE), (true,ex_impE), + (true,disjE), (false,iffI), (true,iffE), (true,not_to_imp) ]; + +val haz_brls = + [ (false,disjI1), (false,disjI2), (false,exI), + (true,allE), (true,not_impE), (true,imp_impE), (true,iff_impE), + (true,all_impE), (true,impE) ]; + +(*0 subgoals vs 1 or more: the p in safep is for positive*) +val (safe0_brls, safep_brls) = + partition (apl(0,op=) o subgoals_of_brl) safe_brls; + +(*Attack subgoals using safe inferences -- matching, not resolution*) +val safe_step_tac = FIRST' [eq_assume_tac, + eq_mp_tac, + bimatch_tac safe0_brls, + hyp_subst_tac, + bimatch_tac safep_brls] ; + +(*Repeatedly attack subgoals using safe inferences -- it's deterministic!*) +val safe_tac = DETERM (REPEAT_FIRST safe_step_tac); + +(*These steps could instantiate variables and are therefore unsafe.*) +val inst_step_tac = + assume_tac APPEND' mp_tac APPEND' + biresolve_tac (safe0_brls @ safep_brls); + +(*One safe or unsafe step. *) +fun step_tac i = FIRST [safe_tac, inst_step_tac i, biresolve_tac haz_brls i]; + +(*Dumb but fast*) +val fast_tac = SELECT_GOAL (DEPTH_SOLVE (step_tac 1)); + +(*Slower but smarter than fast_tac*) +val best_tac = + SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) (step_tac 1)); + +end; + diff -r 000000000000 -r a5a9c433f639 src/FOL/simpdata.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOL/simpdata.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,114 @@ +(* Title: FOL/simpdata + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Simplification data for FOL +*) + +(*** Rewrite rules ***) + +fun int_prove_fun s = + (writeln s; prove_goal IFOL.thy s + (fn prems => [ (cut_facts_tac prems 1), (Int.fast_tac 1) ])); + +val conj_rews = map int_prove_fun + ["P & True <-> P", "True & P <-> P", + "P & False <-> False", "False & P <-> False", + "P & P <-> P", + "P & ~P <-> False", "~P & P <-> False", + "(P & Q) & R <-> P & (Q & R)"]; + +val disj_rews = map int_prove_fun + ["P | True <-> True", "True | P <-> True", + "P | False <-> P", "False | P <-> P", + "P | P <-> P", + "(P | Q) | R <-> P | (Q | R)"]; + +val not_rews = map int_prove_fun + ["~ False <-> True", "~ True <-> False"]; + +val imp_rews = map int_prove_fun + ["(P --> False) <-> ~P", "(P --> True) <-> True", + "(False --> P) <-> True", "(True --> P) <-> P", + "(P --> P) <-> True", "(P --> ~P) <-> ~P"]; + +val iff_rews = map int_prove_fun + ["(True <-> P) <-> P", "(P <-> True) <-> P", + "(P <-> P) <-> True", + "(False <-> P) <-> ~P", "(P <-> False) <-> ~P"]; + +val quant_rews = map int_prove_fun + ["(ALL x.P) <-> P", "(EX x.P) <-> P"]; + +(*These are NOT supplied by default!*) +val distrib_rews = map int_prove_fun + ["~(P|Q) <-> ~P & ~Q", + "P & (Q | R) <-> P&Q | P&R", "(Q | R) & P <-> Q&P | R&P", + "(P | Q --> R) <-> (P --> R) & (Q --> R)"]; + +val P_Imp_P_iff_T = int_prove_fun "P ==> (P <-> True)"; + +fun make_iff_T th = th RS P_Imp_P_iff_T; + +val refl_iff_T = make_iff_T refl; + +val notFalseI = int_prove_fun "~False"; + +(* Conversion into rewrite rules *) + +val not_P_imp_P_iff_F = int_prove_fun "~P ==> (P <-> False)"; + +fun mk_meta_eq th = case concl_of th of + _ $ (Const("op <->",_)$_$_) => th RS iff_reflection + | _ $ (Const("op =",_)$_$_) => th RS eq_reflection + | _ $ (Const("Not",_)$_) => (th RS not_P_imp_P_iff_F) RS iff_reflection + | _ => (make_iff_T th) RS iff_reflection; + +fun atomize th = case concl_of th of (*The operator below is Trueprop*) + _ $ (Const("op -->",_) $ _ $ _) => atomize(th RS mp) + | _ $ (Const("op &",_) $ _ $ _) => atomize(th RS conjunct1) @ + atomize(th RS conjunct2) + | _ $ (Const("All",_) $ _) => atomize(th RS spec) + | _ $ (Const("True",_)) => [] + | _ $ (Const("False",_)) => [] + | _ => [th]; + +fun mk_rew_rules th = map mk_meta_eq (atomize th); + +structure Induction = InductionFun(struct val spec=IFOL.spec end); + +val IFOL_rews = + [refl_iff_T] @ conj_rews @ disj_rews @ not_rews @ + imp_rews @ iff_rews @ quant_rews; + +open Simplifier Induction; + +val IFOL_ss = empty_ss + setmksimps mk_rew_rules + setsolver + (fn prems => resolve_tac (TrueI::refl::iff_refl::notFalseI::prems)) + setsubgoaler asm_simp_tac + addsimps IFOL_rews + addcongs [imp_cong RS iff_reflection]; + +(*Classical version...*) +fun prove_fun s = + (writeln s; prove_goal FOL.thy s + (fn prems => [ (cut_facts_tac prems 1), (Cla.fast_tac FOL_cs 1) ])); + +val cla_rews = map prove_fun + ["P | ~P", "~P | P", + "~ ~ P <-> P", "(~P --> P) <-> P"]; + +val FOL_ss = IFOL_ss addsimps cla_rews; + +(*** case splitting ***) + +val split_tac = + let val eq_reflection2 = prove_goal FOL.thy "x==y ==> x=y" + (fn [prem] => [rewtac prem, rtac refl 1]) + val iff_reflection2 = prove_goal FOL.thy "x==y ==> x<->y" + (fn [prem] => [rewtac prem, rtac iff_refl 1]) + val [iffD] = [eq_reflection2,iff_reflection2] RL [iffD2] + in fn splits => mk_case_split_tac iffD (map mk_meta_eq splits) end; diff -r 000000000000 -r a5a9c433f639 src/FOLP/FOLP.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/FOLP.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,95 @@ +(* Title: FOL/fol.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Tactics and lemmas for fol.thy (classical First-Order Logic) +*) + +open FOLP; + +signature FOLP_LEMMAS = + sig + val disjCI : thm + val excluded_middle : thm + val exCI : thm + val ex_classical : thm + val iffCE : thm + val impCE : thm + val notnotD : thm + val swap : thm + end; + + +structure FOLP_Lemmas : FOLP_LEMMAS = +struct + +(*** Classical introduction rules for | and EX ***) + +val disjCI = prove_goal FOLP.thy + "(!!x.x:~Q ==> f(x):P) ==> ?p : P|Q" + (fn prems=> + [ (resolve_tac [classical] 1), + (REPEAT (ares_tac (prems@[disjI1,notI]) 1)), + (REPEAT (ares_tac (prems@[disjI2,notE]) 1)) ]); + +(*introduction rule involving only EX*) +val ex_classical = prove_goal FOLP.thy + "( !!u.u:~(EX x. P(x)) ==> f(u):P(a)) ==> ?p : EX x.P(x)" + (fn prems=> + [ (resolve_tac [classical] 1), + (eresolve_tac (prems RL [exI]) 1) ]); + +(*version of above, simplifying ~EX to ALL~ *) +val exCI = prove_goal FOLP.thy + "(!!u.u:ALL x. ~P(x) ==> f(u):P(a)) ==> ?p : EX x.P(x)" + (fn [prem]=> + [ (resolve_tac [ex_classical] 1), + (resolve_tac [notI RS allI RS prem] 1), + (eresolve_tac [notE] 1), + (eresolve_tac [exI] 1) ]); + +val excluded_middle = prove_goal FOLP.thy "?p : ~P | P" + (fn _=> [ rtac disjCI 1, assume_tac 1 ]); + + +(*** Special elimination rules *) + + +(*Classical implies (-->) elimination. *) +val impCE = prove_goal FOLP.thy + "[| p:P-->Q; !!x.x:~P ==> f(x):R; !!y.y:Q ==> g(y):R |] ==> ?p : R" + (fn major::prems=> + [ (resolve_tac [excluded_middle RS disjE] 1), + (DEPTH_SOLVE (ares_tac (prems@[major RS mp]) 1)) ]); + +(*Double negation law*) +val notnotD = prove_goal FOLP.thy "p:~~P ==> ?p : P" + (fn [major]=> + [ (resolve_tac [classical] 1), (eresolve_tac [major RS notE] 1) ]); + + +(*** Tactics for implication and contradiction ***) + +(*Classical <-> elimination. Proof substitutes P=Q in + ~P ==> ~Q and P ==> Q *) +val iffCE = prove_goalw FOLP.thy [iff_def] + "[| p:P<->Q; !!x y.[| x:P; y:Q |] ==> f(x,y):R; \ +\ !!x y.[| x:~P; y:~Q |] ==> g(x,y):R |] ==> ?p : R" + (fn prems => + [ (resolve_tac [conjE] 1), + (REPEAT (DEPTH_SOLVE_1 + (etac impCE 1 ORELSE mp_tac 1 ORELSE ares_tac prems 1))) ]); + + +(*Should be used as swap since ~P becomes redundant*) +val swap = prove_goal FOLP.thy + "p:~P ==> (!!x.x:~Q ==> f(x):P) ==> ?p : Q" + (fn major::prems=> + [ (resolve_tac [classical] 1), + (rtac (major RS notE) 1), + (REPEAT (ares_tac prems 1)) ]); + +end; + +open FOLP_Lemmas; diff -r 000000000000 -r a5a9c433f639 src/FOLP/FOLP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/FOLP.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,6 @@ +FOLP = IFOLP + +consts + cla :: "[p=>p]=>p" +rules + classical "(!!x.x:~P ==> f(x):P) ==> cla(f):P" +end diff -r 000000000000 -r a5a9c433f639 src/FOLP/IFOLP.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/IFOLP.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,444 @@ +(* Title: FOLP/ifol.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Tactics and lemmas for ifol.thy (intuitionistic first-order logic) +*) + +open IFOLP; + +signature IFOLP_LEMMAS = + sig + val allE: thm + val all_cong: thm + val all_dupE: thm + val all_impE: thm + val box_equals: thm + val conjE: thm + val conj_cong: thm + val conj_impE: thm + val contrapos: thm + val disj_cong: thm + val disj_impE: thm + val eq_cong: thm + val ex1I: thm + val ex1E: thm + val ex1_equalsE: thm +(* val ex1_cong: thm****) + val ex_cong: thm + val ex_impE: thm + val iffD1: thm + val iffD2: thm + val iffE: thm + val iffI: thm + val iff_cong: thm + val iff_impE: thm + val iff_refl: thm + val iff_sym: thm + val iff_trans: thm + val impE: thm + val imp_cong: thm + val imp_impE: thm + val mp_tac: int -> tactic + val notE: thm + val notI: thm + val not_cong: thm + val not_impE: thm + val not_sym: thm + val not_to_imp: thm + val pred1_cong: thm + val pred2_cong: thm + val pred3_cong: thm + val pred_congs: thm list + val refl: thm + val rev_mp: thm + val simp_equals: thm + val subst: thm + val ssubst: thm + val subst_context: thm + val subst_context2: thm + val subst_context3: thm + val sym: thm + val trans: thm + val TrueI: thm + val uniq_assume_tac: int -> tactic + val uniq_mp_tac: int -> tactic + end; + + +structure IFOLP_Lemmas : IFOLP_LEMMAS = +struct + +val TrueI = TrueI; + +(*** Sequent-style elimination rules for & --> and ALL ***) + +val conjE = prove_goal IFOLP.thy + "[| p:P&Q; !!x y.[| x:P; y:Q |] ==> f(x,y):R |] ==> ?a:R" + (fn prems=> + [ (REPEAT (resolve_tac prems 1 + ORELSE (resolve_tac [conjunct1, conjunct2] 1 THEN + resolve_tac prems 1))) ]); + +val impE = prove_goal IFOLP.thy + "[| p:P-->Q; q:P; !!x.x:Q ==> r(x):R |] ==> ?p:R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); + +val allE = prove_goal IFOLP.thy + "[| p:ALL x.P(x); !!y.y:P(x) ==> q(y):R |] ==> ?p:R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[spec]) 1)) ]); + +(*Duplicates the quantifier; for use with eresolve_tac*) +val all_dupE = prove_goal IFOLP.thy + "[| p:ALL x.P(x); !!y z.[| y:P(x); z:ALL x.P(x) |] ==> q(y,z):R \ +\ |] ==> ?p:R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[spec]) 1)) ]); + + +(*** Negation rules, which translate between ~P and P-->False ***) + +val notI = prove_goalw IFOLP.thy [not_def] "(!!x.x:P ==> q(x):False) ==> ?p:~P" + (fn prems=> [ (REPEAT (ares_tac (prems@[impI]) 1)) ]); + +val notE = prove_goalw IFOLP.thy [not_def] "[| p:~P; q:P |] ==> ?p:R" + (fn prems=> + [ (resolve_tac [mp RS FalseE] 1), + (REPEAT (resolve_tac prems 1)) ]); + +(*This is useful with the special implication rules for each kind of P. *) +val not_to_imp = prove_goal IFOLP.thy + "[| p:~P; !!x.x:(P-->False) ==> q(x):Q |] ==> ?p:Q" + (fn prems=> [ (REPEAT (ares_tac (prems@[impI,notE]) 1)) ]); + + +(* For substitution int an assumption P, reduce Q to P-->Q, substitute into + this implication, then apply impI to move P back into the assumptions. + To specify P use something like + eres_inst_tac [ ("P","ALL y. ?S(x,y)") ] rev_mp 1 *) +val rev_mp = prove_goal IFOLP.thy "[| p:P; q:P --> Q |] ==> ?p:Q" + (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); + + +(*Contrapositive of an inference rule*) +val contrapos = prove_goal IFOLP.thy "[| p:~Q; !!y.y:P==>q(y):Q |] ==> ?a:~P" + (fn [major,minor]=> + [ (rtac (major RS notE RS notI) 1), + (etac minor 1) ]); + +(** Unique assumption tactic. + Ignores proof objects. + Fails unless one assumption is equal and exactly one is unifiable +**) + +local + fun discard_proof (Const("Proof",_) $ P $ _) = P; +in +val uniq_assume_tac = + SUBGOAL + (fn (prem,i) => + let val hyps = map discard_proof (Logic.strip_assums_hyp prem) + and concl = discard_proof (Logic.strip_assums_concl prem) + in + if exists (fn hyp => hyp aconv concl) hyps + then case distinct (filter (fn hyp=> could_unify(hyp,concl)) hyps) of + [_] => assume_tac i + | _ => no_tac + else no_tac + end); +end; + + +(*** Modus Ponens Tactics ***) + +(*Finds P-->Q and P in the assumptions, replaces implication by Q *) +fun mp_tac i = eresolve_tac [notE,make_elim mp] i THEN assume_tac i; + +(*Like mp_tac but instantiates no variables*) +fun uniq_mp_tac i = eresolve_tac [notE,impE] i THEN uniq_assume_tac i; + + +(*** If-and-only-if ***) + +val iffI = prove_goalw IFOLP.thy [iff_def] + "[| !!x.x:P ==> q(x):Q; !!x.x:Q ==> r(x):P |] ==> ?p:P<->Q" + (fn prems=> [ (REPEAT (ares_tac (prems@[conjI, impI]) 1)) ]); + + +(*Observe use of rewrite_rule to unfold "<->" in meta-assumptions (prems) *) +val iffE = prove_goalw IFOLP.thy [iff_def] + "[| p:P <-> Q; !!x y.[| x:P-->Q; y:Q-->P |] ==> q(x,y):R |] ==> ?p:R" + (fn prems => [ (resolve_tac [conjE] 1), (REPEAT (ares_tac prems 1)) ]); + +(* Destruct rules for <-> similar to Modus Ponens *) + +val iffD1 = prove_goalw IFOLP.thy [iff_def] "[| p:P <-> Q; q:P |] ==> ?p:Q" + (fn prems => [ (rtac (conjunct1 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); + +val iffD2 = prove_goalw IFOLP.thy [iff_def] "[| p:P <-> Q; q:Q |] ==> ?p:P" + (fn prems => [ (rtac (conjunct2 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); + +val iff_refl = prove_goal IFOLP.thy "?p:P <-> P" + (fn _ => [ (REPEAT (ares_tac [iffI] 1)) ]); + +val iff_sym = prove_goal IFOLP.thy "p:Q <-> P ==> ?p:P <-> Q" + (fn [major] => + [ (rtac (major RS iffE) 1), + (rtac iffI 1), + (REPEAT (eresolve_tac [asm_rl,mp] 1)) ]); + +val iff_trans = prove_goal IFOLP.thy "[| p:P <-> Q; q:Q<-> R |] ==> ?p:P <-> R" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (REPEAT (eresolve_tac [asm_rl,iffE] 1 ORELSE mp_tac 1)) ]); + + +(*** Unique existence. NOTE THAT the following 2 quantifications + EX!x such that [EX!y such that P(x,y)] (sequential) + EX!x,y such that P(x,y) (simultaneous) + do NOT mean the same thing. The parser treats EX!x y.P(x,y) as sequential. +***) + +val ex1I = prove_goalw IFOLP.thy [ex1_def] + "[| p:P(a); !!x u.u:P(x) ==> f(u) : x=a |] ==> ?p:EX! x. P(x)" + (fn prems => [ (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ]); + +val ex1E = prove_goalw IFOLP.thy [ex1_def] + "[| p:EX! x.P(x); \ +\ !!x u v. [| u:P(x); v:ALL y. P(y) --> y=x |] ==> f(x,u,v):R |] ==>\ +\ ?a : R" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (eresolve_tac [exE,conjE] 1 ORELSE ares_tac prems 1)) ]); + + +(*** <-> congruence rules for simplification ***) + +(*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*) +fun iff_tac prems i = + resolve_tac (prems RL [iffE]) i THEN + REPEAT1 (eresolve_tac [asm_rl,mp] i); + +val conj_cong = prove_goal IFOLP.thy + "[| p:P <-> P'; !!x.x:P' ==> q(x):Q <-> Q' |] ==> ?p:(P&Q) <-> (P'&Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [iffI,conjI] 1 + ORELSE eresolve_tac [iffE,conjE,mp] 1 + ORELSE iff_tac prems 1)) ]); + +val disj_cong = prove_goal IFOLP.thy + "[| p:P <-> P'; q:Q <-> Q' |] ==> ?p:(P|Q) <-> (P'|Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (eresolve_tac [iffE,disjE,disjI1,disjI2] 1 + ORELSE ares_tac [iffI] 1 + ORELSE mp_tac 1)) ]); + +val imp_cong = prove_goal IFOLP.thy + "[| p:P <-> P'; !!x.x:P' ==> q(x):Q <-> Q' |] ==> ?p:(P-->Q) <-> (P'-->Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [iffI,impI] 1 + ORELSE eresolve_tac [iffE] 1 + ORELSE mp_tac 1 ORELSE iff_tac prems 1)) ]); + +val iff_cong = prove_goal IFOLP.thy + "[| p:P <-> P'; q:Q <-> Q' |] ==> ?p:(P<->Q) <-> (P'<->Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (eresolve_tac [iffE] 1 + ORELSE ares_tac [iffI] 1 + ORELSE mp_tac 1)) ]); + +val not_cong = prove_goal IFOLP.thy + "p:P <-> P' ==> ?p:~P <-> ~P'" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [iffI,notI] 1 + ORELSE mp_tac 1 + ORELSE eresolve_tac [iffE,notE] 1)) ]); + +val all_cong = prove_goal IFOLP.thy + "(!!x.f(x):P(x) <-> Q(x)) ==> ?p:(ALL x.P(x)) <-> (ALL x.Q(x))" + (fn prems => + [ (REPEAT (ares_tac [iffI,allI] 1 + ORELSE mp_tac 1 + ORELSE eresolve_tac [allE] 1 ORELSE iff_tac prems 1)) ]); + +val ex_cong = prove_goal IFOLP.thy + "(!!x.f(x):P(x) <-> Q(x)) ==> ?p:(EX x.P(x)) <-> (EX x.Q(x))" + (fn prems => + [ (REPEAT (eresolve_tac [exE] 1 ORELSE ares_tac [iffI,exI] 1 + ORELSE mp_tac 1 + ORELSE iff_tac prems 1)) ]); + +(*NOT PROVED +val ex1_cong = prove_goal IFOLP.thy + "(!!x.f(x):P(x) <-> Q(x)) ==> ?p:(EX! x.P(x)) <-> (EX! x.Q(x))" + (fn prems => + [ (REPEAT (eresolve_tac [ex1E, spec RS mp] 1 ORELSE ares_tac [iffI,ex1I] 1 + ORELSE mp_tac 1 + ORELSE iff_tac prems 1)) ]); +*) + +(*** Equality rules ***) + +val refl = ieqI; + +val subst = prove_goal IFOLP.thy "[| p:a=b; q:P(a) |] ==> ?p : P(b)" + (fn [prem1,prem2] => [ rtac (prem2 RS rev_mp) 1, (rtac (prem1 RS ieqE) 1), + rtac impI 1, atac 1 ]); + +val sym = prove_goal IFOLP.thy "q:a=b ==> ?c:b=a" + (fn [major] => [ (rtac (major RS subst) 1), (rtac refl 1) ]); + +val trans = prove_goal IFOLP.thy "[| p:a=b; q:b=c |] ==> ?d:a=c" + (fn [prem1,prem2] => [ (rtac (prem2 RS subst) 1), (rtac prem1 1) ]); + +(** ~ b=a ==> ~ a=b **) +val not_sym = prove_goal IFOLP.thy "p:~ b=a ==> ?q:~ a=b" + (fn [prem] => [ (rtac (prem RS contrapos) 1), (etac sym 1) ]); + +(*calling "standard" reduces maxidx to 0*) +val ssubst = standard (sym RS subst); + +(*A special case of ex1E that would otherwise need quantifier expansion*) +val ex1_equalsE = prove_goal IFOLP.thy + "[| p:EX! x.P(x); q:P(a); r:P(b) |] ==> ?d:a=b" + (fn prems => + [ (cut_facts_tac prems 1), + (etac ex1E 1), + (rtac trans 1), + (rtac sym 2), + (REPEAT (eresolve_tac [asm_rl, spec RS mp] 1)) ]); + +(** Polymorphic congruence rules **) + +val subst_context = prove_goal IFOLP.thy + "[| p:a=b |] ==> ?d:t(a)=t(b)" + (fn prems=> + [ (resolve_tac (prems RL [ssubst]) 1), + (resolve_tac [refl] 1) ]); + +val subst_context2 = prove_goal IFOLP.thy + "[| p:a=b; q:c=d |] ==> ?p:t(a,c)=t(b,d)" + (fn prems=> + [ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); + +val subst_context3 = prove_goal IFOLP.thy + "[| p:a=b; q:c=d; r:e=f |] ==> ?p:t(a,c,e)=t(b,d,f)" + (fn prems=> + [ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); + +(*Useful with eresolve_tac for proving equalties from known equalities. + a = b + | | + c = d *) +val box_equals = prove_goal IFOLP.thy + "[| p:a=b; q:a=c; r:b=d |] ==> ?p:c=d" + (fn prems=> + [ (resolve_tac [trans] 1), + (resolve_tac [trans] 1), + (resolve_tac [sym] 1), + (REPEAT (resolve_tac prems 1)) ]); + +(*Dual of box_equals: for proving equalities backwards*) +val simp_equals = prove_goal IFOLP.thy + "[| p:a=c; q:b=d; r:c=d |] ==> ?p:a=b" + (fn prems=> + [ (resolve_tac [trans] 1), + (resolve_tac [trans] 1), + (REPEAT (resolve_tac (prems @ (prems RL [sym])) 1)) ]); + +(** Congruence rules for predicate letters **) + +val pred1_cong = prove_goal IFOLP.thy + "p:a=a' ==> ?p:P(a) <-> P(a')" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); + +val pred2_cong = prove_goal IFOLP.thy + "[| p:a=a'; q:b=b' |] ==> ?p:P(a,b) <-> P(a',b')" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); + +val pred3_cong = prove_goal IFOLP.thy + "[| p:a=a'; q:b=b'; r:c=c' |] ==> ?p:P(a,b,c) <-> P(a',b',c')" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); + +(*special cases for free variables P, Q, R, S -- up to 3 arguments*) + +val pred_congs = + flat (map (fn c => + map (fn th => read_instantiate [("P",c)] th) + [pred1_cong,pred2_cong,pred3_cong]) + (explode"PQRS")); + +(*special case for the equality predicate!*) +val eq_cong = read_instantiate [("P","op =")] pred2_cong; + + +(*** Simplifications of assumed implications. + Roy Dyckhoff has proved that conj_impE, disj_impE, and imp_impE + used with mp_tac (restricted to atomic formulae) is COMPLETE for + intuitionistic propositional logic. See + R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic + (preprint, University of St Andrews, 1991) ***) + +val conj_impE = prove_goal IFOLP.thy + "[| p:(P&Q)-->S; !!x.x:P-->(Q-->S) ==> q(x):R |] ==> ?p:R" + (fn major::prems=> + [ (REPEAT (ares_tac ([conjI, impI, major RS mp]@prems) 1)) ]); + +val disj_impE = prove_goal IFOLP.thy + "[| p:(P|Q)-->S; !!x y.[| x:P-->S; y:Q-->S |] ==> q(x,y):R |] ==> ?p:R" + (fn major::prems=> + [ (DEPTH_SOLVE (ares_tac ([disjI1, disjI2, impI, major RS mp]@prems) 1)) ]); + +(*Simplifies the implication. Classical version is stronger. + Still UNSAFE since Q must be provable -- backtracking needed. *) +val imp_impE = prove_goal IFOLP.thy + "[| p:(P-->Q)-->S; !!x y.[| x:P; y:Q-->S |] ==> q(x,y):Q; !!x.x:S ==> r(x):R |] ==> \ +\ ?p:R" + (fn major::prems=> + [ (REPEAT (ares_tac ([impI, major RS mp]@prems) 1)) ]); + +(*Simplifies the implication. Classical version is stronger. + Still UNSAFE since ~P must be provable -- backtracking needed. *) +val not_impE = prove_goal IFOLP.thy + "[| p:~P --> S; !!y.y:P ==> q(y):False; !!y.y:S ==> r(y):R |] ==> ?p:R" + (fn major::prems=> + [ (REPEAT (ares_tac ([notI, impI, major RS mp]@prems) 1)) ]); + +(*Simplifies the implication. UNSAFE. *) +val iff_impE = prove_goal IFOLP.thy + "[| p:(P<->Q)-->S; !!x y.[| x:P; y:Q-->S |] ==> q(x,y):Q; \ +\ !!x y.[| x:Q; y:P-->S |] ==> r(x,y):P; !!x.x:S ==> s(x):R |] ==> ?p:R" + (fn major::prems=> + [ (REPEAT (ares_tac ([iffI, impI, major RS mp]@prems) 1)) ]); + +(*What if (ALL x.~~P(x)) --> ~~(ALL x.P(x)) is an assumption? UNSAFE*) +val all_impE = prove_goal IFOLP.thy + "[| p:(ALL x.P(x))-->S; !!x.q:P(x); !!y.y:S ==> r(y):R |] ==> ?p:R" + (fn major::prems=> + [ (REPEAT (ares_tac ([allI, impI, major RS mp]@prems) 1)) ]); + +(*Unsafe: (EX x.P(x))-->S is equivalent to ALL x.P(x)-->S. *) +val ex_impE = prove_goal IFOLP.thy + "[| p:(EX x.P(x))-->S; !!y.y:P(a)-->S ==> q(y):R |] ==> ?p:R" + (fn major::prems=> + [ (REPEAT (ares_tac ([exI, impI, major RS mp]@prems) 1)) ]); + +end; + +open IFOLP_Lemmas; + diff -r 000000000000 -r a5a9c433f639 src/FOLP/IFOLP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/IFOLP.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,141 @@ +IFOLP = Pure + + +classes term < logic + +default term + +types p,o 0 + +arities p,o :: logic + +consts + (*** Judgements ***) + "@Proof" :: "[p,o]=>prop" ("(_ /: _)" [10,10] 5) + Proof :: "[o,p]=>prop" + EqProof :: "[p,p,o]=>prop" ("(3_ /= _ :/ _)" [10,10,10] 5) + + (*** Logical Connectives -- Type Formers ***) + "=" :: "['a,'a] => o" (infixl 50) + True,False :: "o" + "Not" :: "o => o" ("~ _" [40] 40) + "&" :: "[o,o] => o" (infixr 35) + "|" :: "[o,o] => o" (infixr 30) + "-->" :: "[o,o] => o" (infixr 25) + "<->" :: "[o,o] => o" (infixr 25) + (*Quantifiers*) + All :: "('a => o) => o" (binder "ALL " 10) + Ex :: "('a => o) => o" (binder "EX " 10) + Ex1 :: "('a => o) => o" (binder "EX! " 10) + (*Rewriting gadgets*) + NORM :: "o => o" + norm :: "'a => 'a" + + (*** Proof Term Formers ***) + tt :: "p" + contr :: "p=>p" + fst,snd :: "p=>p" + pair :: "[p,p]=>p" ("(1<_,/_>)") + split :: "[p, [p,p]=>p] =>p" + inl,inr :: "p=>p" + when :: "[p, p=>p, p=>p]=>p" + lambda :: "(p => p) => p" (binder "lam " 20) + "`" :: "[p,p]=>p" (infixl 60) + alll :: "['a=>p]=>p" (binder "all " 15) + "^" :: "[p,'a]=>p" (infixl 50) + exists :: "['a,p]=>p" ("(1[_,/_])") + xsplit :: "[p,['a,p]=>p]=>p" + ideq :: "'a=>p" + idpeel :: "[p,'a=>p]=>p" + nrm, NRM :: "p" + +rules + +(**** Propositional logic ****) + +(*Equality*) +(* Like Intensional Equality in MLTT - but proofs distinct from terms *) + +ieqI "ideq(a) : a=a" +ieqE "[| p : a=b; !!x.f(x) : P(x,x) |] ==> idpeel(p,f) : P(a,b)" + +(* Truth and Falsity *) + +TrueI "tt : True" +FalseE "a:False ==> contr(a):P" + +(* Conjunction *) + +conjI "[| a:P; b:Q |] ==> : P&Q" +conjunct1 "p:P&Q ==> fst(p):P" +conjunct2 "p:P&Q ==> snd(p):Q" + +(* Disjunction *) + +disjI1 "a:P ==> inl(a):P|Q" +disjI2 "b:Q ==> inr(b):P|Q" +disjE "[| a:P|Q; !!x.x:P ==> f(x):R; !!x.x:Q ==> g(x):R \ +\ |] ==> when(a,f,g):R" + +(* Implication *) + +impI "(!!x.x:P ==> f(x):Q) ==> lam x.f(x):P-->Q" +mp "[| f:P-->Q; a:P |] ==> f`a:Q" + +(*Quantifiers*) + +allI "(!!x. f(x) : P(x)) ==> all x.f(x) : ALL x.P(x)" +spec "(f:ALL x.P(x)) ==> f^x : P(x)" + +exI "p : P(x) ==> [x,p] : EX x.P(x)" +exE "[| p: EX x.P(x); !!x u. u:P(x) ==> f(x,u) : R |] ==> xsplit(p,f):R" + +(**** Equality between proofs ****) + +prefl "a : P ==> a = a : P" +psym "a = b : P ==> b = a : P" +ptrans "[| a = b : P; b = c : P |] ==> a = c : P" + +idpeelB "[| !!x.f(x) : P(x,x) |] ==> idpeel(ideq(a),f) = f(a) : P(a,a)" + +fstB "a:P ==> fst() = a : P" +sndB "b:Q ==> snd() = b : Q" +pairEC "p:P&Q ==> p = : P&Q" + +whenBinl "[| a:P; !!x.x:P ==> f(x) : Q |] ==> when(inl(a),f,g) = f(a) : Q" +whenBinr "[| b:P; !!x.x:P ==> g(x) : Q |] ==> when(inr(b),f,g) = g(b) : Q" +plusEC "a:P|Q ==> when(a,%x.inl(x),%y.inr(y)) = p : P|Q" + +applyB "[| a:P; !!x.x:P ==> b(x) : Q |] ==> (lam x.b(x)) ` a = b(a) : Q" +funEC "f:P ==> f = lam x.f`x : P" + +specB "[| !!x.f(x) : P(x) |] ==> (all x.f(x)) ^ a = f(a) : P(a)" + + +(**** Definitions ****) + +not_def "~P == P-->False" +iff_def "P<->Q == (P-->Q) & (Q-->P)" + +(*Unique existence*) +ex1_def "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)" + +(*Rewriting -- special constants to flag normalized terms and formulae*) +norm_eq "nrm : norm(x) = x" +NORM_iff "NRM : NORM(P) <-> P" + +end + +ML + +(*show_proofs:=true displays the proof terms -- they are ENORMOUS*) +val show_proofs = ref false; + +fun proof_tr [p,P] = Const("Proof",dummyT) $ P $ p; + +fun proof_tr' [P,p] = + if !show_proofs then Const("@Proof",dummyT) $ p $ P + else P (*this case discards the proof term*); + +val parse_translation = [("@Proof", proof_tr)]; +val print_translation = [("Proof", proof_tr')]; + diff -r 000000000000 -r a5a9c433f639 src/FOLP/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/Makefile Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,43 @@ +######################################################################### +# # +# Makefile for Isabelle (FOLP) # +# # +######################################################################### + +#To make the system, cd to this directory and type +# make -f Makefile +#To make the system and test it on standard examples, type +# make -f Makefile test + +#Environment variable ISABELLECOMP specifies the compiler. +#Environment variable ISABELLEBIN specifies the destination directory. +#For Poly/ML, ISABELLEBIN must begin with a / + +#Makes pure Isabelle (Pure) if this file is ABSENT -- but not +#if it is out of date, since this Makefile does not know its dependencies! + +BIN = $(ISABELLEBIN) +COMP = $(ISABELLECOMP) +FILES = ROOT.ML ifolp.thy ifolp.ML folp.thy folp.ML int-prover.ML simpdata.ML\ + classical.ML ../Provers/simp.ML ../Provers/ind.ML + +$(BIN)/FOLP: $(BIN)/Pure $(FILES) + case "$(COMP)" in \ + poly*) echo 'make_database"$(BIN)/FOLP"; quit();' \ + | $(COMP) $(BIN)/Pure;\ + echo 'open PolyML; use"ROOT";' | $(COMP) $(BIN)/FOLP;;\ + sml*) echo 'use"ROOT.ML"; xML"$(BIN)/FOLP" banner;' | $(BIN)/Pure;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +$(BIN)/Pure: + cd ../Pure; $(MAKE) + +test: ex/ROOT.ML $(BIN)/FOLP + case "$(COMP)" in \ + poly*) echo 'use"ex/ROOT.ML"; quit();' | $(COMP) $(BIN)/FOLP ;;\ + sml*) echo 'use"ex/ROOT.ML";' | $(BIN)/FOLP;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +.PRECIOUS: $(BIN)/Pure $(BIN)/FOLP diff -r 000000000000 -r a5a9c433f639 src/FOLP/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,78 @@ +(* Title: FOLP/ROOT + ID: $Id$ + Author: martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Modifed version of Lawrence Paulson's FOL that contains proof terms. + +Presence of unknown proof term means that matching does not behave as expected. +*) + +val banner = "First-Order Logic with Natural Deduction with Proof Terms"; + +writeln banner; + +print_depth 1; +use_thy "ifolp"; +use_thy "folp"; + +use "../Provers/hypsubst.ML"; +use "classical.ML"; (* Patched 'cos matching won't instantiate proof *) +use "simp.ML"; (* Patched 'cos matching won't instantiate proof *) +use "../Provers/ind.ML"; + + +(*** Applying HypsubstFun to generate hyp_subst_tac ***) + +structure Hypsubst_Data = + struct + (*Take apart an equality judgement; otherwise raise Match!*) + fun dest_eq (Const("Proof",_) $ (Const("op =",_) $ t $ u) $ _) = (t,u); + + val imp_intr = impI + + (*etac rev_cut_eq moves an equality to be the last premise. *) + val rev_cut_eq = prove_goal IFOLP.thy + "[| p:a=b; !!x.x:a=b ==> f(x):R |] ==> ?p:R" + (fn prems => [ REPEAT(resolve_tac prems 1) ]); + + val rev_mp = rev_mp + val subst = subst + val sym = sym + end; + +structure Hypsubst = HypsubstFun(Hypsubst_Data); +open Hypsubst; + +use "int-prover.ML"; + +(*** Applying ClassicalFun to create a classical prover ***) +structure Classical_Data = + struct + val sizef = size_of_thm + val mp = mp + val not_elim = notE + val swap = swap + val hyp_subst_tacs=[hyp_subst_tac] + end; + +structure Cla = ClassicalFun(Classical_Data); +open Cla; + +(*Propositional rules + -- iffCE might seem better, but in the examples in ex/cla + run about 7% slower than with iffE*) +val prop_cs = empty_cs addSIs [refl,TrueI,conjI,disjCI,impI,notI,iffI] + addSEs [conjE,disjE,impCE,FalseE,iffE]; + +(*Quantifier rules*) +val FOLP_cs = prop_cs addSIs [allI] addIs [exI,ex1I] + addSEs [exE,ex1E] addEs [allE]; + +val FOLP_dup_cs = prop_cs addSIs [allI] addIs [exCI,ex1I] + addSEs [exE,ex1E] addEs [all_dupE]; + +use "simpdata.ML"; + +use "../Pure/install_pp.ML"; +print_depth 8; diff -r 000000000000 -r a5a9c433f639 src/FOLP/change --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/change Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,18 @@ +#! /bin/sh +# +# Usage: +# expandshort FILE1 ... FILEn +# +# leaves previous versions as XXX~~ +# +for f in $* +do +echo Expanding shorthands in $f. \ Backup file is $f~~ +mv $f $f~~; sed -e ' +s/PFOL/FOLP/g +s/PIFOL/IFOLP/g +s/pfol/folp/g +s/pifol/ifolp/g +' $f~~ > $f +done +echo Finished. diff -r 000000000000 -r a5a9c433f639 src/FOLP/classical.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/classical.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,191 @@ +(* Title: FOLP/classical + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Like Provers/classical but modified because match_tac is unsuitable for +proof objects. + +Theorem prover for classical reasoning, including predicate calculus, set +theory, etc. + +Rules must be classified as intr, elim, safe, hazardous. + +A rule is unsafe unless it can be applied blindly without harmful results. +For a rule to be safe, its premises and conclusion should be logically +equivalent. There should be no variables in the premises that are not in +the conclusion. +*) + +signature CLASSICAL_DATA = + sig + val mp: thm (* [| P-->Q; P |] ==> Q *) + val not_elim: thm (* [| ~P; P |] ==> R *) + val swap: thm (* ~P ==> (~Q ==> P) ==> Q *) + val sizef : thm -> int (* size function for BEST_FIRST *) + val hyp_subst_tacs: (int -> tactic) list + end; + +(*Higher precedence than := facilitates use of references*) +infix 4 addSIs addSEs addSDs addIs addEs addDs; + + +signature CLASSICAL = + sig + type claset + val empty_cs: claset + val addDs : claset * thm list -> claset + val addEs : claset * thm list -> claset + val addIs : claset * thm list -> claset + val addSDs: claset * thm list -> claset + val addSEs: claset * thm list -> claset + val addSIs: claset * thm list -> claset + val print_cs: claset -> unit + val rep_claset: claset -> + {safeIs: thm list, safeEs: thm list, hazIs: thm list, hazEs: thm list, + safe0_brls:(bool*thm)list, safep_brls: (bool*thm)list, + haz_brls: (bool*thm)list} + val best_tac : claset -> int -> tactic + val chain_tac : int -> tactic + val contr_tac : int -> tactic + val fast_tac : claset -> int -> tactic + val inst_step_tac : int -> tactic + val joinrules : thm list * thm list -> (bool * thm) list + val mp_tac: int -> tactic + val safe_tac : claset -> tactic + val safe_step_tac : claset -> int -> tactic + val slow_step_tac : claset -> int -> tactic + val step_tac : claset -> int -> tactic + val swapify : thm list -> thm list + val swap_res_tac : thm list -> int -> tactic + val uniq_mp_tac: int -> tactic + end; + + +functor ClassicalFun(Data: CLASSICAL_DATA): CLASSICAL = +struct + +local open Data in + +(** Useful tactics for classical reasoning **) + +val imp_elim = make_elim mp; + +(*Solve goal that assumes both P and ~P. *) +val contr_tac = eresolve_tac [not_elim] THEN' assume_tac; + +(*Finds P-->Q and P in the assumptions, replaces implication by Q *) +fun mp_tac i = eresolve_tac ([not_elim,imp_elim]) i THEN assume_tac i; + +(*Like mp_tac but instantiates no variables*) +fun uniq_mp_tac i = ematch_tac ([not_elim,imp_elim]) i THEN uniq_assume_tac i; + +(*Creates rules to eliminate ~A, from rules to introduce A*) +fun swapify intrs = intrs RLN (2, [swap]); + +(*Uses introduction rules in the normal way, or on negated assumptions, + trying rules in order. *) +fun swap_res_tac rls = + let fun tacf rl = rtac rl ORELSE' etac (rl RSN (2,swap)) + in assume_tac ORELSE' contr_tac ORELSE' FIRST' (map tacf rls) + end; + +(*Given assumption P-->Q, reduces subgoal Q to P [deletes the implication!] *) +fun chain_tac i = + eresolve_tac [imp_elim] i THEN + (assume_tac (i+1) ORELSE contr_tac (i+1)); + +(*** Classical rule sets ***) + +datatype claset = + CS of {safeIs: thm list, + safeEs: thm list, + hazIs: thm list, + hazEs: thm list, + (*the following are computed from the above*) + safe0_brls: (bool*thm)list, + safep_brls: (bool*thm)list, + haz_brls: (bool*thm)list}; + +fun rep_claset (CS x) = x; + +(*For use with biresolve_tac. Combines intrs with swap to catch negated + assumptions. Also pairs elims with true. *) +fun joinrules (intrs,elims) = + map (pair true) (elims @ swapify intrs) @ map (pair false) intrs; + +(*Note that allE precedes exI in haz_brls*) +fun make_cs {safeIs,safeEs,hazIs,hazEs} = + let val (safe0_brls, safep_brls) = (*0 subgoals vs 1 or more*) + partition (apl(0,op=) o subgoals_of_brl) + (sort lessb (joinrules(safeIs, safeEs))) + in CS{safeIs=safeIs, safeEs=safeEs, hazIs=hazIs, hazEs=hazEs, + safe0_brls=safe0_brls, safep_brls=safep_brls, + haz_brls = sort lessb (joinrules(hazIs, hazEs))} + end; + +(*** Manipulation of clasets ***) + +val empty_cs = make_cs{safeIs=[], safeEs=[], hazIs=[], hazEs=[]}; + +fun print_cs (CS{safeIs,safeEs,hazIs,hazEs,...}) = + (writeln"Introduction rules"; prths hazIs; + writeln"Safe introduction rules"; prths safeIs; + writeln"Elimination rules"; prths hazEs; + writeln"Safe elimination rules"; prths safeEs; + ()); + +fun (CS{safeIs,safeEs,hazIs,hazEs,...}) addSIs ths = + make_cs {safeIs=ths@safeIs, safeEs=safeEs, hazIs=hazIs, hazEs=hazEs}; + +fun (CS{safeIs,safeEs,hazIs,hazEs,...}) addSEs ths = + make_cs {safeIs=safeIs, safeEs=ths@safeEs, hazIs=hazIs, hazEs=hazEs}; + +fun cs addSDs ths = cs addSEs (map make_elim ths); + +fun (CS{safeIs,safeEs,hazIs,hazEs,...}) addIs ths = + make_cs {safeIs=safeIs, safeEs=safeEs, hazIs=ths@hazIs, hazEs=hazEs}; + +fun (CS{safeIs,safeEs,hazIs,hazEs,...}) addEs ths = + make_cs {safeIs=safeIs, safeEs=safeEs, hazIs=hazIs, hazEs=ths@hazEs}; + +fun cs addDs ths = cs addEs (map make_elim ths); + +(*** Simple tactics for theorem proving ***) + +(*Attack subgoals using safe inferences*) +fun safe_step_tac (CS{safe0_brls,safep_brls,...}) = + FIRST' [uniq_assume_tac, + uniq_mp_tac, + biresolve_tac safe0_brls, + FIRST' hyp_subst_tacs, + biresolve_tac safep_brls] ; + +(*Repeatedly attack subgoals using safe inferences*) +fun safe_tac cs = DETERM (REPEAT_FIRST (safe_step_tac cs)); + +(*These steps could instantiate variables and are therefore unsafe.*) +val inst_step_tac = assume_tac APPEND' contr_tac; + +(*Single step for the prover. FAILS unless it makes progress. *) +fun step_tac (cs as (CS{haz_brls,...})) i = + FIRST [safe_tac cs, + inst_step_tac i, + biresolve_tac haz_brls i]; + +(*** The following tactics all fail unless they solve one goal ***) + +(*Dumb but fast*) +fun fast_tac cs = SELECT_GOAL (DEPTH_SOLVE (step_tac cs 1)); + +(*Slower but smarter than fast_tac*) +fun best_tac cs = + SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, sizef) (step_tac cs 1)); + +(*Using a "safe" rule to instantiate variables is unsafe. This tactic + allows backtracking from "safe" rules to "unsafe" rules here.*) +fun slow_step_tac (cs as (CS{haz_brls,...})) i = + safe_tac cs ORELSE (assume_tac i APPEND biresolve_tac haz_brls i); + +end; +end; diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/If.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/If.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,43 @@ +(* Title: FOLP/ex/if + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For ex/if.thy. First-Order Logic: the 'if' example +*) + +open If; +open Cla; (*in case structure Int is open!*) + +val prems = goalw If.thy [if_def] + "[| !!x.x:P ==> f(x):Q; !!x.x:~P ==> g(x):R |] ==> ?p:if(P,Q,R)"; +by (fast_tac (FOLP_cs addIs prems) 1); +val ifI = result(); + +val major::prems = goalw If.thy [if_def] + "[| p:if(P,Q,R); !!x y.[| x:P; y:Q |] ==> f(x,y):S; \ +\ !!x y.[| x:~P; y:R |] ==> g(x,y):S |] ==> ?p:S"; +by (cut_facts_tac [major] 1); +by (fast_tac (FOLP_cs addIs prems) 1); +val ifE = result(); + + +goal If.thy + "?p : if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))"; +by (resolve_tac [iffI] 1); +by (eresolve_tac [ifE] 1); +by (eresolve_tac [ifE] 1); +by (resolve_tac [ifI] 1); +by (resolve_tac [ifI] 1); + +choplev 0; +val if_cs = FOLP_cs addSIs [ifI] addSEs[ifE]; +by (fast_tac if_cs 1); +val if_commute = result(); + + +goal If.thy "?p : if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))"; +by (fast_tac if_cs 1); +val nested_ifs = result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/If.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/If.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,5 @@ +If = FOLP + +consts if :: "[o,o,o]=>o" +rules + if_def "if(P,Q,R) == P&Q | ~P&R" +end diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/Nat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/Nat.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,82 @@ +(* Title: FOLP/ex/nat.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Examples for the manual "Introduction to Isabelle" + +Proofs about the natural numbers + +To generate similar output to manual, execute these commands: + Pretty.setmargin 72; print_depth 0; +*) + +open Nat; + +goal Nat.thy "?p : ~ (Suc(k) = k)"; +by (res_inst_tac [("n","k")] induct 1); +by (rtac notI 1); +by (etac Suc_neq_0 1); +by (rtac notI 1); +by (etac notE 1); +by (etac Suc_inject 1); +val Suc_n_not_n = result(); + + +goal Nat.thy "?p : (k+m)+n = k+(m+n)"; +prths ([induct] RL [topthm()]); (*prints all 14 next states!*) +by (rtac induct 1); +back(); +back(); +back(); +back(); +back(); +back(); + +goalw Nat.thy [add_def] "?p : 0+n = n"; +by (rtac rec_0 1); +val add_0 = result(); + +goalw Nat.thy [add_def] "?p : Suc(m)+n = Suc(m+n)"; +by (rtac rec_Suc 1); +val add_Suc = result(); + +(* +val nat_congs = mk_congs Nat.thy ["Suc", "op +"]; +prths nat_congs; +*) +val prems = goal Nat.thy "p: x=y ==> ?p : Suc(x) = Suc(y)"; +by (resolve_tac (prems RL [subst]) 1); +by (rtac refl 1); +val Suc_cong = result(); + +val prems = goal Nat.thy "[| p : a=x; q: b=y |] ==> ?p : a+b=x+y"; +by (resolve_tac (prems RL [subst]) 1 THEN resolve_tac (prems RL [subst]) 1 THEN + rtac refl 1); +val Plus_cong = result(); + +val nat_congs = [Suc_cong,Plus_cong]; + + +val add_ss = FOLP_ss addcongs nat_congs + addrews [add_0, add_Suc]; + +goal Nat.thy "?p : (k+m)+n = k+(m+n)"; +by (res_inst_tac [("n","k")] induct 1); +by (SIMP_TAC add_ss 1); +by (ASM_SIMP_TAC add_ss 1); +val add_assoc = result(); + +goal Nat.thy "?p : m+0 = m"; +by (res_inst_tac [("n","m")] induct 1); +by (SIMP_TAC add_ss 1); +by (ASM_SIMP_TAC add_ss 1); +val add_0_right = result(); + +goal Nat.thy "?p : m+Suc(n) = Suc(m+n)"; +by (res_inst_tac [("n","m")] induct 1); +by (ALLGOALS (ASM_SIMP_TAC add_ss)); +val add_Suc_right = result(); + +(*mk_typed_congs appears not to work with FOLP's version of subst*) + diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/Nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/Nat.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,36 @@ +(* Title: FOLP/ex/nat.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Examples for the manual "Introduction to Isabelle" + +Theory of the natural numbers: Peano's axioms, primitive recursion +*) + +Nat = IFOLP + +types nat 0 +arities nat :: term +consts "0" :: "nat" ("0") + Suc :: "nat=>nat" + rec :: "[nat, 'a, [nat,'a]=>'a] => 'a" + "+" :: "[nat, nat] => nat" (infixl 60) + + (*Proof terms*) + nrec :: "[nat,p,[nat,p]=>p]=>p" + ninj,nneq :: "p=>p" + rec0, recSuc :: "p" + +rules + induct "[| b:P(0); !!x u. u:P(x) ==> c(x,u):P(Suc(x)) \ +\ |] ==> nrec(n,b,c):P(n)" + + Suc_inject "p:Suc(m)=Suc(n) ==> ninj(p) : m=n" + Suc_neq_0 "p:Suc(m)=0 ==> nneq(p) : R" + rec_0 "rec0 : rec(0,a,f) = a" + rec_Suc "recSuc : rec(Suc(m), a, f) = f(m, rec(m,a,f))" + add_def "m+n == rec(m, n, %x y. Suc(y))" + + nrecB0 "b: A ==> nrec(0,b,c) = b : A" + nrecBSuc "c(n,nrec(n,b,c)) : A ==> nrec(Suc(n),b,c) = c(n,nrec(n,b,c)) : A" +end diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/Prolog.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/Prolog.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,74 @@ +(* Title: FOL/ex/prolog.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For ex/prolog.thy. First-Order Logic: PROLOG examples +*) + +open Prolog; + +goal Prolog.thy "app(a:b:c:Nil, d:e:Nil, ?x)"; +by (resolve_tac [appNil,appCons] 1); +by (resolve_tac [appNil,appCons] 1); +by (resolve_tac [appNil,appCons] 1); +by (resolve_tac [appNil,appCons] 1); +prth (result()); + +goal Prolog.thy "app(?x, c:d:Nil, a:b:c:d:Nil)"; +by (REPEAT (resolve_tac [appNil,appCons] 1)); +result(); + + +goal Prolog.thy "app(?x, ?y, a:b:c:d:Nil)"; +by (REPEAT (resolve_tac [appNil,appCons] 1)); +back(); +back(); +back(); +back(); +result(); + + +(*app([x1,...,xn], y, ?z) requires (n+1) inferences*) +(*rev([x1,...,xn], ?y) requires (n+1)(n+2)/2 inferences*) + +goal Prolog.thy "rev(a:b:c:d:Nil, ?x)"; +val rules = [appNil,appCons,revNil,revCons]; +by (REPEAT (resolve_tac rules 1)); +result(); + +goal Prolog.thy "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:Nil, ?w)"; +by (REPEAT (resolve_tac rules 1)); +result(); + +goal Prolog.thy "rev(?x, a:b:c:Nil)"; +by (REPEAT (resolve_tac rules 1)); (*does not solve it directly!*) +back(); +back(); + +(*backtracking version*) +val prolog_tac = DEPTH_FIRST (has_fewer_prems 1) (resolve_tac rules 1); + +choplev 0; +by prolog_tac; +result(); + +goal Prolog.thy "rev(a:?x:c:?y:Nil, d:?z:b:?u)"; +by prolog_tac; +result(); + +(*rev([a..p], ?w) requires 153 inferences *) +goal Prolog.thy "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil, ?w)"; +by (DEPTH_SOLVE (resolve_tac ([refl,conjI]@rules) 1)); +(*Poly/ML: 4 secs >> 38 lips*) +result(); + +(*?x has 16, ?y has 32; rev(?y,?w) requires 561 (rather large) inferences; + total inferences = 2 + 1 + 17 + 561 = 581*) +goal Prolog.thy + "a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil = ?x & app(?x,?x,?y) & rev(?y,?w)"; +by (DEPTH_SOLVE (resolve_tac ([refl,conjI]@rules) 1)); +(*Poly/ML: 29 secs >> 20 lips*) +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/Prolog.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/Prolog.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,22 @@ +(* Title: FOL/ex/prolog.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +First-Order Logic: PROLOG examples + +Inherits from FOL the class term, the type o, and the coercion Trueprop +*) + +Prolog = FOL + +types list 1 +arities list :: (term)term +consts Nil :: "'a list" + ":" :: "['a, 'a list]=> 'a list" (infixr 60) + app :: "['a list, 'a list, 'a list] => o" + rev :: "['a list, 'a list] => o" +rules appNil "app(Nil,ys,ys)" + appCons "app(xs,ys,zs) ==> app(x:xs, ys, x:zs)" + revNil "rev(Nil,Nil)" + revCons "[| rev(xs,ys); app(ys, x:Nil, zs) |] ==> rev(x:xs, zs)" +end diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,33 @@ +(* Title: FOLP/ex/ROOT + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Executes all examples for First-Order Logic. +*) + +writeln"Root file for FOLP examples"; + +proof_timing := true; + +time_use "ex/intro.ML"; +time_use_thy "ex/nat"; +time_use "ex/foundn.ML"; + +writeln"\n** Intuitionistic examples **\n"; +time_use "ex/int.ML"; + +val thy = IFOLP.thy and tac = Int.fast_tac 1; +time_use "ex/prop.ML"; +time_use "ex/quant.ML"; +commit(); + +writeln"\n** Classical examples **\n"; +time_use "ex/cla.ML"; +time_use_thy "ex/if"; + +val thy = FOLP.thy and tac = Cla.fast_tac FOLP_cs 1; +time_use "ex/prop.ML"; +time_use "ex/quant.ML"; + +maketest"END: Root file for FOLP examples"; diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/cla.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/cla.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,360 @@ +(* Title: FOLP/ex/cla + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Classical First-Order Logic +*) + +writeln"File FOL/ex/cla."; + +open Cla; (*in case structure Int is open!*) + +goal FOLP.thy "?p : (P --> Q | R) --> (P-->Q) | (P-->R)"; +by (fast_tac FOLP_cs 1); +result(); + +(*If and only if*) + +goal FOLP.thy "?p : (P<->Q) <-> (Q<->P)"; +by (fast_tac FOLP_cs 1); +result(); + +goal FOLP.thy "?p : ~ (P <-> ~P)"; +by (fast_tac FOLP_cs 1); +result(); + + +(*Sample problems from + F. J. Pelletier, + Seventy-Five Problems for Testing Automatic Theorem Provers, + J. Automated Reasoning 2 (1986), 191-216. + Errata, JAR 4 (1988), 236-236. + +The hardest problems -- judging by experience with several theorem provers, +including matrix ones -- are 34 and 43. +*) + +writeln"Pelletier's examples"; +(*1*) +goal FOLP.thy "?p : (P-->Q) <-> (~Q --> ~P)"; +by (fast_tac FOLP_cs 1); +result(); + +(*2*) +goal FOLP.thy "?p : ~ ~ P <-> P"; +by (fast_tac FOLP_cs 1); +result(); + +(*3*) +goal FOLP.thy "?p : ~(P-->Q) --> (Q-->P)"; +by (fast_tac FOLP_cs 1); +result(); + +(*4*) +goal FOLP.thy "?p : (~P-->Q) <-> (~Q --> P)"; +by (fast_tac FOLP_cs 1); +result(); + +(*5*) +goal FOLP.thy "?p : ((P|Q)-->(P|R)) --> (P|(Q-->R))"; +by (fast_tac FOLP_cs 1); +result(); + +(*6*) +goal FOLP.thy "?p : P | ~ P"; +by (fast_tac FOLP_cs 1); +result(); + +(*7*) +goal FOLP.thy "?p : P | ~ ~ ~ P"; +by (fast_tac FOLP_cs 1); +result(); + +(*8. Peirce's law*) +goal FOLP.thy "?p : ((P-->Q) --> P) --> P"; +by (fast_tac FOLP_cs 1); +result(); + +(*9*) +goal FOLP.thy "?p : ((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"; +by (fast_tac FOLP_cs 1); +result(); + +(*10*) +goal FOLP.thy "?p : (Q-->R) & (R-->P&Q) & (P-->Q|R) --> (P<->Q)"; +by (fast_tac FOLP_cs 1); +result(); + +(*11. Proved in each direction (incorrectly, says Pelletier!!) *) +goal FOLP.thy "?p : P<->P"; +by (fast_tac FOLP_cs 1); +result(); + +(*12. "Dijkstra's law"*) +goal FOLP.thy "?p : ((P <-> Q) <-> R) <-> (P <-> (Q <-> R))"; +by (fast_tac FOLP_cs 1); +result(); + +(*13. Distributive law*) +goal FOLP.thy "?p : P | (Q & R) <-> (P | Q) & (P | R)"; +by (fast_tac FOLP_cs 1); +result(); + +(*14*) +goal FOLP.thy "?p : (P <-> Q) <-> ((Q | ~P) & (~Q|P))"; +by (fast_tac FOLP_cs 1); +result(); + +(*15*) +goal FOLP.thy "?p : (P --> Q) <-> (~P | Q)"; +by (fast_tac FOLP_cs 1); +result(); + +(*16*) +goal FOLP.thy "?p : (P-->Q) | (Q-->P)"; +by (fast_tac FOLP_cs 1); +result(); + +(*17*) +goal FOLP.thy "?p : ((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Classical Logic: examples with quantifiers"; + +goal FOLP.thy "?p : (ALL x. P(x) & Q(x)) <-> (ALL x. P(x)) & (ALL x. Q(x))"; +by (fast_tac FOLP_cs 1); +result(); + +goal FOLP.thy "?p : (EX x. P-->Q(x)) <-> (P --> (EX x.Q(x)))"; +by (fast_tac FOLP_cs 1); +result(); + +goal FOLP.thy "?p : (EX x.P(x)-->Q) <-> (ALL x.P(x)) --> Q"; +by (fast_tac FOLP_cs 1); +result(); + +goal FOLP.thy "?p : (ALL x.P(x)) | Q <-> (ALL x. P(x) | Q)"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problems requiring quantifier duplication"; + +(*Needs multiple instantiation of ALL.*) +(* +goal FOLP.thy "?p : (ALL x. P(x)-->P(f(x))) & P(d)-->P(f(f(f(d))))"; +by (best_tac FOLP_dup_cs 1); +result(); +*) +(*Needs double instantiation of the quantifier*) +goal FOLP.thy "?p : EX x. P(x) --> P(a) & P(b)"; +by (best_tac FOLP_dup_cs 1); +result(); + +goal FOLP.thy "?p : EX z. P(z) --> (ALL x. P(x))"; +by (best_tac FOLP_dup_cs 1); +result(); + + +writeln"Hard examples with quantifiers"; + +writeln"Problem 18"; +goal FOLP.thy "?p : EX y. ALL x. P(y)-->P(x)"; +by (best_tac FOLP_dup_cs 1); +result(); + +writeln"Problem 19"; +goal FOLP.thy "?p : EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))"; +by (best_tac FOLP_dup_cs 1); +result(); + +writeln"Problem 20"; +goal FOLP.thy "?p : (ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w))) \ +\ --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"; +by (fast_tac FOLP_cs 1); +result(); +(* +writeln"Problem 21"; +goal FOLP.thy "?p : (EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> (EX x. P<->Q(x))"; +by (best_tac FOLP_dup_cs 1); +result(); +*) +writeln"Problem 22"; +goal FOLP.thy "?p : (ALL x. P <-> Q(x)) --> (P <-> (ALL x. Q(x)))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 23"; +goal FOLP.thy "?p : (ALL x. P | Q(x)) <-> (P | (ALL x. Q(x)))"; +by (best_tac FOLP_cs 1); +result(); + +writeln"Problem 24"; +goal FOLP.thy "?p : ~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) & \ +\ ~(EX x.P(x)) --> (EX x.Q(x)) & (ALL x. Q(x)|R(x) --> S(x)) \ +\ --> (EX x. P(x)&R(x))"; +by (fast_tac FOLP_cs 1); +result(); +(* +writeln"Problem 25"; +goal FOLP.thy "?p : (EX x. P(x)) & \ +\ (ALL x. L(x) --> ~ (M(x) & R(x))) & \ +\ (ALL x. P(x) --> (M(x) & L(x))) & \ +\ ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x))) \ +\ --> (EX x. Q(x)&P(x))"; +by (best_tac FOLP_cs 1); +result(); + +writeln"Problem 26"; +goal FOLP.thy "?u : ((EX x. p(x)) <-> (EX x. q(x))) & \ +\ (ALL x. ALL y. p(x) & q(y) --> (r(x) <-> s(y))) \ +\ --> ((ALL x. p(x)-->r(x)) <-> (ALL x. q(x)-->s(x)))"; +by (fast_tac FOLP_cs 1); +result(); +*) +writeln"Problem 27"; +goal FOLP.thy "?p : (EX x. P(x) & ~Q(x)) & \ +\ (ALL x. P(x) --> R(x)) & \ +\ (ALL x. M(x) & L(x) --> P(x)) & \ +\ ((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x))) \ +\ --> (ALL x. M(x) --> ~L(x))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 28. AMENDED"; +goal FOLP.thy "?p : (ALL x. P(x) --> (ALL x. Q(x))) & \ +\ ((ALL x. Q(x)|R(x)) --> (EX x. Q(x)&S(x))) & \ +\ ((EX x.S(x)) --> (ALL x. L(x) --> M(x))) \ +\ --> (ALL x. P(x) & L(x) --> M(x))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 29. Essentially the same as Principia Mathematica *11.71"; +goal FOLP.thy "?p : (EX x. P(x)) & (EX y. Q(y)) \ +\ --> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y)) <-> \ +\ (ALL x y. P(x) & Q(y) --> R(x) & S(y)))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 30"; +goal FOLP.thy "?p : (ALL x. P(x) | Q(x) --> ~ R(x)) & \ +\ (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x)) \ +\ --> (ALL x. S(x))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 31"; +goal FOLP.thy "?p : ~(EX x.P(x) & (Q(x) | R(x))) & \ +\ (EX x. L(x) & P(x)) & \ +\ (ALL x. ~ R(x) --> M(x)) \ +\ --> (EX x. L(x) & M(x))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 32"; +goal FOLP.thy "?p : (ALL x. P(x) & (Q(x)|R(x))-->S(x)) & \ +\ (ALL x. S(x) & R(x) --> L(x)) & \ +\ (ALL x. M(x) --> R(x)) \ +\ --> (ALL x. P(x) & M(x) --> L(x))"; +by (best_tac FOLP_cs 1); +result(); + +writeln"Problem 33"; +goal FOLP.thy "?p : (ALL x. P(a) & (P(x)-->P(b))-->P(c)) <-> \ +\ (ALL x. (~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c)))"; +by (best_tac FOLP_cs 1); +result(); + +writeln"Problem 35"; +goal FOLP.thy "?p : EX x y. P(x,y) --> (ALL u v. P(u,v))"; +by (best_tac FOLP_dup_cs 1); +result(); + +writeln"Problem 36"; +goal FOLP.thy +"?p : (ALL x. EX y. J(x,y)) & \ +\ (ALL x. EX y. G(x,y)) & \ +\ (ALL x y. J(x,y) | G(x,y) --> (ALL z. J(y,z) | G(y,z) --> H(x,z))) \ +\ --> (ALL x. EX y. H(x,y))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 37"; +goal FOLP.thy "?p : (ALL z. EX w. ALL x. EX y. \ +\ (P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (EX u.Q(u,w)))) & \ +\ (ALL x z. ~P(x,z) --> (EX y. Q(y,z))) & \ +\ ((EX x y. Q(x,y)) --> (ALL x. R(x,x))) \ +\ --> (ALL x. EX y. R(x,y))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 39"; +goal FOLP.thy "?p : ~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 40. AMENDED"; +goal FOLP.thy "?p : (EX y. ALL x. F(x,y) <-> F(x,x)) --> \ +\ ~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 41"; +goal FOLP.thy "?p : (ALL z. EX y. ALL x. f(x,y) <-> f(x,z) & ~ f(x,x)) \ +\ --> ~ (EX z. ALL x. f(x,z))"; +by (best_tac FOLP_cs 1); +result(); + +writeln"Problem 44"; +goal FOLP.thy "?p : (ALL x. f(x) --> \ +\ (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y)))) & \ +\ (EX x. j(x) & (ALL y. g(y) --> h(x,y))) \ +\ --> (EX x. j(x) & ~f(x))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problems (mainly) involving equality or functions"; + +writeln"Problem 48"; +goal FOLP.thy "?p : (a=b | c=d) & (a=c | b=d) --> a=d | b=c"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 50"; +(*What has this to do with equality?*) +goal FOLP.thy "?p : (ALL x. P(a,x) | (ALL y.P(x,y))) --> (EX x. ALL y.P(x,y))"; +by (best_tac FOLP_dup_cs 1); +result(); + +writeln"Problem 56"; +goal FOLP.thy + "?p : (ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 57"; +goal FOLP.thy +"?p : P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) & \ +\ (ALL x y z. P(x,y) & P(y,z) --> P(x,z)) --> P(f(a,b), f(a,c))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Problem 58 NOT PROVED AUTOMATICALLY"; +goal FOLP.thy "?p : (ALL x y. f(x)=g(y)) --> (ALL x y. f(f(x))=f(g(y)))"; +val f_cong = read_instantiate [("t","f")] subst_context; +by (fast_tac (FOLP_cs addIs [f_cong]) 1); +result(); + +writeln"Problem 59"; +goal FOLP.thy "?p : (ALL x. P(x) <-> ~P(f(x))) --> (EX x. P(x) & ~P(f(x)))"; +by (best_tac FOLP_dup_cs 1); +result(); + +writeln"Problem 60"; +goal FOLP.thy +"?p : ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"; +by (fast_tac FOLP_cs 1); +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/foundn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/foundn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,134 @@ +(* Title: FOL/ex/foundn + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Intuitionistic FOL: Examples from The Foundation of a Generic Theorem Prover +*) + +writeln"File FOL/ex/foundn."; + +goal IFOLP.thy "?p : A&B --> (C-->A&C)"; +by (resolve_tac [impI] 1); +by (resolve_tac [impI] 1); +by (resolve_tac [conjI] 1); +by (assume_tac 2); +by (resolve_tac [conjunct1] 1); +by (assume_tac 1); +result(); + +(*A form of conj-elimination*) +val prems = +goal IFOLP.thy "p : A&B ==> (!!x y.[| x:A; y:B |] ==> f(x,y):C) ==> ?p:C"; +by (resolve_tac prems 1); +by (resolve_tac [conjunct1] 1); +by (resolve_tac prems 1); +by (resolve_tac [conjunct2] 1); +by (resolve_tac prems 1); +result(); + + +val prems = +goal IFOLP.thy "(!!A x. x:~ ~A ==> cla(x):A) ==> ?p:B | ~B"; +by (resolve_tac prems 1); +by (resolve_tac [notI] 1); +by (res_inst_tac [ ("P", "~B") ] notE 1); +by (resolve_tac [notI] 2); +by (res_inst_tac [ ("P", "B | ~B") ] notE 2); +by (assume_tac 2); +by (resolve_tac [disjI1] 2); +by (assume_tac 2); +by (resolve_tac [notI] 1); +by (res_inst_tac [ ("P", "B | ~B") ] notE 1); +by (assume_tac 1); +by (resolve_tac [disjI2] 1); +by (assume_tac 1); +result(); + + +val prems = +goal IFOLP.thy "(!!A x. x:~ ~A ==> cla(x):A) ==> ?p:B | ~B"; +by (resolve_tac prems 1); +by (resolve_tac [notI] 1); +by (resolve_tac [notE] 1); +by (resolve_tac [notI] 2); +by (eresolve_tac [notE] 2); +by (eresolve_tac [disjI1] 2); +by (resolve_tac [notI] 1); +by (eresolve_tac [notE] 1); +by (eresolve_tac [disjI2] 1); +result(); + + +val prems = +goal IFOLP.thy "[| p:A | ~A; q:~ ~A |] ==> ?p:A"; +by (resolve_tac [disjE] 1); +by (resolve_tac prems 1); +by (assume_tac 1); +by (resolve_tac [FalseE] 1); +by (res_inst_tac [ ("P", "~A") ] notE 1); +by (resolve_tac prems 1); +by (assume_tac 1); +result(); + + +writeln"Examples with quantifiers"; + +val prems = +goal IFOLP.thy "p : ALL z. G(z) ==> ?p:ALL z. G(z)|H(z)"; +by (resolve_tac [allI] 1); +by (resolve_tac [disjI1] 1); +by (resolve_tac (prems RL [spec]) 1); + (*can use instead + by (resolve_tac [spec] 1); by (resolve_tac prems 1); *) +result(); + + +goal IFOLP.thy "?p : ALL x. EX y. x=y"; +by (resolve_tac [allI] 1); +by (resolve_tac [exI] 1); +by (resolve_tac [refl] 1); +result(); + + +goal IFOLP.thy "?p : EX y. ALL x. x=y"; +by (resolve_tac [exI] 1); +by (resolve_tac [allI] 1); +by (resolve_tac [refl] 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + + +(*Parallel lifting example. *) +goal IFOLP.thy "?p : EX u.ALL x.EX v.ALL y.EX w. P(u,x,v,y,w)"; +by (resolve_tac [exI, allI] 1); +by (resolve_tac [exI, allI] 1); +by (resolve_tac [exI, allI] 1); +by (resolve_tac [exI, allI] 1); +by (resolve_tac [exI, allI] 1); + + +val prems = +goal IFOLP.thy "p : (EX z.F(z)) & B ==> ?p:(EX z. F(z) & B)"; +by (resolve_tac [conjE] 1); +by (resolve_tac prems 1); +by (resolve_tac [exE] 1); +by (assume_tac 1); +by (resolve_tac [exI] 1); +by (resolve_tac [conjI] 1); +by (assume_tac 1); +by (assume_tac 1); +result(); + + +(*A bigger demonstration of quantifiers -- not in the paper*) +goal IFOLP.thy "?p : (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"; +by (resolve_tac [impI] 1); +by (resolve_tac [allI] 1); +by (resolve_tac [exE] 1 THEN assume_tac 1); +by (resolve_tac [exI] 1); +by (resolve_tac [allE] 1 THEN assume_tac 1); +by (assume_tac 1); +result(); + + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/if.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/if.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,43 @@ +(* Title: FOLP/ex/if + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For ex/if.thy. First-Order Logic: the 'if' example +*) + +open If; +open Cla; (*in case structure Int is open!*) + +val prems = goalw If.thy [if_def] + "[| !!x.x:P ==> f(x):Q; !!x.x:~P ==> g(x):R |] ==> ?p:if(P,Q,R)"; +by (fast_tac (FOLP_cs addIs prems) 1); +val ifI = result(); + +val major::prems = goalw If.thy [if_def] + "[| p:if(P,Q,R); !!x y.[| x:P; y:Q |] ==> f(x,y):S; \ +\ !!x y.[| x:~P; y:R |] ==> g(x,y):S |] ==> ?p:S"; +by (cut_facts_tac [major] 1); +by (fast_tac (FOLP_cs addIs prems) 1); +val ifE = result(); + + +goal If.thy + "?p : if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))"; +by (resolve_tac [iffI] 1); +by (eresolve_tac [ifE] 1); +by (eresolve_tac [ifE] 1); +by (resolve_tac [ifI] 1); +by (resolve_tac [ifI] 1); + +choplev 0; +val if_cs = FOLP_cs addSIs [ifI] addSEs[ifE]; +by (fast_tac if_cs 1); +val if_commute = result(); + + +goal If.thy "?p : if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))"; +by (fast_tac if_cs 1); +val nested_ifs = result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/if.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/if.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,5 @@ +If = FOLP + +consts if :: "[o,o,o]=>o" +rules + if_def "if(P,Q,R) == P&Q | ~P&R" +end diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/int.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/int.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,361 @@ +(* Title: FOL/ex/int + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Intuitionistic First-Order Logic + +Single-step commands: +by (Int.step_tac 1); +by (biresolve_tac safe_brls 1); +by (biresolve_tac haz_brls 1); +by (assume_tac 1); +by (Int.safe_tac 1); +by (Int.mp_tac 1); +by (Int.fast_tac 1); +*) + +writeln"File FOL/ex/int."; + +(*Note: for PROPOSITIONAL formulae... + ~A is classically provable iff it is intuitionistically provable. + Therefore A is classically provable iff ~~A is intuitionistically provable. + +Let Q be the conjuction of the propositions A|~A, one for each atom A in +P. If P is provable classically, then clearly P&Q is provable +intuitionistically, so ~~(P&Q) is also provable intuitionistically. +The latter is intuitionistically equivalent to ~~P&~~Q, hence to ~~P, +since ~~Q is intuitionistically provable. Finally, if P is a negation then +~~P is intuitionstically equivalent to P. [Andy Pitts] +*) + +goal IFOLP.thy "?p : ~~(P&Q) <-> ~~P & ~~Q"; +by (Int.fast_tac 1); +result(); + +goal IFOLP.thy "?p : ~~~P <-> ~P"; +by (Int.fast_tac 1); +result(); + +goal IFOLP.thy "?p : ~~((P --> Q | R) --> (P-->Q) | (P-->R))"; +by (Int.fast_tac 1); +result(); + +goal IFOLP.thy "?p : (P<->Q) <-> (Q<->P)"; +by (Int.fast_tac 1); +result(); + + +writeln"Lemmas for the propositional double-negation translation"; + +goal IFOLP.thy "?p : P --> ~~P"; +by (Int.fast_tac 1); +result(); + +goal IFOLP.thy "?p : ~~(~~P --> P)"; +by (Int.fast_tac 1); +result(); + +goal IFOLP.thy "?p : ~~P & ~~(P --> Q) --> ~~Q"; +by (Int.fast_tac 1); +result(); + + +writeln"The following are classically but not constructively valid."; + +(*The attempt to prove them terminates quickly!*) +goal IFOLP.thy "?p : ((P-->Q) --> P) --> P"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +(*Check that subgoals remain: proof failed.*) +getgoal 1; + +goal IFOLP.thy "?p : (P&Q-->R) --> (P-->R) | (Q-->R)"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + + +writeln"Intuitionistic FOL: propositional problems based on Pelletier."; + +writeln"Problem ~~1"; +goal IFOLP.thy "?p : ~~((P-->Q) <-> (~Q --> ~P))"; +by (Int.fast_tac 1); +result(); +(*5 secs*) + + +writeln"Problem ~~2"; +goal IFOLP.thy "?p : ~~(~~P <-> P)"; +by (Int.fast_tac 1); +result(); +(*1 secs*) + + +writeln"Problem 3"; +goal IFOLP.thy "?p : ~(P-->Q) --> (Q-->P)"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~4"; +goal IFOLP.thy "?p : ~~((~P-->Q) <-> (~Q --> P))"; +by (Int.fast_tac 1); +result(); +(*9 secs*) + +writeln"Problem ~~5"; +goal IFOLP.thy "?p : ~~((P|Q-->P|R) --> P|(Q-->R))"; +by (Int.fast_tac 1); +result(); +(*10 secs*) + + +writeln"Problem ~~6"; +goal IFOLP.thy "?p : ~~(P | ~P)"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~7"; +goal IFOLP.thy "?p : ~~(P | ~~~P)"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~8. Peirce's law"; +goal IFOLP.thy "?p : ~~(((P-->Q) --> P) --> P)"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 9"; +goal IFOLP.thy "?p : ((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"; +by (Int.fast_tac 1); +result(); +(*9 secs*) + + +writeln"Problem 10"; +goal IFOLP.thy "?p : (Q-->R) --> (R-->P&Q) --> (P-->(Q|R)) --> (P<->Q)"; +by (Int.fast_tac 1); +result(); + +writeln"11. Proved in each direction (incorrectly, says Pelletier!!) "; +goal IFOLP.thy "?p : P<->P"; +by (Int.fast_tac 1); + +writeln"Problem ~~12. Dijkstra's law "; +goal IFOLP.thy "?p : ~~(((P <-> Q) <-> R) <-> (P <-> (Q <-> R)))"; +by (Int.fast_tac 1); +result(); + +goal IFOLP.thy "?p : ((P <-> Q) <-> R) --> ~~(P <-> (Q <-> R))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 13. Distributive law"; +goal IFOLP.thy "?p : P | (Q & R) <-> (P | Q) & (P | R)"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~14"; +goal IFOLP.thy "?p : ~~((P <-> Q) <-> ((Q | ~P) & (~Q|P)))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~15"; +goal IFOLP.thy "?p : ~~((P --> Q) <-> (~P | Q))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~16"; +goal IFOLP.thy "?p : ~~((P-->Q) | (Q-->P))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~17"; +goal IFOLP.thy + "?p : ~~(((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S)))"; +by (Int.fast_tac 1); +(*over 5 minutes?? -- printing the proof term takes 40 secs!!*) +result(); + + +writeln"** Examples with quantifiers **"; + +writeln"The converse is classical in the following implications..."; + +goal IFOLP.thy "?p : (EX x.P(x)-->Q) --> (ALL x.P(x)) --> Q"; +by (Int.fast_tac 1); +result(); + +goal IFOLP.thy "?p : ((ALL x.P(x))-->Q) --> ~ (ALL x. P(x) & ~Q)"; +by (Int.fast_tac 1); +result(); + +goal IFOLP.thy "?p : ((ALL x. ~P(x))-->Q) --> ~ (ALL x. ~ (P(x)|Q))"; +by (Int.fast_tac 1); +result(); + +goal IFOLP.thy "?p : (ALL x.P(x)) | Q --> (ALL x. P(x) | Q)"; +by (Int.fast_tac 1); +result(); + +goal IFOLP.thy "?p : (EX x. P --> Q(x)) --> (P --> (EX x. Q(x)))"; +by (Int.fast_tac 1); +result(); + + + + +writeln"The following are not constructively valid!"; +(*The attempt to prove them terminates quickly!*) + +goal IFOLP.thy "?p : ((ALL x.P(x))-->Q) --> (EX x.P(x)-->Q)"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + +goal IFOLP.thy "?p : (P --> (EX x.Q(x))) --> (EX x. P-->Q(x))"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + +goal IFOLP.thy "?p : (ALL x. P(x) | Q) --> ((ALL x.P(x)) | Q)"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + +goal IFOLP.thy "?p : (ALL x. ~~P(x)) --> ~~(ALL x. P(x))"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + +(*Classically but not intuitionistically valid. Proved by a bug in 1986!*) +goal IFOLP.thy "?p : EX x. Q(x) --> (ALL x. Q(x))"; +by (Int.fast_tac 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + + +writeln"Hard examples with quantifiers"; + +(*The ones that have not been proved are not known to be valid! + Some will require quantifier duplication -- not currently available*) + +writeln"Problem ~~18"; +goal IFOLP.thy "?p : ~~(EX y. ALL x. P(y)-->P(x))"; +(*NOT PROVED*) + +writeln"Problem ~~19"; +goal IFOLP.thy "?p : ~~(EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x)))"; +(*NOT PROVED*) + +writeln"Problem 20"; +goal IFOLP.thy "?p : (ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w))) \ +\ --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 21"; +goal IFOLP.thy + "?p : (EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> ~~(EX x. P<->Q(x))"; +(*NOT PROVED*) + +writeln"Problem 22"; +goal IFOLP.thy "?p : (ALL x. P <-> Q(x)) --> (P <-> (ALL x. Q(x)))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~23"; +goal IFOLP.thy "?p : ~~ ((ALL x. P | Q(x)) <-> (P | (ALL x. Q(x))))"; +by (Int.best_tac 1); +result(); + +writeln"Problem 24"; +goal IFOLP.thy "?p : ~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) & \ +\ ~(EX x.P(x)) --> (EX x.Q(x)) & (ALL x. Q(x)|R(x) --> S(x)) \ +\ --> (EX x. P(x)&R(x))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 25"; +goal IFOLP.thy "?p : (EX x. P(x)) & \ +\ (ALL x. L(x) --> ~ (M(x) & R(x))) & \ +\ (ALL x. P(x) --> (M(x) & L(x))) & \ +\ ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x))) \ +\ --> (EX x. Q(x)&P(x))"; +by (Int.best_tac 1); +result(); + +writeln"Problem 29. Essentially the same as Principia Mathematica *11.71"; +goal IFOLP.thy "?p : (EX x. P(x)) & (EX y. Q(y)) \ +\ --> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y)) <-> \ +\ (ALL x y. P(x) & Q(y) --> R(x) & S(y)))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem ~~30"; +goal IFOLP.thy "?p : (ALL x. (P(x) | Q(x)) --> ~ R(x)) & \ +\ (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x)) \ +\ --> (ALL x. ~~S(x))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 31"; +goal IFOLP.thy "?p : ~(EX x.P(x) & (Q(x) | R(x))) & \ +\ (EX x. L(x) & P(x)) & \ +\ (ALL x. ~ R(x) --> M(x)) \ +\ --> (EX x. L(x) & M(x))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 32"; +goal IFOLP.thy "?p : (ALL x. P(x) & (Q(x)|R(x))-->S(x)) & \ +\ (ALL x. S(x) & R(x) --> L(x)) & \ +\ (ALL x. M(x) --> R(x)) \ +\ --> (ALL x. P(x) & M(x) --> L(x))"; +by (Int.best_tac 1); (*SLOW*) +result(); + +writeln"Problem 39"; +goal IFOLP.thy "?p : ~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 40. AMENDED"; +goal IFOLP.thy "?p : (EX y. ALL x. F(x,y) <-> F(x,x)) --> \ +\ ~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 44"; +goal IFOLP.thy "?p : (ALL x. f(x) --> \ +\ (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y)))) & \ +\ (EX x. j(x) & (ALL y. g(y) --> h(x,y))) \ +\ --> (EX x. j(x) & ~f(x))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 48"; +goal IFOLP.thy "?p : (a=b | c=d) & (a=c | b=d) --> a=d | b=c"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 51"; +goal IFOLP.thy + "?p : (EX z w. ALL x y. P(x,y) <-> (x=z & y=w)) --> \ +\ (EX z. ALL x. EX w. (ALL y. P(x,y) <-> y=w) <-> x=z)"; +by (Int.best_tac 1); (*60 seconds*) +result(); + +writeln"Problem 56"; +goal IFOLP.thy + "?p : (ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 57"; +goal IFOLP.thy + "?p : P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) & \ +\ (ALL x y z. P(x,y) & P(y,z) --> P(x,z)) --> P(f(a,b), f(a,c))"; +by (Int.fast_tac 1); +result(); + +writeln"Problem 60"; +goal IFOLP.thy + "?p : ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"; +by (Int.fast_tac 1); +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/intro.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/intro.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,92 @@ +(* Title: FOL/ex/intro + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Examples for the manual "Introduction to Isabelle" + +Derives some inference rules, illustrating the use of definitions + +To generate similar output to manual, execute these commands: + Pretty.setmargin 72; print_depth 0; +*) + + +(**** Some simple backward proofs ****) + +goal FOLP.thy "?p:P|P --> P"; +by (resolve_tac [impI] 1); +by (resolve_tac [disjE] 1); +by (assume_tac 3); +by (assume_tac 2); +by (assume_tac 1); +val mythm = result(); + +goal FOLP.thy "?p:(P & Q) | R --> (P | R)"; +by (resolve_tac [impI] 1); +by (eresolve_tac [disjE] 1); +by (dresolve_tac [conjunct1] 1); +by (resolve_tac [disjI1] 1); +by (resolve_tac [disjI2] 2); +by (REPEAT (assume_tac 1)); +result(); + +(*Correct version, delaying use of "spec" until last*) +goal FOLP.thy "?p:(ALL x y.P(x,y)) --> (ALL z w.P(w,z))"; +by (resolve_tac [impI] 1); +by (resolve_tac [allI] 1); +by (resolve_tac [allI] 1); +by (dresolve_tac [spec] 1); +by (dresolve_tac [spec] 1); +by (assume_tac 1); +result(); + + +(**** Demonstration of fast_tac ****) + +goal FOLP.thy "?p:(EX y. ALL x. J(y,x) <-> ~J(x,x)) \ +\ --> ~ (ALL x. EX y. ALL z. J(z,y) <-> ~ J(z,x))"; +by (fast_tac FOLP_cs 1); +result(); + +goal FOLP.thy "?p:ALL x. P(x,f(x)) <-> \ +\ (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"; +by (fast_tac FOLP_cs 1); +result(); + + +(**** Derivation of conjunction elimination rule ****) + +val [major,minor] = goal FOLP.thy "[| p:P&Q; !!x y.[| x:P; y:Q |] ==>f(x,y):R |] ==> ?p:R"; +by (resolve_tac [minor] 1); +by (resolve_tac [major RS conjunct1] 1); +by (resolve_tac [major RS conjunct2] 1); +prth (topthm()); +result(); + + +(**** Derived rules involving definitions ****) + +(** Derivation of negation introduction **) + +val prems = goal FOLP.thy "(!!x.x:P ==> f(x):False) ==> ?p:~P"; +by (rewrite_goals_tac [not_def]); +by (resolve_tac [impI] 1); +by (resolve_tac prems 1); +by (assume_tac 1); +result(); + +val [major,minor] = goal FOLP.thy "[| p:~P; q:P |] ==> ?p:R"; +by (resolve_tac [FalseE] 1); +by (resolve_tac [mp] 1); +by (resolve_tac [rewrite_rule [not_def] major] 1); +by (resolve_tac [minor] 1); +result(); + +(*Alternative proof of above result*) +val [major,minor] = goalw FOLP.thy [not_def] + "[| p:~P; q:P |] ==> ?p:R"; +by (resolve_tac [minor RS (major RS mp RS FalseE)] 1); +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/nat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/nat.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,82 @@ +(* Title: FOLP/ex/nat.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Examples for the manual "Introduction to Isabelle" + +Proofs about the natural numbers + +To generate similar output to manual, execute these commands: + Pretty.setmargin 72; print_depth 0; +*) + +open Nat; + +goal Nat.thy "?p : ~ (Suc(k) = k)"; +by (res_inst_tac [("n","k")] induct 1); +by (rtac notI 1); +by (etac Suc_neq_0 1); +by (rtac notI 1); +by (etac notE 1); +by (etac Suc_inject 1); +val Suc_n_not_n = result(); + + +goal Nat.thy "?p : (k+m)+n = k+(m+n)"; +prths ([induct] RL [topthm()]); (*prints all 14 next states!*) +by (rtac induct 1); +back(); +back(); +back(); +back(); +back(); +back(); + +goalw Nat.thy [add_def] "?p : 0+n = n"; +by (rtac rec_0 1); +val add_0 = result(); + +goalw Nat.thy [add_def] "?p : Suc(m)+n = Suc(m+n)"; +by (rtac rec_Suc 1); +val add_Suc = result(); + +(* +val nat_congs = mk_congs Nat.thy ["Suc", "op +"]; +prths nat_congs; +*) +val prems = goal Nat.thy "p: x=y ==> ?p : Suc(x) = Suc(y)"; +by (resolve_tac (prems RL [subst]) 1); +by (rtac refl 1); +val Suc_cong = result(); + +val prems = goal Nat.thy "[| p : a=x; q: b=y |] ==> ?p : a+b=x+y"; +by (resolve_tac (prems RL [subst]) 1 THEN resolve_tac (prems RL [subst]) 1 THEN + rtac refl 1); +val Plus_cong = result(); + +val nat_congs = [Suc_cong,Plus_cong]; + + +val add_ss = FOLP_ss addcongs nat_congs + addrews [add_0, add_Suc]; + +goal Nat.thy "?p : (k+m)+n = k+(m+n)"; +by (res_inst_tac [("n","k")] induct 1); +by (SIMP_TAC add_ss 1); +by (ASM_SIMP_TAC add_ss 1); +val add_assoc = result(); + +goal Nat.thy "?p : m+0 = m"; +by (res_inst_tac [("n","m")] induct 1); +by (SIMP_TAC add_ss 1); +by (ASM_SIMP_TAC add_ss 1); +val add_0_right = result(); + +goal Nat.thy "?p : m+Suc(n) = Suc(m+n)"; +by (res_inst_tac [("n","m")] induct 1); +by (ALLGOALS (ASM_SIMP_TAC add_ss)); +val add_Suc_right = result(); + +(*mk_typed_congs appears not to work with FOLP's version of subst*) + diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/nat.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,36 @@ +(* Title: FOLP/ex/nat.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Examples for the manual "Introduction to Isabelle" + +Theory of the natural numbers: Peano's axioms, primitive recursion +*) + +Nat = IFOLP + +types nat 0 +arities nat :: term +consts "0" :: "nat" ("0") + Suc :: "nat=>nat" + rec :: "[nat, 'a, [nat,'a]=>'a] => 'a" + "+" :: "[nat, nat] => nat" (infixl 60) + + (*Proof terms*) + nrec :: "[nat,p,[nat,p]=>p]=>p" + ninj,nneq :: "p=>p" + rec0, recSuc :: "p" + +rules + induct "[| b:P(0); !!x u. u:P(x) ==> c(x,u):P(Suc(x)) \ +\ |] ==> nrec(n,b,c):P(n)" + + Suc_inject "p:Suc(m)=Suc(n) ==> ninj(p) : m=n" + Suc_neq_0 "p:Suc(m)=0 ==> nneq(p) : R" + rec_0 "rec0 : rec(0,a,f) = a" + rec_Suc "recSuc : rec(Suc(m), a, f) = f(m, rec(m,a,f))" + add_def "m+n == rec(m, n, %x y. Suc(y))" + + nrecB0 "b: A ==> nrec(0,b,c) = b : A" + nrecBSuc "c(n,nrec(n,b,c)) : A ==> nrec(Suc(n),b,c) = c(n,nrec(n,b,c)) : A" +end diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/prolog.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/prolog.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,74 @@ +(* Title: FOL/ex/prolog.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For ex/prolog.thy. First-Order Logic: PROLOG examples +*) + +open Prolog; + +goal Prolog.thy "app(a:b:c:Nil, d:e:Nil, ?x)"; +by (resolve_tac [appNil,appCons] 1); +by (resolve_tac [appNil,appCons] 1); +by (resolve_tac [appNil,appCons] 1); +by (resolve_tac [appNil,appCons] 1); +prth (result()); + +goal Prolog.thy "app(?x, c:d:Nil, a:b:c:d:Nil)"; +by (REPEAT (resolve_tac [appNil,appCons] 1)); +result(); + + +goal Prolog.thy "app(?x, ?y, a:b:c:d:Nil)"; +by (REPEAT (resolve_tac [appNil,appCons] 1)); +back(); +back(); +back(); +back(); +result(); + + +(*app([x1,...,xn], y, ?z) requires (n+1) inferences*) +(*rev([x1,...,xn], ?y) requires (n+1)(n+2)/2 inferences*) + +goal Prolog.thy "rev(a:b:c:d:Nil, ?x)"; +val rules = [appNil,appCons,revNil,revCons]; +by (REPEAT (resolve_tac rules 1)); +result(); + +goal Prolog.thy "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:Nil, ?w)"; +by (REPEAT (resolve_tac rules 1)); +result(); + +goal Prolog.thy "rev(?x, a:b:c:Nil)"; +by (REPEAT (resolve_tac rules 1)); (*does not solve it directly!*) +back(); +back(); + +(*backtracking version*) +val prolog_tac = DEPTH_FIRST (has_fewer_prems 1) (resolve_tac rules 1); + +choplev 0; +by prolog_tac; +result(); + +goal Prolog.thy "rev(a:?x:c:?y:Nil, d:?z:b:?u)"; +by prolog_tac; +result(); + +(*rev([a..p], ?w) requires 153 inferences *) +goal Prolog.thy "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil, ?w)"; +by (DEPTH_SOLVE (resolve_tac ([refl,conjI]@rules) 1)); +(*Poly/ML: 4 secs >> 38 lips*) +result(); + +(*?x has 16, ?y has 32; rev(?y,?w) requires 561 (rather large) inferences; + total inferences = 2 + 1 + 17 + 561 = 581*) +goal Prolog.thy + "a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil = ?x & app(?x,?x,?y) & rev(?y,?w)"; +by (DEPTH_SOLVE (resolve_tac ([refl,conjI]@rules) 1)); +(*Poly/ML: 29 secs >> 20 lips*) +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/prolog.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/prolog.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,22 @@ +(* Title: FOL/ex/prolog.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +First-Order Logic: PROLOG examples + +Inherits from FOL the class term, the type o, and the coercion Trueprop +*) + +Prolog = FOL + +types list 1 +arities list :: (term)term +consts Nil :: "'a list" + ":" :: "['a, 'a list]=> 'a list" (infixr 60) + app :: "['a list, 'a list, 'a list] => o" + rev :: "['a list, 'a list] => o" +rules appNil "app(Nil,ys,ys)" + appCons "app(xs,ys,zs) ==> app(x:xs, ys, x:zs)" + revNil "rev(Nil,Nil)" + revCons "[| rev(xs,ys); app(ys, x:Nil, zs) |] ==> rev(x:xs, zs)" +end diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/prop.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/prop.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,153 @@ +(* Title: FOL/ex/prop + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +First-Order Logic: propositional examples (intuitionistic and classical) +Needs declarations of the theory "thy" and the tactic "tac" +*) + +writeln"File FOL/ex/prop."; + + +writeln"commutative laws of & and | "; +goal thy "?p : P & Q --> Q & P"; +by tac; +result(); + +goal thy "?p : P | Q --> Q | P"; +by tac; +result(); + + +writeln"associative laws of & and | "; +goal thy "?p : (P & Q) & R --> P & (Q & R)"; +by tac; +result(); + +goal thy "?p : (P | Q) | R --> P | (Q | R)"; +by tac; +result(); + + + +writeln"distributive laws of & and | "; +goal thy "?p : (P & Q) | R --> (P | R) & (Q | R)"; +by tac; +result(); + +goal thy "?p : (P | R) & (Q | R) --> (P & Q) | R"; +by tac; +result(); + +goal thy "?p : (P | Q) & R --> (P & R) | (Q & R)"; +by tac; +result(); + + +goal thy "?p : (P & R) | (Q & R) --> (P | Q) & R"; +by tac; +result(); + + +writeln"Laws involving implication"; + +goal thy "?p : (P-->R) & (Q-->R) <-> (P|Q --> R)"; +by tac; +result(); + + +goal thy "?p : (P & Q --> R) <-> (P--> (Q-->R))"; +by tac; +result(); + + +goal thy "?p : ((P-->R)-->R) --> ((Q-->R)-->R) --> (P&Q-->R) --> R"; +by tac; +result(); + +goal thy "?p : ~(P-->R) --> ~(Q-->R) --> ~(P&Q-->R)"; +by tac; +result(); + +goal thy "?p : (P --> Q & R) <-> (P-->Q) & (P-->R)"; +by tac; +result(); + + +writeln"Propositions-as-types"; + +(*The combinator K*) +goal thy "?p : P --> (Q --> P)"; +by tac; +result(); + +(*The combinator S*) +goal thy "?p : (P-->Q-->R) --> (P-->Q) --> (P-->R)"; +by tac; +result(); + + +(*Converse is classical*) +goal thy "?p : (P-->Q) | (P-->R) --> (P --> Q | R)"; +by tac; +result(); + +goal thy "?p : (P-->Q) --> (~Q --> ~P)"; +by tac; +result(); + + +writeln"Schwichtenberg's examples (via T. Nipkow)"; + +(* stab-imp *) +goal thy "?p : (((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"; +by tac; +result(); + +(* stab-to-peirce *) +goal thy "?p : (((P --> R) --> R) --> P) --> (((Q --> R) --> R) --> Q) \ +\ --> ((P --> Q) --> P) --> P"; +by tac; +result(); + +(* peirce-imp1 *) +goal thy "?p : (((Q --> R) --> Q) --> Q) \ +\ --> (((P --> Q) --> R) --> P --> Q) --> P --> Q"; +by tac; +result(); + +(* peirce-imp2 *) +goal thy "?p : (((P --> R) --> P) --> P) --> ((P --> Q --> R) --> P) --> P"; +by tac; +result(); + +(* mints *) +goal thy "?p : ((((P --> Q) --> P) --> P) --> Q) --> Q"; +by tac; +result(); + +(* mints-solovev *) +goal thy "?p : (P --> (Q --> R) --> Q) --> ((P --> Q) --> R) --> R"; +by tac; +result(); + +(* tatsuta *) +goal thy "?p : (((P7 --> P1) --> P10) --> P4 --> P5) \ +\ --> (((P8 --> P2) --> P9) --> P3 --> P10) \ +\ --> (P1 --> P8) --> P6 --> P7 \ +\ --> (((P3 --> P2) --> P9) --> P4) \ +\ --> (P1 --> P3) --> (((P6 --> P1) --> P2) --> P9) --> P5"; +by tac; +result(); + +(* tatsuta1 *) +goal thy "?p : (((P8 --> P2) --> P9) --> P3 --> P10) \ +\ --> (((P3 --> P2) --> P9) --> P4) \ +\ --> (((P6 --> P1) --> P2) --> P9) \ +\ --> (((P7 --> P1) --> P10) --> P4 --> P5) \ +\ --> (P1 --> P3) --> (P1 --> P8) --> P6 --> P7 --> P5"; +by tac; +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/FOLP/ex/quant.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ex/quant.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,129 @@ +(* Title: FOL/ex/quant + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +First-Order Logic: quantifier examples (intuitionistic and classical) +Needs declarations of the theory "thy" and the tactic "tac" +*) + +writeln"File FOL/ex/quant."; + +goal thy "?p : (ALL x y.P(x,y)) --> (ALL y x.P(x,y))"; +by tac; +result(); + + +goal thy "?p : (EX x y.P(x,y)) --> (EX y x.P(x,y))"; +by tac; +result(); + + +(*Converse is false*) +goal thy "?p : (ALL x.P(x)) | (ALL x.Q(x)) --> (ALL x. P(x) | Q(x))"; +by tac; +result(); + +goal thy "?p : (ALL x. P-->Q(x)) <-> (P--> (ALL x.Q(x)))"; +by tac; +result(); + + +goal thy "?p : (ALL x.P(x)-->Q) <-> ((EX x.P(x)) --> Q)"; +by tac; +result(); + + +writeln"Some harder ones"; + +goal thy "?p : (EX x. P(x) | Q(x)) <-> (EX x.P(x)) | (EX x.Q(x))"; +by tac; +result(); +(*6 secs*) + +(*Converse is false*) +goal thy "?p : (EX x. P(x)&Q(x)) --> (EX x.P(x)) & (EX x.Q(x))"; +by tac; +result(); + + +writeln"Basic test of quantifier reasoning"; +(*TRUE*) +goal thy "?p : (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"; +by tac; +result(); + + +goal thy "?p : (ALL x. Q(x)) --> (EX x. Q(x))"; +by tac; +result(); + + +writeln"The following should fail, as they are false!"; + +goal thy "?p : (ALL x. EX y. Q(x,y)) --> (EX y. ALL x. Q(x,y))"; +by tac handle ERROR => writeln"Failed, as expected"; +(*Check that subgoals remain: proof failed.*) +getgoal 1; + +goal thy "?p : (EX x. Q(x)) --> (ALL x. Q(x))"; +by tac handle ERROR => writeln"Failed, as expected"; +getgoal 1; + +goal thy "?p : P(?a) --> (ALL x.P(x))"; +by tac handle ERROR => writeln"Failed, as expected"; +(*Check that subgoals remain: proof failed.*) +getgoal 1; + +goal thy + "?p : (P(?a) --> (ALL x.Q(x))) --> (ALL x. P(x) --> Q(x))"; +by tac handle ERROR => writeln"Failed, as expected"; +getgoal 1; + + +writeln"Back to things that are provable..."; + +goal thy "?p : (ALL x.P(x)-->Q(x)) & (EX x.P(x)) --> (EX x.Q(x))"; +by tac; +result(); + + +(*An example of why exI should be delayed as long as possible*) +goal thy "?p : (P --> (EX x.Q(x))) & P --> (EX x.Q(x))"; +by tac; +result(); + +goal thy "?p : (ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"; +by tac; +(*Verify that no subgoals remain.*) +uresult(); + + +goal thy "?p : (ALL x. Q(x)) --> (EX x. Q(x))"; +by tac; +result(); + + +writeln"Some slow ones"; + + +(*Principia Mathematica *11.53 *) +goal thy "?p : (ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))"; +by tac; +result(); +(*6 secs*) + +(*Principia Mathematica *11.55 *) +goal thy "?p : (EX x y. P(x) & Q(x,y)) <-> (EX x. P(x) & (EX y. Q(x,y)))"; +by tac; +result(); +(*9 secs*) + +(*Principia Mathematica *11.61 *) +goal thy "?p : (EX y. ALL x. P(x) --> Q(x,y)) --> (ALL x. P(x) --> (EX y. Q(x,y)))"; +by tac; +result(); +(*3 secs*) + +writeln"Reached end of file."; + diff -r 000000000000 -r a5a9c433f639 src/FOLP/folp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/folp.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,95 @@ +(* Title: FOL/fol.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Tactics and lemmas for fol.thy (classical First-Order Logic) +*) + +open FOLP; + +signature FOLP_LEMMAS = + sig + val disjCI : thm + val excluded_middle : thm + val exCI : thm + val ex_classical : thm + val iffCE : thm + val impCE : thm + val notnotD : thm + val swap : thm + end; + + +structure FOLP_Lemmas : FOLP_LEMMAS = +struct + +(*** Classical introduction rules for | and EX ***) + +val disjCI = prove_goal FOLP.thy + "(!!x.x:~Q ==> f(x):P) ==> ?p : P|Q" + (fn prems=> + [ (resolve_tac [classical] 1), + (REPEAT (ares_tac (prems@[disjI1,notI]) 1)), + (REPEAT (ares_tac (prems@[disjI2,notE]) 1)) ]); + +(*introduction rule involving only EX*) +val ex_classical = prove_goal FOLP.thy + "( !!u.u:~(EX x. P(x)) ==> f(u):P(a)) ==> ?p : EX x.P(x)" + (fn prems=> + [ (resolve_tac [classical] 1), + (eresolve_tac (prems RL [exI]) 1) ]); + +(*version of above, simplifying ~EX to ALL~ *) +val exCI = prove_goal FOLP.thy + "(!!u.u:ALL x. ~P(x) ==> f(u):P(a)) ==> ?p : EX x.P(x)" + (fn [prem]=> + [ (resolve_tac [ex_classical] 1), + (resolve_tac [notI RS allI RS prem] 1), + (eresolve_tac [notE] 1), + (eresolve_tac [exI] 1) ]); + +val excluded_middle = prove_goal FOLP.thy "?p : ~P | P" + (fn _=> [ rtac disjCI 1, assume_tac 1 ]); + + +(*** Special elimination rules *) + + +(*Classical implies (-->) elimination. *) +val impCE = prove_goal FOLP.thy + "[| p:P-->Q; !!x.x:~P ==> f(x):R; !!y.y:Q ==> g(y):R |] ==> ?p : R" + (fn major::prems=> + [ (resolve_tac [excluded_middle RS disjE] 1), + (DEPTH_SOLVE (ares_tac (prems@[major RS mp]) 1)) ]); + +(*Double negation law*) +val notnotD = prove_goal FOLP.thy "p:~~P ==> ?p : P" + (fn [major]=> + [ (resolve_tac [classical] 1), (eresolve_tac [major RS notE] 1) ]); + + +(*** Tactics for implication and contradiction ***) + +(*Classical <-> elimination. Proof substitutes P=Q in + ~P ==> ~Q and P ==> Q *) +val iffCE = prove_goalw FOLP.thy [iff_def] + "[| p:P<->Q; !!x y.[| x:P; y:Q |] ==> f(x,y):R; \ +\ !!x y.[| x:~P; y:~Q |] ==> g(x,y):R |] ==> ?p : R" + (fn prems => + [ (resolve_tac [conjE] 1), + (REPEAT (DEPTH_SOLVE_1 + (etac impCE 1 ORELSE mp_tac 1 ORELSE ares_tac prems 1))) ]); + + +(*Should be used as swap since ~P becomes redundant*) +val swap = prove_goal FOLP.thy + "p:~P ==> (!!x.x:~Q ==> f(x):P) ==> ?p : Q" + (fn major::prems=> + [ (resolve_tac [classical] 1), + (rtac (major RS notE) 1), + (REPEAT (ares_tac prems 1)) ]); + +end; + +open FOLP_Lemmas; diff -r 000000000000 -r a5a9c433f639 src/FOLP/folp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/folp.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,6 @@ +FOLP = IFOLP + +consts + cla :: "[p=>p]=>p" +rules + classical "(!!x.x:~P ==> f(x):P) ==> cla(f):P" +end diff -r 000000000000 -r a5a9c433f639 src/FOLP/ifolp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ifolp.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,444 @@ +(* Title: FOLP/ifol.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Tactics and lemmas for ifol.thy (intuitionistic first-order logic) +*) + +open IFOLP; + +signature IFOLP_LEMMAS = + sig + val allE: thm + val all_cong: thm + val all_dupE: thm + val all_impE: thm + val box_equals: thm + val conjE: thm + val conj_cong: thm + val conj_impE: thm + val contrapos: thm + val disj_cong: thm + val disj_impE: thm + val eq_cong: thm + val ex1I: thm + val ex1E: thm + val ex1_equalsE: thm +(* val ex1_cong: thm****) + val ex_cong: thm + val ex_impE: thm + val iffD1: thm + val iffD2: thm + val iffE: thm + val iffI: thm + val iff_cong: thm + val iff_impE: thm + val iff_refl: thm + val iff_sym: thm + val iff_trans: thm + val impE: thm + val imp_cong: thm + val imp_impE: thm + val mp_tac: int -> tactic + val notE: thm + val notI: thm + val not_cong: thm + val not_impE: thm + val not_sym: thm + val not_to_imp: thm + val pred1_cong: thm + val pred2_cong: thm + val pred3_cong: thm + val pred_congs: thm list + val refl: thm + val rev_mp: thm + val simp_equals: thm + val subst: thm + val ssubst: thm + val subst_context: thm + val subst_context2: thm + val subst_context3: thm + val sym: thm + val trans: thm + val TrueI: thm + val uniq_assume_tac: int -> tactic + val uniq_mp_tac: int -> tactic + end; + + +structure IFOLP_Lemmas : IFOLP_LEMMAS = +struct + +val TrueI = TrueI; + +(*** Sequent-style elimination rules for & --> and ALL ***) + +val conjE = prove_goal IFOLP.thy + "[| p:P&Q; !!x y.[| x:P; y:Q |] ==> f(x,y):R |] ==> ?a:R" + (fn prems=> + [ (REPEAT (resolve_tac prems 1 + ORELSE (resolve_tac [conjunct1, conjunct2] 1 THEN + resolve_tac prems 1))) ]); + +val impE = prove_goal IFOLP.thy + "[| p:P-->Q; q:P; !!x.x:Q ==> r(x):R |] ==> ?p:R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); + +val allE = prove_goal IFOLP.thy + "[| p:ALL x.P(x); !!y.y:P(x) ==> q(y):R |] ==> ?p:R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[spec]) 1)) ]); + +(*Duplicates the quantifier; for use with eresolve_tac*) +val all_dupE = prove_goal IFOLP.thy + "[| p:ALL x.P(x); !!y z.[| y:P(x); z:ALL x.P(x) |] ==> q(y,z):R \ +\ |] ==> ?p:R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[spec]) 1)) ]); + + +(*** Negation rules, which translate between ~P and P-->False ***) + +val notI = prove_goalw IFOLP.thy [not_def] "(!!x.x:P ==> q(x):False) ==> ?p:~P" + (fn prems=> [ (REPEAT (ares_tac (prems@[impI]) 1)) ]); + +val notE = prove_goalw IFOLP.thy [not_def] "[| p:~P; q:P |] ==> ?p:R" + (fn prems=> + [ (resolve_tac [mp RS FalseE] 1), + (REPEAT (resolve_tac prems 1)) ]); + +(*This is useful with the special implication rules for each kind of P. *) +val not_to_imp = prove_goal IFOLP.thy + "[| p:~P; !!x.x:(P-->False) ==> q(x):Q |] ==> ?p:Q" + (fn prems=> [ (REPEAT (ares_tac (prems@[impI,notE]) 1)) ]); + + +(* For substitution int an assumption P, reduce Q to P-->Q, substitute into + this implication, then apply impI to move P back into the assumptions. + To specify P use something like + eres_inst_tac [ ("P","ALL y. ?S(x,y)") ] rev_mp 1 *) +val rev_mp = prove_goal IFOLP.thy "[| p:P; q:P --> Q |] ==> ?p:Q" + (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); + + +(*Contrapositive of an inference rule*) +val contrapos = prove_goal IFOLP.thy "[| p:~Q; !!y.y:P==>q(y):Q |] ==> ?a:~P" + (fn [major,minor]=> + [ (rtac (major RS notE RS notI) 1), + (etac minor 1) ]); + +(** Unique assumption tactic. + Ignores proof objects. + Fails unless one assumption is equal and exactly one is unifiable +**) + +local + fun discard_proof (Const("Proof",_) $ P $ _) = P; +in +val uniq_assume_tac = + SUBGOAL + (fn (prem,i) => + let val hyps = map discard_proof (Logic.strip_assums_hyp prem) + and concl = discard_proof (Logic.strip_assums_concl prem) + in + if exists (fn hyp => hyp aconv concl) hyps + then case distinct (filter (fn hyp=> could_unify(hyp,concl)) hyps) of + [_] => assume_tac i + | _ => no_tac + else no_tac + end); +end; + + +(*** Modus Ponens Tactics ***) + +(*Finds P-->Q and P in the assumptions, replaces implication by Q *) +fun mp_tac i = eresolve_tac [notE,make_elim mp] i THEN assume_tac i; + +(*Like mp_tac but instantiates no variables*) +fun uniq_mp_tac i = eresolve_tac [notE,impE] i THEN uniq_assume_tac i; + + +(*** If-and-only-if ***) + +val iffI = prove_goalw IFOLP.thy [iff_def] + "[| !!x.x:P ==> q(x):Q; !!x.x:Q ==> r(x):P |] ==> ?p:P<->Q" + (fn prems=> [ (REPEAT (ares_tac (prems@[conjI, impI]) 1)) ]); + + +(*Observe use of rewrite_rule to unfold "<->" in meta-assumptions (prems) *) +val iffE = prove_goalw IFOLP.thy [iff_def] + "[| p:P <-> Q; !!x y.[| x:P-->Q; y:Q-->P |] ==> q(x,y):R |] ==> ?p:R" + (fn prems => [ (resolve_tac [conjE] 1), (REPEAT (ares_tac prems 1)) ]); + +(* Destruct rules for <-> similar to Modus Ponens *) + +val iffD1 = prove_goalw IFOLP.thy [iff_def] "[| p:P <-> Q; q:P |] ==> ?p:Q" + (fn prems => [ (rtac (conjunct1 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); + +val iffD2 = prove_goalw IFOLP.thy [iff_def] "[| p:P <-> Q; q:Q |] ==> ?p:P" + (fn prems => [ (rtac (conjunct2 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); + +val iff_refl = prove_goal IFOLP.thy "?p:P <-> P" + (fn _ => [ (REPEAT (ares_tac [iffI] 1)) ]); + +val iff_sym = prove_goal IFOLP.thy "p:Q <-> P ==> ?p:P <-> Q" + (fn [major] => + [ (rtac (major RS iffE) 1), + (rtac iffI 1), + (REPEAT (eresolve_tac [asm_rl,mp] 1)) ]); + +val iff_trans = prove_goal IFOLP.thy "[| p:P <-> Q; q:Q<-> R |] ==> ?p:P <-> R" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (REPEAT (eresolve_tac [asm_rl,iffE] 1 ORELSE mp_tac 1)) ]); + + +(*** Unique existence. NOTE THAT the following 2 quantifications + EX!x such that [EX!y such that P(x,y)] (sequential) + EX!x,y such that P(x,y) (simultaneous) + do NOT mean the same thing. The parser treats EX!x y.P(x,y) as sequential. +***) + +val ex1I = prove_goalw IFOLP.thy [ex1_def] + "[| p:P(a); !!x u.u:P(x) ==> f(u) : x=a |] ==> ?p:EX! x. P(x)" + (fn prems => [ (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ]); + +val ex1E = prove_goalw IFOLP.thy [ex1_def] + "[| p:EX! x.P(x); \ +\ !!x u v. [| u:P(x); v:ALL y. P(y) --> y=x |] ==> f(x,u,v):R |] ==>\ +\ ?a : R" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (eresolve_tac [exE,conjE] 1 ORELSE ares_tac prems 1)) ]); + + +(*** <-> congruence rules for simplification ***) + +(*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*) +fun iff_tac prems i = + resolve_tac (prems RL [iffE]) i THEN + REPEAT1 (eresolve_tac [asm_rl,mp] i); + +val conj_cong = prove_goal IFOLP.thy + "[| p:P <-> P'; !!x.x:P' ==> q(x):Q <-> Q' |] ==> ?p:(P&Q) <-> (P'&Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [iffI,conjI] 1 + ORELSE eresolve_tac [iffE,conjE,mp] 1 + ORELSE iff_tac prems 1)) ]); + +val disj_cong = prove_goal IFOLP.thy + "[| p:P <-> P'; q:Q <-> Q' |] ==> ?p:(P|Q) <-> (P'|Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (eresolve_tac [iffE,disjE,disjI1,disjI2] 1 + ORELSE ares_tac [iffI] 1 + ORELSE mp_tac 1)) ]); + +val imp_cong = prove_goal IFOLP.thy + "[| p:P <-> P'; !!x.x:P' ==> q(x):Q <-> Q' |] ==> ?p:(P-->Q) <-> (P'-->Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [iffI,impI] 1 + ORELSE eresolve_tac [iffE] 1 + ORELSE mp_tac 1 ORELSE iff_tac prems 1)) ]); + +val iff_cong = prove_goal IFOLP.thy + "[| p:P <-> P'; q:Q <-> Q' |] ==> ?p:(P<->Q) <-> (P'<->Q')" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (eresolve_tac [iffE] 1 + ORELSE ares_tac [iffI] 1 + ORELSE mp_tac 1)) ]); + +val not_cong = prove_goal IFOLP.thy + "p:P <-> P' ==> ?p:~P <-> ~P'" + (fn prems => + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [iffI,notI] 1 + ORELSE mp_tac 1 + ORELSE eresolve_tac [iffE,notE] 1)) ]); + +val all_cong = prove_goal IFOLP.thy + "(!!x.f(x):P(x) <-> Q(x)) ==> ?p:(ALL x.P(x)) <-> (ALL x.Q(x))" + (fn prems => + [ (REPEAT (ares_tac [iffI,allI] 1 + ORELSE mp_tac 1 + ORELSE eresolve_tac [allE] 1 ORELSE iff_tac prems 1)) ]); + +val ex_cong = prove_goal IFOLP.thy + "(!!x.f(x):P(x) <-> Q(x)) ==> ?p:(EX x.P(x)) <-> (EX x.Q(x))" + (fn prems => + [ (REPEAT (eresolve_tac [exE] 1 ORELSE ares_tac [iffI,exI] 1 + ORELSE mp_tac 1 + ORELSE iff_tac prems 1)) ]); + +(*NOT PROVED +val ex1_cong = prove_goal IFOLP.thy + "(!!x.f(x):P(x) <-> Q(x)) ==> ?p:(EX! x.P(x)) <-> (EX! x.Q(x))" + (fn prems => + [ (REPEAT (eresolve_tac [ex1E, spec RS mp] 1 ORELSE ares_tac [iffI,ex1I] 1 + ORELSE mp_tac 1 + ORELSE iff_tac prems 1)) ]); +*) + +(*** Equality rules ***) + +val refl = ieqI; + +val subst = prove_goal IFOLP.thy "[| p:a=b; q:P(a) |] ==> ?p : P(b)" + (fn [prem1,prem2] => [ rtac (prem2 RS rev_mp) 1, (rtac (prem1 RS ieqE) 1), + rtac impI 1, atac 1 ]); + +val sym = prove_goal IFOLP.thy "q:a=b ==> ?c:b=a" + (fn [major] => [ (rtac (major RS subst) 1), (rtac refl 1) ]); + +val trans = prove_goal IFOLP.thy "[| p:a=b; q:b=c |] ==> ?d:a=c" + (fn [prem1,prem2] => [ (rtac (prem2 RS subst) 1), (rtac prem1 1) ]); + +(** ~ b=a ==> ~ a=b **) +val not_sym = prove_goal IFOLP.thy "p:~ b=a ==> ?q:~ a=b" + (fn [prem] => [ (rtac (prem RS contrapos) 1), (etac sym 1) ]); + +(*calling "standard" reduces maxidx to 0*) +val ssubst = standard (sym RS subst); + +(*A special case of ex1E that would otherwise need quantifier expansion*) +val ex1_equalsE = prove_goal IFOLP.thy + "[| p:EX! x.P(x); q:P(a); r:P(b) |] ==> ?d:a=b" + (fn prems => + [ (cut_facts_tac prems 1), + (etac ex1E 1), + (rtac trans 1), + (rtac sym 2), + (REPEAT (eresolve_tac [asm_rl, spec RS mp] 1)) ]); + +(** Polymorphic congruence rules **) + +val subst_context = prove_goal IFOLP.thy + "[| p:a=b |] ==> ?d:t(a)=t(b)" + (fn prems=> + [ (resolve_tac (prems RL [ssubst]) 1), + (resolve_tac [refl] 1) ]); + +val subst_context2 = prove_goal IFOLP.thy + "[| p:a=b; q:c=d |] ==> ?p:t(a,c)=t(b,d)" + (fn prems=> + [ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); + +val subst_context3 = prove_goal IFOLP.thy + "[| p:a=b; q:c=d; r:e=f |] ==> ?p:t(a,c,e)=t(b,d,f)" + (fn prems=> + [ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); + +(*Useful with eresolve_tac for proving equalties from known equalities. + a = b + | | + c = d *) +val box_equals = prove_goal IFOLP.thy + "[| p:a=b; q:a=c; r:b=d |] ==> ?p:c=d" + (fn prems=> + [ (resolve_tac [trans] 1), + (resolve_tac [trans] 1), + (resolve_tac [sym] 1), + (REPEAT (resolve_tac prems 1)) ]); + +(*Dual of box_equals: for proving equalities backwards*) +val simp_equals = prove_goal IFOLP.thy + "[| p:a=c; q:b=d; r:c=d |] ==> ?p:a=b" + (fn prems=> + [ (resolve_tac [trans] 1), + (resolve_tac [trans] 1), + (REPEAT (resolve_tac (prems @ (prems RL [sym])) 1)) ]); + +(** Congruence rules for predicate letters **) + +val pred1_cong = prove_goal IFOLP.thy + "p:a=a' ==> ?p:P(a) <-> P(a')" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); + +val pred2_cong = prove_goal IFOLP.thy + "[| p:a=a'; q:b=b' |] ==> ?p:P(a,b) <-> P(a',b')" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); + +val pred3_cong = prove_goal IFOLP.thy + "[| p:a=a'; q:b=b'; r:c=c' |] ==> ?p:P(a,b,c) <-> P(a',b',c')" + (fn prems => + [ (cut_facts_tac prems 1), + (rtac iffI 1), + (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); + +(*special cases for free variables P, Q, R, S -- up to 3 arguments*) + +val pred_congs = + flat (map (fn c => + map (fn th => read_instantiate [("P",c)] th) + [pred1_cong,pred2_cong,pred3_cong]) + (explode"PQRS")); + +(*special case for the equality predicate!*) +val eq_cong = read_instantiate [("P","op =")] pred2_cong; + + +(*** Simplifications of assumed implications. + Roy Dyckhoff has proved that conj_impE, disj_impE, and imp_impE + used with mp_tac (restricted to atomic formulae) is COMPLETE for + intuitionistic propositional logic. See + R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic + (preprint, University of St Andrews, 1991) ***) + +val conj_impE = prove_goal IFOLP.thy + "[| p:(P&Q)-->S; !!x.x:P-->(Q-->S) ==> q(x):R |] ==> ?p:R" + (fn major::prems=> + [ (REPEAT (ares_tac ([conjI, impI, major RS mp]@prems) 1)) ]); + +val disj_impE = prove_goal IFOLP.thy + "[| p:(P|Q)-->S; !!x y.[| x:P-->S; y:Q-->S |] ==> q(x,y):R |] ==> ?p:R" + (fn major::prems=> + [ (DEPTH_SOLVE (ares_tac ([disjI1, disjI2, impI, major RS mp]@prems) 1)) ]); + +(*Simplifies the implication. Classical version is stronger. + Still UNSAFE since Q must be provable -- backtracking needed. *) +val imp_impE = prove_goal IFOLP.thy + "[| p:(P-->Q)-->S; !!x y.[| x:P; y:Q-->S |] ==> q(x,y):Q; !!x.x:S ==> r(x):R |] ==> \ +\ ?p:R" + (fn major::prems=> + [ (REPEAT (ares_tac ([impI, major RS mp]@prems) 1)) ]); + +(*Simplifies the implication. Classical version is stronger. + Still UNSAFE since ~P must be provable -- backtracking needed. *) +val not_impE = prove_goal IFOLP.thy + "[| p:~P --> S; !!y.y:P ==> q(y):False; !!y.y:S ==> r(y):R |] ==> ?p:R" + (fn major::prems=> + [ (REPEAT (ares_tac ([notI, impI, major RS mp]@prems) 1)) ]); + +(*Simplifies the implication. UNSAFE. *) +val iff_impE = prove_goal IFOLP.thy + "[| p:(P<->Q)-->S; !!x y.[| x:P; y:Q-->S |] ==> q(x,y):Q; \ +\ !!x y.[| x:Q; y:P-->S |] ==> r(x,y):P; !!x.x:S ==> s(x):R |] ==> ?p:R" + (fn major::prems=> + [ (REPEAT (ares_tac ([iffI, impI, major RS mp]@prems) 1)) ]); + +(*What if (ALL x.~~P(x)) --> ~~(ALL x.P(x)) is an assumption? UNSAFE*) +val all_impE = prove_goal IFOLP.thy + "[| p:(ALL x.P(x))-->S; !!x.q:P(x); !!y.y:S ==> r(y):R |] ==> ?p:R" + (fn major::prems=> + [ (REPEAT (ares_tac ([allI, impI, major RS mp]@prems) 1)) ]); + +(*Unsafe: (EX x.P(x))-->S is equivalent to ALL x.P(x)-->S. *) +val ex_impE = prove_goal IFOLP.thy + "[| p:(EX x.P(x))-->S; !!y.y:P(a)-->S ==> q(y):R |] ==> ?p:R" + (fn major::prems=> + [ (REPEAT (ares_tac ([exI, impI, major RS mp]@prems) 1)) ]); + +end; + +open IFOLP_Lemmas; + diff -r 000000000000 -r a5a9c433f639 src/FOLP/ifolp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/ifolp.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,141 @@ +IFOLP = Pure + + +classes term < logic + +default term + +types p,o 0 + +arities p,o :: logic + +consts + (*** Judgements ***) + "@Proof" :: "[p,o]=>prop" ("(_ /: _)" [10,10] 5) + Proof :: "[o,p]=>prop" + EqProof :: "[p,p,o]=>prop" ("(3_ /= _ :/ _)" [10,10,10] 5) + + (*** Logical Connectives -- Type Formers ***) + "=" :: "['a,'a] => o" (infixl 50) + True,False :: "o" + "Not" :: "o => o" ("~ _" [40] 40) + "&" :: "[o,o] => o" (infixr 35) + "|" :: "[o,o] => o" (infixr 30) + "-->" :: "[o,o] => o" (infixr 25) + "<->" :: "[o,o] => o" (infixr 25) + (*Quantifiers*) + All :: "('a => o) => o" (binder "ALL " 10) + Ex :: "('a => o) => o" (binder "EX " 10) + Ex1 :: "('a => o) => o" (binder "EX! " 10) + (*Rewriting gadgets*) + NORM :: "o => o" + norm :: "'a => 'a" + + (*** Proof Term Formers ***) + tt :: "p" + contr :: "p=>p" + fst,snd :: "p=>p" + pair :: "[p,p]=>p" ("(1<_,/_>)") + split :: "[p, [p,p]=>p] =>p" + inl,inr :: "p=>p" + when :: "[p, p=>p, p=>p]=>p" + lambda :: "(p => p) => p" (binder "lam " 20) + "`" :: "[p,p]=>p" (infixl 60) + alll :: "['a=>p]=>p" (binder "all " 15) + "^" :: "[p,'a]=>p" (infixl 50) + exists :: "['a,p]=>p" ("(1[_,/_])") + xsplit :: "[p,['a,p]=>p]=>p" + ideq :: "'a=>p" + idpeel :: "[p,'a=>p]=>p" + nrm, NRM :: "p" + +rules + +(**** Propositional logic ****) + +(*Equality*) +(* Like Intensional Equality in MLTT - but proofs distinct from terms *) + +ieqI "ideq(a) : a=a" +ieqE "[| p : a=b; !!x.f(x) : P(x,x) |] ==> idpeel(p,f) : P(a,b)" + +(* Truth and Falsity *) + +TrueI "tt : True" +FalseE "a:False ==> contr(a):P" + +(* Conjunction *) + +conjI "[| a:P; b:Q |] ==> : P&Q" +conjunct1 "p:P&Q ==> fst(p):P" +conjunct2 "p:P&Q ==> snd(p):Q" + +(* Disjunction *) + +disjI1 "a:P ==> inl(a):P|Q" +disjI2 "b:Q ==> inr(b):P|Q" +disjE "[| a:P|Q; !!x.x:P ==> f(x):R; !!x.x:Q ==> g(x):R \ +\ |] ==> when(a,f,g):R" + +(* Implication *) + +impI "(!!x.x:P ==> f(x):Q) ==> lam x.f(x):P-->Q" +mp "[| f:P-->Q; a:P |] ==> f`a:Q" + +(*Quantifiers*) + +allI "(!!x. f(x) : P(x)) ==> all x.f(x) : ALL x.P(x)" +spec "(f:ALL x.P(x)) ==> f^x : P(x)" + +exI "p : P(x) ==> [x,p] : EX x.P(x)" +exE "[| p: EX x.P(x); !!x u. u:P(x) ==> f(x,u) : R |] ==> xsplit(p,f):R" + +(**** Equality between proofs ****) + +prefl "a : P ==> a = a : P" +psym "a = b : P ==> b = a : P" +ptrans "[| a = b : P; b = c : P |] ==> a = c : P" + +idpeelB "[| !!x.f(x) : P(x,x) |] ==> idpeel(ideq(a),f) = f(a) : P(a,a)" + +fstB "a:P ==> fst() = a : P" +sndB "b:Q ==> snd() = b : Q" +pairEC "p:P&Q ==> p = : P&Q" + +whenBinl "[| a:P; !!x.x:P ==> f(x) : Q |] ==> when(inl(a),f,g) = f(a) : Q" +whenBinr "[| b:P; !!x.x:P ==> g(x) : Q |] ==> when(inr(b),f,g) = g(b) : Q" +plusEC "a:P|Q ==> when(a,%x.inl(x),%y.inr(y)) = p : P|Q" + +applyB "[| a:P; !!x.x:P ==> b(x) : Q |] ==> (lam x.b(x)) ` a = b(a) : Q" +funEC "f:P ==> f = lam x.f`x : P" + +specB "[| !!x.f(x) : P(x) |] ==> (all x.f(x)) ^ a = f(a) : P(a)" + + +(**** Definitions ****) + +not_def "~P == P-->False" +iff_def "P<->Q == (P-->Q) & (Q-->P)" + +(*Unique existence*) +ex1_def "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)" + +(*Rewriting -- special constants to flag normalized terms and formulae*) +norm_eq "nrm : norm(x) = x" +NORM_iff "NRM : NORM(P) <-> P" + +end + +ML + +(*show_proofs:=true displays the proof terms -- they are ENORMOUS*) +val show_proofs = ref false; + +fun proof_tr [p,P] = Const("Proof",dummyT) $ P $ p; + +fun proof_tr' [P,p] = + if !show_proofs then Const("@Proof",dummyT) $ p $ P + else P (*this case discards the proof term*); + +val parse_translation = [("@Proof", proof_tr)]; +val print_translation = [("Proof", proof_tr')]; + diff -r 000000000000 -r a5a9c433f639 src/FOLP/int-prover.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/int-prover.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,79 @@ +(* Title: FOL/int-prover + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +A naive prover for intuitionistic logic + +BEWARE OF NAME CLASHES WITH CLASSICAL TACTICS -- use Int.fast_tac ... + +Completeness (for propositional logic) is proved in + +Roy Dyckhoff. +Contraction-Free Sequent Calculi for Intuitionistic Logic. +J. Symbolic Logic (in press) +*) + +signature INT_PROVER = + sig + val best_tac: int -> tactic + val fast_tac: int -> tactic + val inst_step_tac: int -> tactic + val safe_step_tac: int -> tactic + val safe_brls: (bool * thm) list + val safe_tac: tactic + val step_tac: int -> tactic + val haz_brls: (bool * thm) list + end; + + +structure Int : INT_PROVER = +struct + +(*Negation is treated as a primitive symbol, with rules notI (introduction), + not_to_imp (converts the assumption ~P to P-->False), and not_impE + (handles double negations). Could instead rewrite by not_def as the first + step of an intuitionistic proof. +*) +val safe_brls = sort lessb + [ (true,FalseE), (false,TrueI), (false,refl), + (false,impI), (false,notI), (false,allI), + (true,conjE), (true,exE), + (false,conjI), (true,conj_impE), + (true,disj_impE), (true,ex_impE), + (true,disjE), (false,iffI), (true,iffE), (true,not_to_imp) ]; + +val haz_brls = + [ (false,disjI1), (false,disjI2), (false,exI), + (true,allE), (true,not_impE), (true,imp_impE), (true,iff_impE), + (true,all_impE), (true,impE) ]; + +(*0 subgoals vs 1 or more: the p in safep is for positive*) +val (safe0_brls, safep_brls) = + partition (apl(0,op=) o subgoals_of_brl) safe_brls; + +(*Attack subgoals using safe inferences*) +val safe_step_tac = FIRST' [uniq_assume_tac, + IFOLP_Lemmas.uniq_mp_tac, + biresolve_tac safe0_brls, + hyp_subst_tac, + biresolve_tac safep_brls] ; + +(*Repeatedly attack subgoals using safe inferences*) +val safe_tac = DETERM (REPEAT_FIRST safe_step_tac); + +(*These steps could instantiate variables and are therefore unsafe.*) +val inst_step_tac = assume_tac APPEND' mp_tac; + +(*One safe or unsafe step. *) +fun step_tac i = FIRST [safe_tac, inst_step_tac i, biresolve_tac haz_brls i]; + +(*Dumb but fast*) +val fast_tac = SELECT_GOAL (DEPTH_SOLVE (step_tac 1)); + +(*Slower but smarter than fast_tac*) +val best_tac = + SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) (step_tac 1)); + +end; + diff -r 000000000000 -r a5a9c433f639 src/FOLP/intprover.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/intprover.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,79 @@ +(* Title: FOL/int-prover + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +A naive prover for intuitionistic logic + +BEWARE OF NAME CLASHES WITH CLASSICAL TACTICS -- use Int.fast_tac ... + +Completeness (for propositional logic) is proved in + +Roy Dyckhoff. +Contraction-Free Sequent Calculi for Intuitionistic Logic. +J. Symbolic Logic (in press) +*) + +signature INT_PROVER = + sig + val best_tac: int -> tactic + val fast_tac: int -> tactic + val inst_step_tac: int -> tactic + val safe_step_tac: int -> tactic + val safe_brls: (bool * thm) list + val safe_tac: tactic + val step_tac: int -> tactic + val haz_brls: (bool * thm) list + end; + + +structure Int : INT_PROVER = +struct + +(*Negation is treated as a primitive symbol, with rules notI (introduction), + not_to_imp (converts the assumption ~P to P-->False), and not_impE + (handles double negations). Could instead rewrite by not_def as the first + step of an intuitionistic proof. +*) +val safe_brls = sort lessb + [ (true,FalseE), (false,TrueI), (false,refl), + (false,impI), (false,notI), (false,allI), + (true,conjE), (true,exE), + (false,conjI), (true,conj_impE), + (true,disj_impE), (true,ex_impE), + (true,disjE), (false,iffI), (true,iffE), (true,not_to_imp) ]; + +val haz_brls = + [ (false,disjI1), (false,disjI2), (false,exI), + (true,allE), (true,not_impE), (true,imp_impE), (true,iff_impE), + (true,all_impE), (true,impE) ]; + +(*0 subgoals vs 1 or more: the p in safep is for positive*) +val (safe0_brls, safep_brls) = + partition (apl(0,op=) o subgoals_of_brl) safe_brls; + +(*Attack subgoals using safe inferences*) +val safe_step_tac = FIRST' [uniq_assume_tac, + IFOLP_Lemmas.uniq_mp_tac, + biresolve_tac safe0_brls, + hyp_subst_tac, + biresolve_tac safep_brls] ; + +(*Repeatedly attack subgoals using safe inferences*) +val safe_tac = DETERM (REPEAT_FIRST safe_step_tac); + +(*These steps could instantiate variables and are therefore unsafe.*) +val inst_step_tac = assume_tac APPEND' mp_tac; + +(*One safe or unsafe step. *) +fun step_tac i = FIRST [safe_tac, inst_step_tac i, biresolve_tac haz_brls i]; + +(*Dumb but fast*) +val fast_tac = SELECT_GOAL (DEPTH_SOLVE (step_tac 1)); + +(*Slower but smarter than fast_tac*) +val best_tac = + SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) (step_tac 1)); + +end; + diff -r 000000000000 -r a5a9c433f639 src/FOLP/simp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/simp.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,611 @@ +(* Title: FOLP/simp + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 University of Cambridge + +FOLP version of... + +Generic simplifier, suitable for most logics. (from Provers) + +This version allows instantiation of Vars in the subgoal, since the proof +term must change. +*) + +signature SIMP_DATA = +sig + val case_splits : (thm * string) list + val dest_red : term -> term * term * term + val mk_rew_rules : thm -> thm list + val norm_thms : (thm*thm) list (* [(?x>>norm(?x), norm(?x)>>?x), ...] *) + val red1 : thm (* ?P>>?Q ==> ?P ==> ?Q *) + val red2 : thm (* ?P>>?Q ==> ?Q ==> ?P *) + val refl_thms : thm list + val subst_thms : thm list (* [ ?a>>?b ==> ?P(?a) ==> ?P(?b), ...] *) + val trans_thms : thm list +end; + + +infix 4 addrews addcongs delrews delcongs setauto; + +signature SIMP = +sig + type simpset + val empty_ss : simpset + val addcongs : simpset * thm list -> simpset + val addrews : simpset * thm list -> simpset + val delcongs : simpset * thm list -> simpset + val delrews : simpset * thm list -> simpset + val dest_ss : simpset -> thm list * thm list + val print_ss : simpset -> unit + val setauto : simpset * (int -> tactic) -> simpset + val ASM_SIMP_CASE_TAC : simpset -> int -> tactic + val ASM_SIMP_TAC : simpset -> int -> tactic + val CASE_TAC : simpset -> int -> tactic + val SIMP_CASE2_TAC : simpset -> int -> tactic + val SIMP_THM : simpset -> thm -> thm + val SIMP_TAC : simpset -> int -> tactic + val SIMP_CASE_TAC : simpset -> int -> tactic + val mk_congs : theory -> string list -> thm list + val mk_typed_congs : theory -> (string * string) list -> thm list +(* temporarily disabled: + val extract_free_congs : unit -> thm list +*) + val tracing : bool ref +end; + +functor SimpFun (Simp_data: SIMP_DATA) : SIMP = +struct + +local open Simp_data Logic in + +(*For taking apart reductions into left, right hand sides*) +val lhs_of = #2 o dest_red; +val rhs_of = #3 o dest_red; + +(*** Indexing and filtering of theorems ***) + +fun eq_brl ((b1,th1),(b2,th2)) = b1=b2 andalso eq_thm(th1,th2); + +(*insert a thm in a discrimination net by its lhs*) +fun lhs_insert_thm (th,net) = + Net.insert_term((lhs_of (concl_of th), (false,th)), net, eq_brl) + handle Net.INSERT => net; + +(*match subgoal i against possible theorems in the net. + Similar to match_from_nat_tac, but the net does not contain numbers; + rewrite rules are not ordered.*) +fun net_tac net = + SUBGOAL(fn (prem,i) => + resolve_tac (Net.unify_term net (strip_assums_concl prem)) i); + +(*match subgoal i against possible theorems indexed by lhs in the net*) +fun lhs_net_tac net = + SUBGOAL(fn (prem,i) => + biresolve_tac (Net.unify_term net + (lhs_of (strip_assums_concl prem))) i); + +fun nth_subgoal i thm = nth_elem(i-1,prems_of thm); + +fun goal_concl i thm = strip_assums_concl(nth_subgoal i thm); + +fun lhs_of_eq i thm = lhs_of(goal_concl i thm) +and rhs_of_eq i thm = rhs_of(goal_concl i thm); + +fun var_lhs(thm,i) = +let fun var(Var _) = true + | var(Abs(_,_,t)) = var t + | var(f$_) = var f + | var _ = false; +in var(lhs_of_eq i thm) end; + +fun contains_op opns = + let fun contains(Const(s,_)) = s mem opns | + contains(s$t) = contains s orelse contains t | + contains(Abs(_,_,t)) = contains t | + contains _ = false; + in contains end; + +fun may_match(match_ops,i) = contains_op match_ops o lhs_of_eq i; + +val (normI_thms,normE_thms) = split_list norm_thms; + +(*Get the norm constants from norm_thms*) +val norms = + let fun norm thm = + case lhs_of(concl_of thm) of + Const(n,_)$_ => n + | _ => (prths normE_thms; error"No constant in lhs of a norm_thm") + in map norm normE_thms end; + +fun lhs_is_NORM(thm,i) = case lhs_of_eq i thm of + Const(s,_)$_ => s mem norms | _ => false; + +val refl_tac = resolve_tac refl_thms; + +fun find_res thms thm = + let fun find [] = (prths thms; error"Check Simp_Data") + | find(th::thms) = thm RS th handle _ => find thms + in find thms end; + +val mk_trans = find_res trans_thms; + +fun mk_trans2 thm = +let fun mk[] = error"Check transitivity" + | mk(t::ts) = (thm RSN (2,t)) handle _ => mk ts +in mk trans_thms end; + +(*Applies tactic and returns the first resulting state, FAILS if none!*) +fun one_result(tac,thm) = case Sequence.pull(tapply(tac,thm)) of + Some(thm',_) => thm' + | None => raise THM("Simplifier: could not continue", 0, [thm]); + +fun res1(thm,thms,i) = one_result(resolve_tac thms i,thm); + + +(**** Adding "NORM" tags ****) + +(*get name of the constant from conclusion of a congruence rule*) +fun cong_const cong = + case head_of (lhs_of (concl_of cong)) of + Const(c,_) => c + | _ => "" (*a placeholder distinct from const names*); + +(*true if the term is an atomic proposition (no ==> signs) *) +val atomic = null o strip_assums_hyp; + +(*ccs contains the names of the constants possessing congruence rules*) +fun add_hidden_vars ccs = + let fun add_hvars(tm,hvars) = case tm of + Abs(_,_,body) => add_term_vars(body,hvars) + | _$_ => let val (f,args) = strip_comb tm + in case f of + Const(c,T) => + if c mem ccs + then foldr add_hvars (args,hvars) + else add_term_vars(tm,hvars) + | _ => add_term_vars(tm,hvars) + end + | _ => hvars; + in add_hvars end; + +fun add_new_asm_vars new_asms = + let fun itf((tm,at),vars) = + if at then vars else add_term_vars(tm,vars) + fun add_list(tm,al,vars) = let val (_,tml) = strip_comb tm + in if length(tml)=length(al) + then foldr itf (tml~~al,vars) + else vars + end + fun add_vars (tm,vars) = case tm of + Abs (_,_,body) => add_vars(body,vars) + | r$s => (case head_of tm of + Const(c,T) => (case assoc(new_asms,c) of + None => add_vars(r,add_vars(s,vars)) + | Some(al) => add_list(tm,al,vars)) + | _ => add_vars(r,add_vars(s,vars))) + | _ => vars + in add_vars end; + + +fun add_norms(congs,ccs,new_asms) thm = +let val thm' = mk_trans2 thm; +(* thm': [?z -> l; Prems; r -> ?t] ==> ?z -> ?t *) + val nops = nprems_of thm' + val lhs = rhs_of_eq 1 thm' + val rhs = lhs_of_eq nops thm' + val asms = tl(rev(tl(prems_of thm'))) + val hvars = foldr (add_hidden_vars ccs) (lhs::rhs::asms,[]) + val hvars = add_new_asm_vars new_asms (rhs,hvars) + fun it_asms (asm,hvars) = + if atomic asm then add_new_asm_vars new_asms (asm,hvars) + else add_term_frees(asm,hvars) + val hvars = foldr it_asms (asms,hvars) + val hvs = map (#1 o dest_Var) hvars + val refl1_tac = refl_tac 1 + val add_norm_tac = DEPTH_FIRST (has_fewer_prems nops) + (STATE(fn thm => + case head_of(rhs_of_eq 1 thm) of + Var(ixn,_) => if ixn mem hvs then refl1_tac + else resolve_tac normI_thms 1 ORELSE refl1_tac + | Const _ => resolve_tac normI_thms 1 ORELSE + resolve_tac congs 1 ORELSE refl1_tac + | Free _ => resolve_tac congs 1 ORELSE refl1_tac + | _ => refl1_tac)) + val Some(thm'',_) = Sequence.pull(tapply(add_norm_tac,thm')) +in thm'' end; + +fun add_norm_tags congs = + let val ccs = map cong_const congs + val new_asms = filter (exists not o #2) + (ccs ~~ (map (map atomic o prems_of) congs)); + in add_norms(congs,ccs,new_asms) end; + +fun normed_rews congs = + let val add_norms = add_norm_tags congs; + in fn thm => map (varifyT o add_norms o mk_trans) (mk_rew_rules(freezeT thm)) + end; + +fun NORM norm_lhs_tac = EVERY'[resolve_tac [red2], norm_lhs_tac, refl_tac]; + +val trans_norms = map mk_trans normE_thms; + + +(* SIMPSET *) + +datatype simpset = + SS of {auto_tac: int -> tactic, + congs: thm list, + cong_net: thm Net.net, + mk_simps: thm -> thm list, + simps: (thm * thm list) list, + simp_net: thm Net.net} + +val empty_ss = SS{auto_tac= K no_tac, congs=[], cong_net=Net.empty, + mk_simps=normed_rews[], simps=[], simp_net=Net.empty}; + +(** Insertion of congruences and rewrites **) + +(*insert a thm in a thm net*) +fun insert_thm_warn (th,net) = + Net.insert_term((concl_of th, th), net, eq_thm) + handle Net.INSERT => + (writeln"\nDuplicate rewrite or congruence rule:"; print_thm th; + net); + +val insert_thms = foldr insert_thm_warn; + +fun addrew(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thm) = +let val thms = mk_simps thm +in SS{auto_tac=auto_tac,congs=congs, cong_net=cong_net, mk_simps=mk_simps, + simps = (thm,thms)::simps, simp_net = insert_thms(thms,simp_net)} +end; + +val op addrews = foldl addrew; + +fun op addcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thms) = +let val congs' = thms @ congs; +in SS{auto_tac=auto_tac, congs= congs', + cong_net= insert_thms (map mk_trans thms,cong_net), + mk_simps= normed_rews congs', simps=simps, simp_net=simp_net} +end; + +(** Deletion of congruences and rewrites **) + +(*delete a thm from a thm net*) +fun delete_thm_warn (th,net) = + Net.delete_term((concl_of th, th), net, eq_thm) + handle Net.DELETE => + (writeln"\nNo such rewrite or congruence rule:"; print_thm th; + net); + +val delete_thms = foldr delete_thm_warn; + +fun op delcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thms) = +let val congs' = foldl (gen_rem eq_thm) (congs,thms) +in SS{auto_tac=auto_tac, congs= congs', + cong_net= delete_thms(map mk_trans thms,cong_net), + mk_simps= normed_rews congs', simps=simps, simp_net=simp_net} +end; + +fun delrew(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thm) = +let fun find((p as (th,ths))::ps',ps) = + if eq_thm(thm,th) then (ths,ps@ps') else find(ps',p::ps) + | find([],simps') = (writeln"\nNo such rewrite or congruence rule:"; + print_thm thm; + ([],simps')) + val (thms,simps') = find(simps,[]) +in SS{auto_tac=auto_tac, congs=congs, cong_net=cong_net, mk_simps=mk_simps, + simps = simps', simp_net = delete_thms(thms,simp_net)} +end; + +val op delrews = foldl delrew; + + +fun op setauto(SS{congs,cong_net,mk_simps,simps,simp_net,...}, auto_tac) = + SS{auto_tac=auto_tac, congs=congs, cong_net=cong_net, mk_simps=mk_simps, + simps=simps, simp_net=simp_net}; + + +(** Inspection of a simpset **) + +fun dest_ss(SS{congs,simps,...}) = (congs, map #1 simps); + +fun print_ss(SS{congs,simps,...}) = + (writeln"Congruences:"; prths congs; + writeln"Rewrite Rules:"; prths (map #1 simps); ()); + + +(* Rewriting with conditionals *) + +val (case_thms,case_consts) = split_list case_splits; +val case_rews = map mk_trans case_thms; + +fun if_rewritable ifc i thm = + let val tm = goal_concl i thm + fun nobound(Abs(_,_,tm),j,k) = nobound(tm,j,k+1) + | nobound(s$t,j,k) = nobound(s,j,k) andalso nobound(t,j,k) + | nobound(Bound n,j,k) = n < k orelse k+j <= n + | nobound(_) = true; + fun check_args(al,j) = forall (fn t => nobound(t,j,0)) al + fun find_if(Abs(_,_,tm),j) = find_if(tm,j+1) + | find_if(tm as s$t,j) = let val (f,al) = strip_comb tm in + case f of Const(c,_) => if c=ifc then check_args(al,j) + else find_if(s,j) orelse find_if(t,j) + | _ => find_if(s,j) orelse find_if(t,j) end + | find_if(_) = false; + in find_if(tm,0) end; + +fun IF1_TAC cong_tac i = + let fun seq_try (ifth::ifths,ifc::ifcs) thm = tapply( + COND (if_rewritable ifc i) (DETERM(resolve_tac[ifth]i)) + (Tactic(seq_try(ifths,ifcs))), thm) + | seq_try([],_) thm = tapply(no_tac,thm) + and try_rew thm = tapply(Tactic(seq_try(case_rews,case_consts)) + ORELSE Tactic one_subt, thm) + and one_subt thm = + let val test = has_fewer_prems (nprems_of thm + 1) + fun loop thm = tapply(COND test no_tac + ((Tactic try_rew THEN DEPTH_FIRST test (refl_tac i)) + ORELSE (refl_tac i THEN Tactic loop)), thm) + in tapply(cong_tac THEN Tactic loop, thm) end + in COND (may_match(case_consts,i)) (Tactic try_rew) no_tac end; + +fun CASE_TAC (SS{cong_net,...}) i = +let val cong_tac = net_tac cong_net i +in NORM (IF1_TAC cong_tac) i end; + +(* Rewriting Automaton *) + +datatype cntrl = STOP | MK_EQ | ASMS of int | SIMP_LHS | REW | REFL | TRUE + | PROVE | POP_CS | POP_ARTR | IF; +(* +fun pr_cntrl c = case c of STOP => prs("STOP") | MK_EQ => prs("MK_EQ") | +ASMS i => print_int i | POP_ARTR => prs("POP_ARTR") | +SIMP_LHS => prs("SIMP_LHS") | REW => prs("REW") | REFL => prs("REFL") | +TRUE => prs("TRUE") | PROVE => prs("PROVE") | POP_CS => prs("POP_CS") | IF +=> prs("IF"); +*) +fun simp_refl([],_,ss) = ss + | simp_refl(a'::ns,a,ss) = if a'=a then simp_refl(ns,a,SIMP_LHS::REFL::ss) + else simp_refl(ns,a,ASMS(a)::SIMP_LHS::REFL::POP_ARTR::ss); + +(** Tracing **) + +val tracing = ref false; + +(*Replace parameters by Free variables in P*) +fun variants_abs ([],P) = P + | variants_abs ((a,T)::aTs, P) = + variants_abs (aTs, #2 (variant_abs(a,T,P))); + +(*Select subgoal i from proof state; substitute parameters, for printing*) +fun prepare_goal i st = + let val subgi = nth_subgoal i st + val params = rev(strip_params subgi) + in variants_abs (params, strip_assums_concl subgi) end; + +(*print lhs of conclusion of subgoal i*) +fun pr_goal_lhs i st = + writeln (Sign.string_of_term (#sign(rep_thm st)) + (lhs_of (prepare_goal i st))); + +(*print conclusion of subgoal i*) +fun pr_goal_concl i st = + writeln (Sign.string_of_term (#sign(rep_thm st)) (prepare_goal i st)) + +(*print subgoals i to j (inclusive)*) +fun pr_goals (i,j) st = + if i>j then () + else (pr_goal_concl i st; pr_goals (i+1,j) st); + +(*Print rewrite for tracing; i=subgoal#, n=number of new subgoals, + thm=old state, thm'=new state *) +fun pr_rew (i,n,thm,thm',not_asms) = + if !tracing + then (if not_asms then () else writeln"Assumption used in"; + pr_goal_lhs i thm; writeln"->"; pr_goal_lhs (i+n) thm'; + if n>0 then (writeln"Conditions:"; pr_goals (i, i+n-1) thm') + else (); + writeln"" ) + else (); + +(* Skip the first n hyps of a goal, and return the rest in generalized form *) +fun strip_varify(Const("==>", _) $ H $ B, n, vs) = + if n=0 then subst_bounds(vs,H)::strip_varify(B,0,vs) + else strip_varify(B,n-1,vs) + | strip_varify(Const("all",_)$Abs(_,T,t), n, vs) = + strip_varify(t,n,Var(("?",length vs),T)::vs) + | strip_varify _ = []; + +fun execute(ss,if_fl,auto_tac,cong_tac,net,i,thm) = let + +fun simp_lhs(thm,ss,anet,ats,cs) = + if var_lhs(thm,i) then (ss,thm,anet,ats,cs) else + if lhs_is_NORM(thm,i) then (ss, res1(thm,trans_norms,i), anet,ats,cs) + else case Sequence.pull(tapply(cong_tac i,thm)) of + Some(thm',_) => + let val ps = prems_of thm and ps' = prems_of thm'; + val n = length(ps')-length(ps); + val a = length(strip_assums_hyp(nth_elem(i-1,ps))) + val l = map (fn p => length(strip_assums_hyp(p))) + (take(n,drop(i-1,ps'))); + in (simp_refl(rev(l),a,REW::ss),thm',anet,ats,cs) end + | None => (REW::ss,thm,anet,ats,cs); + +(*NB: the "Adding rewrites:" trace will look strange because assumptions + are represented by rules, generalized over their parameters*) +fun add_asms(ss,thm,a,anet,ats,cs) = + let val As = strip_varify(nth_subgoal i thm, a, []); + val thms = map (trivial o Sign.cterm_of(#sign(rep_thm(thm))))As; + val new_rws = flat(map mk_rew_rules thms); + val rwrls = map mk_trans (flat(map mk_rew_rules thms)); + val anet' = foldr lhs_insert_thm (rwrls,anet) + in if !tracing andalso not(null new_rws) + then (writeln"Adding rewrites:"; prths new_rws; ()) + else (); + (ss,thm,anet',anet::ats,cs) + end; + +fun rew(seq,thm,ss,anet,ats,cs, more) = case Sequence.pull seq of + Some(thm',seq') => + let val n = (nprems_of thm') - (nprems_of thm) + in pr_rew(i,n,thm,thm',more); + if n=0 then (SIMP_LHS::ss, thm', anet, ats, cs) + else ((replicate n PROVE) @ (POP_CS::SIMP_LHS::ss), + thm', anet, ats, (ss,thm,anet,ats,seq',more)::cs) + end + | None => if more + then rew(tapply(lhs_net_tac anet i THEN assume_tac i,thm), + thm,ss,anet,ats,cs,false) + else (ss,thm,anet,ats,cs); + +fun try_true(thm,ss,anet,ats,cs) = + case Sequence.pull(tapply(auto_tac i,thm)) of + Some(thm',_) => (ss,thm',anet,ats,cs) + | None => let val (ss0,thm0,anet0,ats0,seq,more)::cs0 = cs + in if !tracing + then (writeln"*** Failed to prove precondition. Normal form:"; + pr_goal_concl i thm; writeln"") + else (); + rew(seq,thm0,ss0,anet0,ats0,cs0,more) + end; + +fun if_exp(thm,ss,anet,ats,cs) = + case Sequence.pull(tapply(IF1_TAC (cong_tac i) i,thm)) of + Some(thm',_) => (SIMP_LHS::IF::ss,thm',anet,ats,cs) + | None => (ss,thm,anet,ats,cs); + +fun step(s::ss, thm, anet, ats, cs) = case s of + MK_EQ => (ss, res1(thm,[red2],i), anet, ats, cs) + | ASMS(a) => add_asms(ss,thm,a,anet,ats,cs) + | SIMP_LHS => simp_lhs(thm,ss,anet,ats,cs) + | REW => rew(tapply(net_tac net i,thm),thm,ss,anet,ats,cs,true) + | REFL => (ss, res1(thm,refl_thms,i), anet, ats, cs) + | TRUE => try_true(res1(thm,refl_thms,i),ss,anet,ats,cs) + | PROVE => (if if_fl then MK_EQ::SIMP_LHS::IF::TRUE::ss + else MK_EQ::SIMP_LHS::TRUE::ss, thm, anet, ats, cs) + | POP_ARTR => (ss,thm,hd ats,tl ats,cs) + | POP_CS => (ss,thm,anet,ats,tl cs) + | IF => if_exp(thm,ss,anet,ats,cs); + +fun exec(state as (s::ss, thm, _, _, _)) = + if s=STOP then thm else exec(step(state)); + +in exec(ss, thm, Net.empty, [], []) end; + + +fun EXEC_TAC(ss,fl) (SS{auto_tac,cong_net,simp_net,...}) = +let val cong_tac = net_tac cong_net +in fn i => Tactic(fn thm => + if i <= 0 orelse nprems_of thm < i then Sequence.null + else Sequence.single(execute(ss,fl,auto_tac,cong_tac,simp_net,i,thm))) + THEN TRY(auto_tac i) +end; + +val SIMP_TAC = EXEC_TAC([MK_EQ,SIMP_LHS,REFL,STOP],false); +val SIMP_CASE_TAC = EXEC_TAC([MK_EQ,SIMP_LHS,IF,REFL,STOP],false); + +val ASM_SIMP_TAC = EXEC_TAC([ASMS(0),MK_EQ,SIMP_LHS,REFL,STOP],false); +val ASM_SIMP_CASE_TAC = EXEC_TAC([ASMS(0),MK_EQ,SIMP_LHS,IF,REFL,STOP],false); + +val SIMP_CASE2_TAC = EXEC_TAC([MK_EQ,SIMP_LHS,IF,REFL,STOP],true); + +fun REWRITE (ss,fl) (SS{auto_tac,cong_net,simp_net,...}) = +let val cong_tac = net_tac cong_net +in fn thm => let val state = thm RSN (2,red1) + in execute(ss,fl,auto_tac,cong_tac,simp_net,1,state) end +end; + +val SIMP_THM = REWRITE ([ASMS(0),SIMP_LHS,IF,REFL,STOP],false); + + +(* Compute Congruence rules for individual constants using the substition + rules *) + +val subst_thms = map standard subst_thms; + + +fun exp_app(0,t) = t + | exp_app(i,t) = exp_app(i-1,t $ Bound (i-1)); + +fun exp_abs(Type("fun",[T1,T2]),t,i) = + Abs("x"^string_of_int i,T1,exp_abs(T2,t,i+1)) + | exp_abs(T,t,i) = exp_app(i,t); + +fun eta_Var(ixn,T) = exp_abs(T,Var(ixn,T),0); + + +fun Pinst(f,fT,(eq,eqT),k,i,T,yik,Ts) = +let fun xn_list(x,n) = + let val ixs = map (fn i => (x^(radixstring(26,"a",i)),0)) (0 upto n); + in map eta_Var (ixs ~~ (take(n+1,Ts))) end + val lhs = list_comb(f,xn_list("X",k-1)) + val rhs = list_comb(f,xn_list("X",i-1) @ [Bound 0] @ yik) +in Abs("", T, Const(eq,[fT,fT]--->eqT) $ lhs $ rhs) end; + +fun find_subst tsig T = +let fun find (thm::thms) = + let val (Const(_,cT), va, vb) = dest_red(hd(prems_of thm)); + val [P] = add_term_vars(concl_of thm,[]) \\ [va,vb] + val eqT::_ = binder_types cT + in if Type.typ_instance(tsig,T,eqT) then Some(thm,va,vb,P) + else find thms + end + | find [] = None +in find subst_thms end; + +fun mk_cong sg (f,aTs,rT) (refl,eq) = +let val tsig = #tsig(Sign.rep_sg sg); + val k = length aTs; + fun ri((subst,va as Var(_,Ta),vb as Var(_,Tb),P),i,si,T,yik) = + let val ca = Sign.cterm_of sg va + and cx = Sign.cterm_of sg (eta_Var(("X"^si,0),T)) + val cb = Sign.cterm_of sg vb + and cy = Sign.cterm_of sg (eta_Var(("Y"^si,0),T)) + val cP = Sign.cterm_of sg P + and cp = Sign.cterm_of sg (Pinst(f,rT,eq,k,i,T,yik,aTs)) + in cterm_instantiate [(ca,cx),(cb,cy),(cP,cp)] subst end; + fun mk(c,T::Ts,i,yik) = + let val si = radixstring(26,"a",i) + in case find_subst tsig T of + None => mk(c,Ts,i-1,eta_Var(("X"^si,0),T)::yik) + | Some s => let val c' = c RSN (2,ri(s,i,si,T,yik)) + in mk(c',Ts,i-1,eta_Var(("Y"^si,0),T)::yik) end + end + | mk(c,[],_,_) = c; +in mk(refl,rev aTs,k-1,[]) end; + +fun mk_cong_type sg (f,T) = +let val (aTs,rT) = strip_type T; + val tsig = #tsig(Sign.rep_sg sg); + fun find_refl(r::rs) = + let val (Const(eq,eqT),_,_) = dest_red(concl_of r) + in if Type.typ_instance(tsig, rT, hd(binder_types eqT)) + then Some(r,(eq,body_type eqT)) else find_refl rs + end + | find_refl([]) = None; +in case find_refl refl_thms of + None => [] | Some(refl) => [mk_cong sg (f,aTs,rT) refl] +end; + +fun mk_cong_thy thy f = +let val sg = sign_of thy; + val T = case Sign.Symtab.lookup(#const_tab(Sign.rep_sg sg),f) of + None => error(f^" not declared") | Some(T) => T; + val T' = incr_tvar 9 T; +in mk_cong_type sg (Const(f,T'),T') end; + +fun mk_congs thy = flat o map (mk_cong_thy thy); + +fun mk_typed_congs thy = +let val sg = sign_of thy; + val S0 = Type.defaultS(#tsig(Sign.rep_sg sg)) + fun readfT(f,s) = + let val T = incr_tvar 9 (Sign.read_typ(sg,K(Some(S0))) s); + val t = case Sign.Symtab.lookup(#const_tab(Sign.rep_sg sg),f) of + Some(_) => Const(f,T) | None => Free(f,T) + in (t,T) end +in flat o map (mk_cong_type sg o readfT) end; + +end (* local *) +end (* SIMP *); diff -r 000000000000 -r a5a9c433f639 src/FOLP/simpdata.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/FOLP/simpdata.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,135 @@ +(* Title: FOL/simpdata + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Simplification data for FOL +*) + +(*** Rewrite rules ***) + +fun int_prove_fun_raw s = + (writeln s; prove_goal IFOLP.thy s + (fn prems => [ (cut_facts_tac prems 1), (Int.fast_tac 1) ])); + +fun int_prove_fun s = int_prove_fun_raw ("?p : "^s); + +val conj_rews = map int_prove_fun + ["P & True <-> P", "True & P <-> P", + "P & False <-> False", "False & P <-> False", + "P & P <-> P", + "P & ~P <-> False", "~P & P <-> False", + "(P & Q) & R <-> P & (Q & R)"]; + +val disj_rews = map int_prove_fun + ["P | True <-> True", "True | P <-> True", + "P | False <-> P", "False | P <-> P", + "P | P <-> P", + "(P | Q) | R <-> P | (Q | R)"]; + +val not_rews = map int_prove_fun + ["~ False <-> True", "~ True <-> False"]; + +val imp_rews = map int_prove_fun + ["(P --> False) <-> ~P", "(P --> True) <-> True", + "(False --> P) <-> True", "(True --> P) <-> P", + "(P --> P) <-> True", "(P --> ~P) <-> ~P"]; + +val iff_rews = map int_prove_fun + ["(True <-> P) <-> P", "(P <-> True) <-> P", + "(P <-> P) <-> True", + "(False <-> P) <-> ~P", "(P <-> False) <-> ~P"]; + +val quant_rews = map int_prove_fun + ["(ALL x.P) <-> P", "(EX x.P) <-> P"]; + +(*These are NOT supplied by default!*) +val distrib_rews = map int_prove_fun + ["~(P|Q) <-> ~P & ~Q", + "P & (Q | R) <-> P&Q | P&R", "(Q | R) & P <-> Q&P | R&P", + "(P | Q --> R) <-> (P --> R) & (Q --> R)", + "~(EX x.NORM(P(x))) <-> (ALL x. ~NORM(P(x)))", + "((EX x.NORM(P(x))) --> Q) <-> (ALL x. NORM(P(x)) --> Q)", + "(EX x.NORM(P(x))) & NORM(Q) <-> (EX x. NORM(P(x)) & NORM(Q))", + "NORM(Q) & (EX x.NORM(P(x))) <-> (EX x. NORM(Q) & NORM(P(x)))"]; + +val P_Imp_P_iff_T = int_prove_fun_raw "p:P ==> ?p:(P <-> True)"; + +fun make_iff_T th = th RS P_Imp_P_iff_T; + +val refl_iff_T = make_iff_T refl; + +val norm_thms = [(norm_eq RS sym, norm_eq), + (NORM_iff RS iff_sym, NORM_iff)]; + + +(* Conversion into rewrite rules *) + +val not_P_imp_P_iff_F = int_prove_fun_raw "p:~P ==> ?p:(P <-> False)"; + +fun mk_eq th = case concl_of th of + _ $ (Const("op <->",_)$_$_) $ _ => th + | _ $ (Const("op =",_)$_$_) $ _ => th + | _ $ (Const("Not",_)$_) $ _ => th RS not_P_imp_P_iff_F + | _ => make_iff_T th; + +fun atomize th = case concl_of th of (*The operator below is "Proof $ P $ p"*) + _ $ (Const("op -->",_) $ _ $ _) $ _ => atomize(th RS mp) + | _ $ (Const("op &",_) $ _ $ _) $ _ => atomize(th RS conjunct1) @ + atomize(th RS conjunct2) + | _ $ (Const("All",_) $ _) $ _ => atomize(th RS spec) + | _ $ (Const("True",_)) $ _ => [] + | _ $ (Const("False",_)) $ _ => [] + | _ => [th]; + +fun mk_rew_rules th = map mk_eq (atomize th); + +(*destruct function for analysing equations*) +fun dest_red(_ $ (red $ lhs $ rhs) $ _) = (red,lhs,rhs) + | dest_red t = raise TERM("FOL/dest_red", [t]); + +structure FOLP_SimpData : SIMP_DATA = + struct + val refl_thms = [refl, iff_refl] + val trans_thms = [trans, iff_trans] + val red1 = iffD1 + val red2 = iffD2 + val mk_rew_rules = mk_rew_rules + val case_splits = [] (*NO IF'S!*) + val norm_thms = norm_thms + val subst_thms = [subst]; + val dest_red = dest_red + end; + +structure FOLP_Simp = SimpFun(FOLP_SimpData); +structure Induction = InductionFun(struct val spec=IFOLP.spec end); + +(*not a component of SIMP_DATA, but an argument of SIMP_TAC *) +val FOLP_congs = + [all_cong,ex_cong,eq_cong, + conj_cong,disj_cong,imp_cong,iff_cong,not_cong] @ pred_congs; + +val IFOLP_rews = + [refl_iff_T] @ conj_rews @ disj_rews @ not_rews @ + imp_rews @ iff_rews @ quant_rews; + +open FOLP_Simp Induction; + +val auto_ss = empty_ss setauto ares_tac [TrueI]; + +val IFOLP_ss = auto_ss addcongs FOLP_congs addrews IFOLP_rews; + +(*Classical version...*) +fun prove_fun s = + (writeln s; prove_goal FOLP.thy s + (fn prems => [ (cut_facts_tac prems 1), (Cla.fast_tac FOLP_cs 1) ])); + +val cla_rews = map prove_fun + ["?p:P | ~P", "?p:~P | P", + "?p:~ ~ P <-> P", "?p:(~P --> P) <-> P"]; + +val FOLP_rews = IFOLP_rews@cla_rews; + +val FOLP_ss = auto_ss addcongs FOLP_congs addrews FOLP_rews; + + diff -r 000000000000 -r a5a9c433f639 src/LCF/LCF.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LCF/LCF.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,129 @@ +(* Title: LCF/lcf.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1992 University of Cambridge + +For lcf.thy. Basic lemmas about LCF +*) + +open LCF; + +signature LCF_LEMMAS = +sig + val ap_term: thm + val ap_thm: thm + val COND_cases: thm + val COND_cases_iff: thm + val Contrapos: thm + val cong: thm + val ext: thm + val eq_imp_less1: thm + val eq_imp_less2: thm + val less_anti_sym: thm + val less_ap_term: thm + val less_ap_thm: thm + val less_refl: thm + val less_UU: thm + val not_UU_eq_TT: thm + val not_UU_eq_FF: thm + val not_TT_eq_UU: thm + val not_TT_eq_FF: thm + val not_FF_eq_UU: thm + val not_FF_eq_TT: thm + val rstac: thm list -> int -> tactic + val stac: thm -> int -> tactic + val sstac: thm list -> int -> tactic + val strip_tac: int -> tactic + val tr_induct: thm + val UU_abs: thm + val UU_app: thm +end; + + +structure LCF_Lemmas : LCF_LEMMAS = + +struct + +(* Standard abbreviations *) + +val rstac = resolve_tac; +fun stac th = rtac(th RS sym RS subst); +fun sstac ths = EVERY' (map stac ths); + +fun strip_tac i = REPEAT(rstac [impI,allI] i); + +val eq_imp_less1 = prove_goal thy "x=y ==> x << y" + (fn prems => [rtac (rewrite_rule[eq_def](hd prems) RS conjunct1) 1]); + +val eq_imp_less2 = prove_goal thy "x=y ==> y << x" + (fn prems => [rtac (rewrite_rule[eq_def](hd prems) RS conjunct2) 1]); + +val less_refl = refl RS eq_imp_less1; + +val less_anti_sym = prove_goal thy "[| x << y; y << x |] ==> x=y" + (fn prems => [rewrite_goals_tac[eq_def], + REPEAT(rstac(conjI::prems)1)]); + +val ext = prove_goal thy + "(!!x::'a::cpo. f(x)=g(x)::'b::cpo) ==> (%x.f(x))=(%x.g(x))" + (fn [prem] => [REPEAT(rstac[less_anti_sym, less_ext, allI, + prem RS eq_imp_less1, + prem RS eq_imp_less2]1)]); + +val cong = prove_goal thy "[| f=g; x=y |] ==> f(x)=g(y)" + (fn prems => [cut_facts_tac prems 1, etac subst 1, etac subst 1, + rtac refl 1]); + +val less_ap_term = less_refl RS mono; +val less_ap_thm = less_refl RSN (2,mono); +val ap_term = refl RS cong; +val ap_thm = refl RSN (2,cong); + +val UU_abs = prove_goal thy "(%x::'a::cpo.UU) = UU" + (fn _ => [rtac less_anti_sym 1, rtac minimal 2, + rtac less_ext 1, rtac allI 1, rtac minimal 1]); + +val UU_app = UU_abs RS sym RS ap_thm; + +val less_UU = prove_goal thy "x << UU ==> x=UU" + (fn prems=> [rtac less_anti_sym 1,rstac prems 1,rtac minimal 1]); + + +val tr_induct = prove_goal thy "[| P(UU); P(TT); P(FF) |] ==> ALL b.P(b)" + (fn prems => [rtac allI 1, rtac mp 1, + res_inst_tac[("p","b")]tr_cases 2, + fast_tac (FOL_cs addIs prems) 1]); + + +val Contrapos = prove_goal thy "(A ==> B) ==> (~B ==> ~A)" + (fn prems => [rtac notI 1, rtac notE 1, rstac prems 1, + rstac prems 1, atac 1]); + +val not_less_imp_not_eq1 = eq_imp_less1 COMP Contrapos; +val not_less_imp_not_eq2 = eq_imp_less2 COMP Contrapos; + +val not_UU_eq_TT = not_TT_less_UU RS not_less_imp_not_eq2; +val not_UU_eq_FF = not_FF_less_UU RS not_less_imp_not_eq2; +val not_TT_eq_UU = not_TT_less_UU RS not_less_imp_not_eq1; +val not_TT_eq_FF = not_TT_less_FF RS not_less_imp_not_eq1; +val not_FF_eq_UU = not_FF_less_UU RS not_less_imp_not_eq1; +val not_FF_eq_TT = not_FF_less_TT RS not_less_imp_not_eq1; + + +val COND_cases_iff = (prove_goal thy + "ALL b. P(b=>x|y) <-> (b=UU-->P(UU)) & (b=TT-->P(x)) & (b=FF-->P(y))" + (fn _ => [cut_facts_tac [not_UU_eq_TT,not_UU_eq_FF,not_TT_eq_UU, + not_TT_eq_FF,not_FF_eq_UU,not_FF_eq_TT]1, + rtac tr_induct 1, stac COND_UU 1, stac COND_TT 2, + stac COND_FF 3, REPEAT(fast_tac FOL_cs 1)])) RS spec; + +val lemma = prove_goal thy "A<->B ==> B ==> A" + (fn prems => [cut_facts_tac prems 1, rewrite_goals_tac [iff_def], + fast_tac FOL_cs 1]); + +val COND_cases = conjI RSN (2,conjI RS (COND_cases_iff RS lemma)); + +end; + +open LCF_Lemmas; + diff -r 000000000000 -r a5a9c433f639 src/LCF/LCF.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LCF/LCF.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,107 @@ +(* Title: LCF/lcf.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1992 University of Cambridge + +Natural Deduction Rules for LCF +*) + +LCF = FOL + + +classes cpo < term + +default cpo + +types tr,void 0 + "*" 2 (infixl 6) + "+" 2 (infixl 5) + +arities fun, "*", "+" :: (cpo,cpo)cpo + tr,void :: cpo + +consts + UU :: "'a" + TT,FF :: "tr" + FIX :: "('a => 'a) => 'a" + FST :: "'a*'b => 'a" + SND :: "'a*'b => 'b" + INL :: "'a => 'a+'b" + INR :: "'b => 'a+'b" + WHEN :: "['a=>'c, 'b=>'c, 'a+'b] => 'c" + adm :: "('a => o) => o" + VOID :: "void" ("()") + PAIR :: "['a,'b] => 'a*'b" ("(1<_,/_>)" [0,0] 100) + COND :: "[tr,'a,'a] => 'a" ("(_ =>/ (_ |/ _))" [60,60,60] 60) + "<<" :: "['a,'a] => o" (infixl 50) +rules + (** DOMAIN THEORY **) + + eq_def "x=y == x << y & y << x" + + less_trans "[| x << y; y << z |] ==> x << z" + + less_ext "(ALL x. f(x) << g(x)) ==> f << g" + + mono "[| f << g; x << y |] ==> f(x) << g(y)" + + minimal "UU << x" + + FIX_eq "f(FIX(f)) = FIX(f)" + + (** TR **) + + tr_cases "p=UU | p=TT | p=FF" + + not_TT_less_FF "~ TT << FF" + not_FF_less_TT "~ FF << TT" + not_TT_less_UU "~ TT << UU" + not_FF_less_UU "~ FF << UU" + + COND_UU "UU => x | y = UU" + COND_TT "TT => x | y = x" + COND_FF "FF => x | y = y" + + (** PAIRS **) + + surj_pairing " = z" + + FST "FST() = x" + SND "SND() = y" + + (*** STRICT SUM ***) + + INL_DEF "~x=UU ==> ~INL(x)=UU" + INR_DEF "~x=UU ==> ~INR(x)=UU" + + INL_STRICT "INL(UU) = UU" + INR_STRICT "INR(UU) = UU" + + WHEN_UU "WHEN(f,g,UU) = UU" + WHEN_INL "~x=UU ==> WHEN(f,g,INL(x)) = f(x)" + WHEN_INR "~x=UU ==> WHEN(f,g,INR(x)) = g(x)" + + SUM_EXHAUSTION + "z = UU | (EX x. ~x=UU & z = INL(x)) | (EX y. ~y=UU & z = INR(y))" + + (** VOID **) + + void_cases "(x::void) = UU" + + (** INDUCTION **) + + induct "[| adm(P); P(UU); ALL x. P(x) --> P(f(x)) |] ==> P(FIX(f))" + + (** Admissibility / Chain Completeness **) + (* All rules can be found on pages 199--200 of Larry's LCF book. + Note that "easiness" of types is not taken into account + because it cannot be expressed schematically; flatness could be. *) + + adm_less "adm(%x.t(x) << u(x))" + adm_not_less "adm(%x.~ t(x) << u)" + adm_not_free "adm(%x.A)" + adm_subst "adm(P) ==> adm(%x.P(t(x)))" + adm_conj "[| adm(P); adm(Q) |] ==> adm(%x.P(x)&Q(x))" + adm_disj "[| adm(P); adm(Q) |] ==> adm(%x.P(x)|Q(x))" + adm_imp "[| adm(%x.~P(x)); adm(Q) |] ==> adm(%x.P(x)-->Q(x))" + adm_all "(!!y.adm(P(y))) ==> adm(%x.ALL y.P(y,x))" +end diff -r 000000000000 -r a5a9c433f639 src/LCF/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LCF/Makefile Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,42 @@ +######################################################################### +# # +# Makefile for Isabelle (LCF) # +# # +######################################################################### + +#To make the system, cd to this directory and type +# make +#To make the system and test it on standard examples, type +# make test + +#Environment variable ISABELLECOMP specifies the compiler. +#Environment variable ISABELLEBIN specifies the destination directory. +#For Poly/ML, ISABELLEBIN must begin with a / + +#Makes FOL if this file is ABSENT -- but not +#if it is out of date, since this Makefile does not know its dependencies! + +BIN = $(ISABELLEBIN) +COMP = $(ISABELLECOMP) +FILES = ROOT.ML lcf.thy lcf.ML simpdata.ML pair.ML fix.ML ex.ML + +#Uses cp rather than make_database because Poly/ML allows only 3 levels +$(BIN)/LCF: $(BIN)/FOL $(FILES) + case "$(COMP)" in \ + poly*) cp $(BIN)/FOL $(BIN)/LCF;\ + echo 'open PolyML; use"ROOT";' | $(COMP) $(BIN)/LCF ;;\ + sml*) echo 'use"ROOT.ML"; xML"$(BIN)/LCF" banner;' | $(BIN)/FOL;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +(BIN)/FOL: + cd ../FOL; $(MAKE) + +test: ex.ML $(BIN)/LCF + case "$(COMP)" in \ + poly*) echo 'use"ex.ML"; quit();' | $(COMP) $(BIN)/LCF ;;\ + sml*) echo 'use"ex.ML";' | $(BIN)/LCF;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +.PRECIOUS: $(BIN)/FOL $(BIN)/LCF diff -r 000000000000 -r a5a9c433f639 src/LCF/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LCF/README Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,18 @@ + LCF: Logic for Computable Functions + +This directory contains the Standard ML sources of the Isabelle system for +LCF. It is loaded upon FOL, not Pure Isabelle. Important files include + +ROOT.ML -- loads all source files. Enter an ML image containing FOL and +type: use "ROOT.ML"; + +Makefile -- compiles the files under Poly/ML or SML of New Jersey + +ex.ML -- files containing examples. To execute them, enter an ML image +containing LCF and type: use "ex.ML"; + +Useful references on LCF: + + Lawrence C. Paulson, + Logic and Computation: Interactive proof with Cambridge LCF (CUP, 1987) + diff -r 000000000000 -r a5a9c433f639 src/LCF/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LCF/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,20 @@ +(* Title: LCF/ROOT + ID: $Id$ + Author: Tobias Nipkow + Copyright 1992 University of Cambridge + +Adds LCF to a database containing First-Order Logic. + +This theory is based on Lawrence Paulson's book Logic and Computation. +*) + +val banner = "Logic for Computable Functions (in FOL)"; +writeln banner; + +print_depth 1; +use_thy "lcf"; +use"simpdata.ML"; +use"pair.ML"; +use"fix.ML"; + +val LCF_build_completed = (); (*indicate successful build*) diff -r 000000000000 -r a5a9c433f639 src/LCF/ex.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LCF/ex.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,126 @@ +(* Title: LCF/ex.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +Some examples from Lawrence Paulson's book Logic and Computation. +*) + + +LCF_build_completed; (*Cause examples to fail if LCF did*) + +proof_timing := true; + +(*** Section 10.4 ***) + +val ex_thy = extend_theory thy "Ex 10.4" +([], [], [], [], + [(["P"], "'a => tr"), + (["G","H"], "'a => 'a"), + (["K"], "('a => 'a) => ('a => 'a)") + ], + None) +[ ("P_strict", "P(UU) = UU"), + ("K", "K = (%h x. P(x) => x | h(h(G(x))))"), + ("H", "H = FIX(K)") +]; +val ax = get_axiom ex_thy; + +val P_strict = ax"P_strict"; +val K = ax"K"; +val H = ax"H"; + +val ex_ss = LCF_ss addsimps [P_strict,K]; + + +val H_unfold = prove_goal ex_thy "H = K(H)" + (fn _ => [stac H 1, rtac (FIX_eq RS sym) 1]); + +val H_strict = prove_goal ex_thy "H(UU)=UU" + (fn _ => [stac H_unfold 1, simp_tac ex_ss 1]); + +val ex_ss = ex_ss addsimps [H_strict]; + +goal ex_thy "ALL x. H(FIX(K,x)) = FIX(K,x)"; +by(induct_tac "K" 1); +by(simp_tac ex_ss 1); +by(simp_tac (ex_ss setloop (split_tac [COND_cases_iff])) 1); +by(strip_tac 1); +by(stac H_unfold 1); +by(asm_simp_tac ex_ss 1); +val H_idemp_lemma = topthm(); + +val H_idemp = rewrite_rule [mk_meta_eq (H RS sym)] H_idemp_lemma; + + +(*** Example 3.8 ***) + +val ex_thy = extend_theory thy "Ex 3.8" +([], [], [], [], + [(["P"], "'a => tr"), + (["F","G"], "'a => 'a"), + (["H"], "'a => 'b => 'b"), + (["K"], "('a => 'b => 'b) => ('a => 'b => 'b)") + ], + None) +[ ("F_strict", "F(UU) = UU"), + ("K", "K = (%h x y. P(x) => y | F(h(G(x),y)))"), + ("H", "H = FIX(K)") +]; +val ax = get_axiom ex_thy; + +val F_strict = ax"F_strict"; +val K = ax"K"; +val H = ax"H"; + +val ex_ss = LCF_ss addsimps [F_strict,K]; + +goal ex_thy "ALL x. F(H(x::'a,y::'b)) = H(x,F(y))"; +by(stac H 1); +by(induct_tac "K::('a=>'b=>'b)=>('a=>'b=>'b)" 1); +by(simp_tac ex_ss 1); +by(asm_simp_tac (ex_ss setloop (split_tac [COND_cases_iff])) 1); +result(); + + +(*** Addition with fixpoint of successor ***) + +val ex_thy = extend_theory thy "fix ex" +([], [], [], [], + [(["s"], "'a => 'a"), + (["p"], "'a => 'a => 'a") + ], + None) +[ ("p_strict", "p(UU) = UU"), + ("p_s", "p(s(x),y) = s(p(x,y))") +]; +val ax = get_axiom ex_thy; + +val p_strict = ax"p_strict"; +val p_s = ax"p_s"; + +val ex_ss = LCF_ss addsimps [p_strict,p_s]; + +goal ex_thy "p(FIX(s),y) = FIX(s)"; +by(induct_tac "s" 1); +by(simp_tac ex_ss 1); +by(simp_tac ex_ss 1); +result(); + + +(*** Prefixpoints ***) + +val asms = goal thy "[| f(p) << p; !!q. f(q) << q ==> p << q |] ==> FIX(f)=p"; +by(rewtac eq_def); +by (rtac conjI 1); +by(induct_tac "f" 1); +by (rtac minimal 1); +by(strip_tac 1); +by (rtac less_trans 1); +by (resolve_tac asms 2); +by (etac less_ap_term 1); +by (resolve_tac asms 1); +by (rtac (FIX_eq RS eq_imp_less1) 1); +result(); + +maketest"END: file for LCF examples"; diff -r 000000000000 -r a5a9c433f639 src/LCF/fix.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LCF/fix.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,106 @@ +signature FIX = +sig + val adm_eq: thm + val adm_not_eq_tr: thm + val adm_not_not: thm + val not_eq_TT: thm + val not_eq_FF: thm + val not_eq_UU: thm + val induct2: thm + val induct_tac: string -> int -> tactic + val induct2_tac: string*string -> int -> tactic +end; + +structure Fix:FIX = +struct + +val adm_eq = prove_goal LCF.thy "adm(%x.t(x)=u(x)::'a::cpo)" + (fn _ => [rewrite_goals_tac [eq_def], + REPEAT(rstac[adm_conj,adm_less]1)]); + +val adm_not_not = prove_goal LCF.thy "adm(P) ==> adm(%x.~~P(x))" + (fn prems => [simp_tac (LCF_ss addsimps prems) 1]); + + +val tac = rtac tr_induct 1 THEN REPEAT(simp_tac LCF_ss 1); + +val not_eq_TT = prove_goal LCF.thy "ALL p. ~p=TT <-> (p=FF | p=UU)" + (fn _ => [tac]) RS spec; + +val not_eq_FF = prove_goal LCF.thy "ALL p. ~p=FF <-> (p=TT | p=UU)" + (fn _ => [tac]) RS spec; + +val not_eq_UU = prove_goal LCF.thy "ALL p. ~p=UU <-> (p=TT | p=FF)" + (fn _ => [tac]) RS spec; + +val adm_not_eq_tr = prove_goal LCF.thy "ALL p::tr.adm(%x. ~t(x)=p)" + (fn _ => [rtac tr_induct 1, + REPEAT(simp_tac (LCF_ss addsimps [not_eq_TT,not_eq_FF,not_eq_UU]) 1 THEN + REPEAT(rstac [adm_disj,adm_eq] 1))]) RS spec; + +val adm_lemmas = [adm_not_free,adm_eq,adm_less,adm_not_less,adm_not_eq_tr, + adm_conj,adm_disj,adm_imp,adm_all]; + +fun induct_tac v i = res_inst_tac[("f",v)] induct i THEN + REPEAT(rstac adm_lemmas i); + + +val least_FIX = prove_goal LCF.thy "f(p) = p ==> FIX(f) << p" + (fn [prem] => [induct_tac "f" 1, rtac minimal 1, strip_tac 1, + stac (prem RS sym) 1, etac less_ap_term 1]); + +(*Generates unification messages for some reason*) +val lfp_is_FIX = prove_goal LCF.thy + "[| f(p) = p; ALL q. f(q)=q --> p << q |] ==> p = FIX(f)" + (fn [prem1,prem2] => [rtac less_anti_sym 1, + rtac (FIX_eq RS (prem2 RS spec RS mp)) 1, + rtac least_FIX 1, rtac prem1 1]); + +val ffix = read_instantiate [("f","f::?'a=>?'a")] FIX_eq; +val gfix = read_instantiate [("f","g::?'a=>?'a")] FIX_eq; +val ss = LCF_ss addsimps [ffix,gfix]; + +(* Do not use prove_goal because the result is ?ed which leads to divergence + when submitted as an argument to SIMP_THM *) +(* +local +val thm = trivial(Sign.read_cterm(sign_of LCF.thy) + (" = FIX(%p.)",propT)); +val tac = EVERY1[rtac lfp_is_FIX, simp_tac ss, + strip_tac, simp_tac (LCF_ss addsimps [PROD_less]), + rtac conjI, rtac least_FIX, etac subst, rtac (FST RS sym), + rtac least_FIX, etac subst, rtac (SND RS sym)]; +in +val Some(FIX_pair,_) = Sequence.pull(tapply(tac,thm)); +end; + +val FIX_pair_conj = SIMP_THM (LCF_ss addsimps [PROD_eq]) FIX_pair; +*) +val FIX_pair = prove_goal LCF.thy + " = FIX(%p.)" + (fn _ => [rtac lfp_is_FIX 1, simp_tac ss 1, + strip_tac 1, simp_tac (LCF_ss addsimps [PROD_less]) 1, + rtac conjI 1, rtac least_FIX 1, etac subst 1, rtac (FST RS sym) 1, + rtac least_FIX 1, etac subst 1, rtac (SND RS sym) 1]); + +val FIX_pair_conj = rewrite_rule (map mk_meta_eq [PROD_eq,FST,SND]) FIX_pair; + +val FIX1 = FIX_pair_conj RS conjunct1; +val FIX2 = FIX_pair_conj RS conjunct2; + +val induct2 = prove_goal LCF.thy + "[| adm(%p.P(FST(p),SND(p))); P(UU::'a,UU::'b);\ +\ ALL x y. P(x,y) --> P(f(x),g(y)) |] ==> P(FIX(f),FIX(g))" + (fn prems => [EVERY1 + [res_inst_tac [("f","f"),("g","g")] (standard(FIX1 RS ssubst)), + res_inst_tac [("f","f"),("g","g")] (standard(FIX2 RS ssubst)), + res_inst_tac [("f","%x. ")] induct, + rstac prems, simp_tac ss, rstac prems, + simp_tac (LCF_ss addsimps [expand_all_PROD]), rstac prems]]); + +fun induct2_tac (f,g) i = res_inst_tac[("f",f),("g",g)] induct2 i THEN + REPEAT(rstac adm_lemmas i); + +end; + +open Fix; diff -r 000000000000 -r a5a9c433f639 src/LCF/lcf.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LCF/lcf.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,129 @@ +(* Title: LCF/lcf.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1992 University of Cambridge + +For lcf.thy. Basic lemmas about LCF +*) + +open LCF; + +signature LCF_LEMMAS = +sig + val ap_term: thm + val ap_thm: thm + val COND_cases: thm + val COND_cases_iff: thm + val Contrapos: thm + val cong: thm + val ext: thm + val eq_imp_less1: thm + val eq_imp_less2: thm + val less_anti_sym: thm + val less_ap_term: thm + val less_ap_thm: thm + val less_refl: thm + val less_UU: thm + val not_UU_eq_TT: thm + val not_UU_eq_FF: thm + val not_TT_eq_UU: thm + val not_TT_eq_FF: thm + val not_FF_eq_UU: thm + val not_FF_eq_TT: thm + val rstac: thm list -> int -> tactic + val stac: thm -> int -> tactic + val sstac: thm list -> int -> tactic + val strip_tac: int -> tactic + val tr_induct: thm + val UU_abs: thm + val UU_app: thm +end; + + +structure LCF_Lemmas : LCF_LEMMAS = + +struct + +(* Standard abbreviations *) + +val rstac = resolve_tac; +fun stac th = rtac(th RS sym RS subst); +fun sstac ths = EVERY' (map stac ths); + +fun strip_tac i = REPEAT(rstac [impI,allI] i); + +val eq_imp_less1 = prove_goal thy "x=y ==> x << y" + (fn prems => [rtac (rewrite_rule[eq_def](hd prems) RS conjunct1) 1]); + +val eq_imp_less2 = prove_goal thy "x=y ==> y << x" + (fn prems => [rtac (rewrite_rule[eq_def](hd prems) RS conjunct2) 1]); + +val less_refl = refl RS eq_imp_less1; + +val less_anti_sym = prove_goal thy "[| x << y; y << x |] ==> x=y" + (fn prems => [rewrite_goals_tac[eq_def], + REPEAT(rstac(conjI::prems)1)]); + +val ext = prove_goal thy + "(!!x::'a::cpo. f(x)=g(x)::'b::cpo) ==> (%x.f(x))=(%x.g(x))" + (fn [prem] => [REPEAT(rstac[less_anti_sym, less_ext, allI, + prem RS eq_imp_less1, + prem RS eq_imp_less2]1)]); + +val cong = prove_goal thy "[| f=g; x=y |] ==> f(x)=g(y)" + (fn prems => [cut_facts_tac prems 1, etac subst 1, etac subst 1, + rtac refl 1]); + +val less_ap_term = less_refl RS mono; +val less_ap_thm = less_refl RSN (2,mono); +val ap_term = refl RS cong; +val ap_thm = refl RSN (2,cong); + +val UU_abs = prove_goal thy "(%x::'a::cpo.UU) = UU" + (fn _ => [rtac less_anti_sym 1, rtac minimal 2, + rtac less_ext 1, rtac allI 1, rtac minimal 1]); + +val UU_app = UU_abs RS sym RS ap_thm; + +val less_UU = prove_goal thy "x << UU ==> x=UU" + (fn prems=> [rtac less_anti_sym 1,rstac prems 1,rtac minimal 1]); + + +val tr_induct = prove_goal thy "[| P(UU); P(TT); P(FF) |] ==> ALL b.P(b)" + (fn prems => [rtac allI 1, rtac mp 1, + res_inst_tac[("p","b")]tr_cases 2, + fast_tac (FOL_cs addIs prems) 1]); + + +val Contrapos = prove_goal thy "(A ==> B) ==> (~B ==> ~A)" + (fn prems => [rtac notI 1, rtac notE 1, rstac prems 1, + rstac prems 1, atac 1]); + +val not_less_imp_not_eq1 = eq_imp_less1 COMP Contrapos; +val not_less_imp_not_eq2 = eq_imp_less2 COMP Contrapos; + +val not_UU_eq_TT = not_TT_less_UU RS not_less_imp_not_eq2; +val not_UU_eq_FF = not_FF_less_UU RS not_less_imp_not_eq2; +val not_TT_eq_UU = not_TT_less_UU RS not_less_imp_not_eq1; +val not_TT_eq_FF = not_TT_less_FF RS not_less_imp_not_eq1; +val not_FF_eq_UU = not_FF_less_UU RS not_less_imp_not_eq1; +val not_FF_eq_TT = not_FF_less_TT RS not_less_imp_not_eq1; + + +val COND_cases_iff = (prove_goal thy + "ALL b. P(b=>x|y) <-> (b=UU-->P(UU)) & (b=TT-->P(x)) & (b=FF-->P(y))" + (fn _ => [cut_facts_tac [not_UU_eq_TT,not_UU_eq_FF,not_TT_eq_UU, + not_TT_eq_FF,not_FF_eq_UU,not_FF_eq_TT]1, + rtac tr_induct 1, stac COND_UU 1, stac COND_TT 2, + stac COND_FF 3, REPEAT(fast_tac FOL_cs 1)])) RS spec; + +val lemma = prove_goal thy "A<->B ==> B ==> A" + (fn prems => [cut_facts_tac prems 1, rewrite_goals_tac [iff_def], + fast_tac FOL_cs 1]); + +val COND_cases = conjI RSN (2,conjI RS (COND_cases_iff RS lemma)); + +end; + +open LCF_Lemmas; + diff -r 000000000000 -r a5a9c433f639 src/LCF/lcf.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LCF/lcf.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,107 @@ +(* Title: LCF/lcf.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1992 University of Cambridge + +Natural Deduction Rules for LCF +*) + +LCF = FOL + + +classes cpo < term + +default cpo + +types tr,void 0 + "*" 2 (infixl 6) + "+" 2 (infixl 5) + +arities fun, "*", "+" :: (cpo,cpo)cpo + tr,void :: cpo + +consts + UU :: "'a" + TT,FF :: "tr" + FIX :: "('a => 'a) => 'a" + FST :: "'a*'b => 'a" + SND :: "'a*'b => 'b" + INL :: "'a => 'a+'b" + INR :: "'b => 'a+'b" + WHEN :: "['a=>'c, 'b=>'c, 'a+'b] => 'c" + adm :: "('a => o) => o" + VOID :: "void" ("()") + PAIR :: "['a,'b] => 'a*'b" ("(1<_,/_>)" [0,0] 100) + COND :: "[tr,'a,'a] => 'a" ("(_ =>/ (_ |/ _))" [60,60,60] 60) + "<<" :: "['a,'a] => o" (infixl 50) +rules + (** DOMAIN THEORY **) + + eq_def "x=y == x << y & y << x" + + less_trans "[| x << y; y << z |] ==> x << z" + + less_ext "(ALL x. f(x) << g(x)) ==> f << g" + + mono "[| f << g; x << y |] ==> f(x) << g(y)" + + minimal "UU << x" + + FIX_eq "f(FIX(f)) = FIX(f)" + + (** TR **) + + tr_cases "p=UU | p=TT | p=FF" + + not_TT_less_FF "~ TT << FF" + not_FF_less_TT "~ FF << TT" + not_TT_less_UU "~ TT << UU" + not_FF_less_UU "~ FF << UU" + + COND_UU "UU => x | y = UU" + COND_TT "TT => x | y = x" + COND_FF "FF => x | y = y" + + (** PAIRS **) + + surj_pairing " = z" + + FST "FST() = x" + SND "SND() = y" + + (*** STRICT SUM ***) + + INL_DEF "~x=UU ==> ~INL(x)=UU" + INR_DEF "~x=UU ==> ~INR(x)=UU" + + INL_STRICT "INL(UU) = UU" + INR_STRICT "INR(UU) = UU" + + WHEN_UU "WHEN(f,g,UU) = UU" + WHEN_INL "~x=UU ==> WHEN(f,g,INL(x)) = f(x)" + WHEN_INR "~x=UU ==> WHEN(f,g,INR(x)) = g(x)" + + SUM_EXHAUSTION + "z = UU | (EX x. ~x=UU & z = INL(x)) | (EX y. ~y=UU & z = INR(y))" + + (** VOID **) + + void_cases "(x::void) = UU" + + (** INDUCTION **) + + induct "[| adm(P); P(UU); ALL x. P(x) --> P(f(x)) |] ==> P(FIX(f))" + + (** Admissibility / Chain Completeness **) + (* All rules can be found on pages 199--200 of Larry's LCF book. + Note that "easiness" of types is not taken into account + because it cannot be expressed schematically; flatness could be. *) + + adm_less "adm(%x.t(x) << u(x))" + adm_not_less "adm(%x.~ t(x) << u)" + adm_not_free "adm(%x.A)" + adm_subst "adm(P) ==> adm(%x.P(t(x)))" + adm_conj "[| adm(P); adm(Q) |] ==> adm(%x.P(x)&Q(x))" + adm_disj "[| adm(P); adm(Q) |] ==> adm(%x.P(x)|Q(x))" + adm_imp "[| adm(%x.~P(x)); adm(Q) |] ==> adm(%x.P(x)-->Q(x))" + adm_all "(!!y.adm(P(y))) ==> adm(%x.ALL y.P(y,x))" +end diff -r 000000000000 -r a5a9c433f639 src/LCF/pair.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LCF/pair.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,33 @@ +val expand_all_PROD = prove_goal LCF.thy + "(ALL p. P(p)) <-> (ALL x y. P())" + (fn _ => [rtac iffI 1, fast_tac FOL_cs 1, rtac allI 1, + rtac (surj_pairing RS subst) 1, fast_tac FOL_cs 1]); + +local +val ppair = read_instantiate [("z","p::'a*'b")] surj_pairing; +val qpair = read_instantiate [("z","q::'a*'b")] surj_pairing; +in +val PROD_less = prove_goal LCF.thy + "p::'a*'b << q <-> FST(p) << FST(q) & SND(p) << SND(q)" + (fn _ => [EVERY1[rtac iffI, + rtac conjI, etac less_ap_term, etac less_ap_term, + rtac (ppair RS subst), rtac (qpair RS subst), + etac conjE, rtac mono, etac less_ap_term, atac]]); +end; + +val PROD_eq = prove_goal LCF.thy "p=q <-> FST(p)=FST(q) & SND(p)=SND(q)" + (fn _ => [rtac iffI 1, asm_simp_tac LCF_ss 1, + rewrite_goals_tac [eq_def], + asm_simp_tac (LCF_ss addsimps [PROD_less]) 1]); + +val PAIR_less = prove_goal LCF.thy " << <-> a< [simp_tac (LCF_ss addsimps [PROD_less])1]); + +val PAIR_eq = prove_goal LCF.thy " = <-> a=c & b=d" + (fn _ => [simp_tac (LCF_ss addsimps [PROD_eq])1]); + +val UU_is_UU_UU = prove_goal LCF.thy " << UU" + (fn _ => [simp_tac (LCF_ss addsimps [PROD_less]) 1]) + RS less_UU RS sym; + +val LCF_ss = LCF_ss addsimps [PAIR_less,PAIR_eq,UU_is_UU_UU]; diff -r 000000000000 -r a5a9c433f639 src/LCF/simpdata.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LCF/simpdata.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,28 @@ +(* Title: LCF/simpdata + ID: $Id$ + Author: Tobias Nipkow, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Simplification data for LCF +*) + + +fun mk_rew_rules r = +let fun basify thm = + (case concl_of thm of + _$(_$t$_) => (case fastype_of([],t) of + Type("fun",_) => basify(thm RS ap_thm) + | _ => thm) + | _ => thm) +in map (mk_meta_eq o basify) (atomize r) end; + +val LCF_rews = [minimal, + UU_app, UU_app RS ap_thm, UU_app RS ap_thm RS ap_thm, + not_TT_less_FF,not_FF_less_TT,not_TT_less_UU,not_FF_less_UU, + not_UU_eq_TT,not_UU_eq_FF,not_TT_eq_UU,not_TT_eq_FF, + not_FF_eq_UU,not_FF_eq_TT, + COND_UU,COND_TT,COND_FF, + surj_pairing,FST,SND]; + +val LCF_ss = FOL_ss setmksimps mk_rew_rules + addsimps LCF_rews; diff -r 000000000000 -r a5a9c433f639 src/LK/LK.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LK/LK.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,237 @@ +(* Title: LK/lk.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Tactics and lemmas for lk.thy (thanks also to Philippe de Groote) +*) + +open LK; + +(*Higher precedence than := facilitates use of references*) +infix 4 add_safes add_unsafes; + +signature LK_RESOLVE = + sig + datatype pack = Pack of thm list * thm list + val add_safes: pack * thm list -> pack + val add_unsafes: pack * thm list -> pack + val allL_thin: thm + val best_tac: pack -> int -> tactic + val could_res: term * term -> bool + val could_resolve_seq: term * term -> bool + val cutL_tac: string -> int -> tactic + val cutR_tac: string -> int -> tactic + val conL: thm + val conR: thm + val empty_pack: pack + val exR_thin: thm + val fast_tac: pack -> int -> tactic + val filseq_resolve_tac: thm list -> int -> int -> tactic + val forms_of_seq: term -> term list + val has_prems: int -> thm -> bool + val iffL: thm + val iffR: thm + val less: thm * thm -> bool + val LK_dup_pack: pack + val LK_pack: pack + val pc_tac: pack -> int -> tactic + val prop_pack: pack + val repeat_goal_tac: pack -> int -> tactic + val reresolve_tac: thm list -> int -> tactic + val RESOLVE_THEN: thm list -> (int -> tactic) -> int -> tactic + val safe_goal_tac: pack -> int -> tactic + val step_tac: pack -> int -> tactic + val symL: thm + val TrueR: thm + end; + + +structure LK_Resolve : LK_RESOLVE = +struct + +(*Cut and thin, replacing the right-side formula*) +fun cutR_tac (sP: string) i = + res_inst_tac [ ("P",sP) ] cut i THEN rtac thinR i; + +(*Cut and thin, replacing the left-side formula*) +fun cutL_tac (sP: string) i = + res_inst_tac [ ("P",sP) ] cut i THEN rtac thinL (i+1); + + +(** If-and-only-if rules **) +val iffR = prove_goalw LK.thy [iff_def] + "[| $H,P |- $E,Q,$F; $H,Q |- $E,P,$F |] ==> $H |- $E, P <-> Q, $F" + (fn prems=> [ (REPEAT (resolve_tac (prems@[conjR,impR]) 1)) ]); + +val iffL = prove_goalw LK.thy [iff_def] + "[| $H,$G |- $E,P,Q; $H,Q,P,$G |- $E |] ==> $H, P <-> Q, $G |- $E" + (fn prems=> [ (REPEAT (resolve_tac (prems@[conjL,impL,basic]) 1)) ]); + +val TrueR = prove_goalw LK.thy [True_def] + "$H |- $E, True, $F" + (fn _=> [ rtac impR 1, rtac basic 1 ]); + + +(** Weakened quantifier rules. Incomplete, they let the search terminate.**) + +val allL_thin = prove_goal LK.thy + "$H, P(x), $G |- $E ==> $H, ALL x.P(x), $G |- $E" + (fn prems=> [ (rtac allL 1), (rtac thinL 1), (resolve_tac prems 1) ]); + +val exR_thin = prove_goal LK.thy + "$H |- $E, P(x), $F ==> $H |- $E, EX x.P(x), $F" + (fn prems=> [ (rtac exR 1), (rtac thinR 1), (resolve_tac prems 1) ]); + +(* Symmetry of equality in hypotheses *) +val symL = prove_goal LK.thy + "$H, $G, B = A |- $E ==> $H, A = B, $G |- $E" + (fn prems=> + [ (rtac cut 1), + (rtac thinL 2), + (resolve_tac prems 2), + (resolve_tac [basic RS sym] 1) ]); + + +(**** Theorem Packs ****) + +datatype pack = Pack of thm list * thm list; + +(*A theorem pack has the form (safe rules, unsafe rules) + An unsafe rule is incomplete or introduces variables in subgoals, + and is tried only when the safe rules are not applicable. *) + +fun less (rl1,rl2) = (nprems_of rl1) < (nprems_of rl2); + +val empty_pack = Pack([],[]); + +fun (Pack(safes,unsafes)) add_safes ths = + Pack(sort less (ths@safes), unsafes); + +fun (Pack(safes,unsafes)) add_unsafes ths = + Pack(safes, sort less (ths@unsafes)); + +(*The rules of LK*) +val prop_pack = empty_pack add_safes + [basic, refl, conjL, conjR, disjL, disjR, impL, impR, + notL, notR, iffL, iffR]; + +val LK_pack = prop_pack add_safes [allR, exL] + add_unsafes [allL_thin, exR_thin]; + +val LK_dup_pack = prop_pack add_safes [allR, exL] + add_unsafes [allL, exR]; + + + +(*Returns the list of all formulas in the sequent*) +fun forms_of_seq (Const("Seqof",_) $ P $ u) = P :: forms_of_seq u + | forms_of_seq (H $ u) = forms_of_seq u + | forms_of_seq _ = []; + +(*Tests whether two sequences (left or right sides) could be resolved. + seqp is a premise (subgoal), seqc is a conclusion of an object-rule. + Assumes each formula in seqc is surrounded by sequence variables + -- checks that each concl formula looks like some subgoal formula. + It SHOULD check order as well, using recursion rather than forall/exists*) +fun could_res (seqp,seqc) = + forall (fn Qc => exists (fn Qp => could_unify (Qp,Qc)) + (forms_of_seq seqp)) + (forms_of_seq seqc); + +(*Tests whether two sequents G|-H could be resolved, comparing each side.*) +fun could_resolve_seq (prem,conc) = + case (prem,conc) of + (_ $ Abs(_,_,leftp) $ Abs(_,_,rightp), + _ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) => + could_res (leftp,leftc) andalso could_res (rightp,rightc) + | _ => false; + + +(*Like filt_resolve_tac, using could_resolve_seq + Much faster than resolve_tac when there are many rules. + Resolve subgoal i using the rules, unless more than maxr are compatible. *) +fun filseq_resolve_tac rules maxr = SUBGOAL(fn (prem,i) => + let val rls = filter_thms could_resolve_seq (maxr+1, prem, rules) + in if length rls > maxr then no_tac else resolve_tac rls i + end); + + +(*Predicate: does the rule have n premises? *) +fun has_prems n rule = (nprems_of rule = n); + +(*Continuation-style tactical for resolution. + The list of rules is partitioned into 0, 1, 2 premises. + The resulting tactic, gtac, tries to resolve with rules. + If successful, it recursively applies nextac to the new subgoals only. + Else fails. (Treatment of goals due to Ph. de Groote) + Bind (RESOLVE_THEN rules) to a variable: it preprocesses the rules. *) + +(*Takes rule lists separated in to 0, 1, 2, >2 premises. + The abstraction over state prevents needless divergence in recursion. + The 9999 should be a parameter, to delay treatment of flexible goals. *) +fun RESOLVE_THEN rules = + let val [rls0,rls1,rls2] = partition_list has_prems 0 2 rules; + fun tac nextac i = STATE (fn state => + filseq_resolve_tac rls0 9999 i + ORELSE + (DETERM(filseq_resolve_tac rls1 9999 i) THEN TRY(nextac i)) + ORELSE + (DETERM(filseq_resolve_tac rls2 9999 i) THEN TRY(nextac(i+1)) + THEN TRY(nextac i)) ) + in tac end; + + +(*repeated resolution applied to the designated goal*) +fun reresolve_tac rules = + let val restac = RESOLVE_THEN rules; (*preprocessing done now*) + fun gtac i = restac gtac i + in gtac end; + +(*tries the safe rules repeatedly before the unsafe rules. *) +fun repeat_goal_tac (Pack(safes,unsafes)) = + let val restac = RESOLVE_THEN safes + and lastrestac = RESOLVE_THEN unsafes; + fun gtac i = restac gtac i ORELSE lastrestac gtac i + in gtac end; + + +(*Tries safe rules only*) +fun safe_goal_tac (Pack(safes,unsafes)) = reresolve_tac safes; + +(*Tries a safe rule or else a unsafe rule. Single-step for tracing. *) +fun step_tac (thm_pack as Pack(safes,unsafes)) = + safe_goal_tac thm_pack ORELSE' + filseq_resolve_tac unsafes 9999; + + +(* Tactic for reducing a goal, using Predicate Calculus rules. + A decision procedure for Propositional Calculus, it is incomplete + for Predicate-Calculus because of allL_thin and exR_thin. + Fails if it can do nothing. *) +fun pc_tac thm_pack = SELECT_GOAL (DEPTH_SOLVE (repeat_goal_tac thm_pack 1)); + +(*The following two tactics are analogous to those provided by + Provers/classical. In fact, pc_tac is usually FASTER than fast_tac!*) +fun fast_tac thm_pack = + SELECT_GOAL (DEPTH_SOLVE (step_tac thm_pack 1)); + +fun best_tac thm_pack = + SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) + (step_tac thm_pack 1)); + +(** Contraction. Useful since some rules are not complete. **) + +val conR = prove_goal LK.thy + "$H |- $E, P, $F, P ==> $H |- $E, P, $F" + (fn prems=> + [ (rtac cut 1), (REPEAT (resolve_tac (prems@[basic]) 1)) ]); + +val conL = prove_goal LK.thy + "$H, P, $G, P |- $E ==> $H, P, $G |- $E" + (fn prems=> + [ (rtac cut 1), (REPEAT (resolve_tac (prems@[basic]) 1)) ]); + +end; + +open LK_Resolve; diff -r 000000000000 -r a5a9c433f639 src/LK/LK.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LK/LK.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,124 @@ +(* Title: LK/lk.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Classical First-Order Sequent Calculus +*) + +LK = Pure + +classes term < logic +default term +types o 0 + sequence,seqobj,seqcont,sequ,sobj 0 +arities o :: logic +consts + True,False :: "o" + "=" :: "['a,'a] => o" (infixl 50) + "Not" :: "o => o" ("~ _" [40] 40) + "&" :: "[o,o] => o" (infixr 35) + "|" :: "[o,o] => o" (infixr 30) + "-->","<->" :: "[o,o] => o" (infixr 25) + The :: "('a => o) => 'a" (binder "THE " 10) + All :: "('a => o) => o" (binder "ALL " 10) + Ex :: "('a => o) => o" (binder "EX " 10) + + (*Representation of sequents*) + Trueprop :: "[sobj=>sobj,sobj=>sobj] => prop" + Seqof :: "o => sobj=>sobj" + "@Trueprop" :: "[sequence,sequence] => prop" ("((_)/ |- (_))" [6,6] 5) + "@MtSeq" :: "sequence" ("" [] 1000) + "@NmtSeq" :: "[seqobj,seqcont] => sequence" ("__" [] 1000) + "@MtSeqCont" :: "seqcont" ("" [] 1000) + "@SeqCont" :: "[seqobj,seqcont] => seqcont" (",/ __" [] 1000) + "" :: "o => seqobj" ("_" [] 1000) + "@SeqId" :: "id => seqobj" ("$_" [] 1000) + "@SeqVar" :: "var => seqobj" ("$_") + +rules + (*Structural rules*) + + basic "$H, P, $G |- $E, P, $F" + + thinR "$H |- $E, $F ==> $H |- $E, P, $F" + thinL "$H, $G |- $E ==> $H, P, $G |- $E" + + cut "[| $H |- $E, P; $H, P |- $E |] ==> $H |- $E" + + (*Propositional rules*) + + conjR "[| $H|- $E, P, $F; $H|- $E, Q, $F |] ==> $H|- $E, P&Q, $F" + conjL "$H, P, Q, $G |- $E ==> $H, P & Q, $G |- $E" + + disjR "$H |- $E, P, Q, $F ==> $H |- $E, P|Q, $F" + disjL "[| $H, P, $G |- $E; $H, Q, $G |- $E |] ==> $H, P|Q, $G |- $E" + + impR "$H, P |- $E, Q, $F ==> $H |- $E, P-->Q, $F" + impL "[| $H,$G |- $E,P; $H, Q, $G |- $E |] ==> $H, P-->Q, $G |- $E" + + notR "$H, P |- $E, $F ==> $H |- $E, ~P, $F" + notL "$H, $G |- $E, P ==> $H, ~P, $G |- $E" + + FalseL "$H, False, $G |- $E" + + True_def "True == False-->False" + iff_def "P<->Q == (P-->Q) & (Q-->P)" + + (*Quantifiers*) + + allR "(!!x.$H |- $E, P(x), $F) ==> $H |- $E, ALL x.P(x), $F" + allL "$H, P(x), $G, ALL x.P(x) |- $E ==> $H, ALL x.P(x), $G |- $E" + + exR "$H |- $E, P(x), $F, EX x.P(x) ==> $H |- $E, EX x.P(x), $F" + exL "(!!x.$H, P(x), $G |- $E) ==> $H, EX x.P(x), $G |- $E" + + (*Equality*) + + refl "$H |- $E, a=a, $F" + sym "$H |- $E, a=b, $F ==> $H |- $E, b=a, $F" + trans "[| $H|- $E, a=b, $F; $H|- $E, b=c, $F |] ==> $H|- $E, a=c, $F" + + + (*Descriptions*) + + The "[| $H |- $E, P(a), $F; !!x.$H, P(x) |- $E, x=a, $F |] ==> \ +\ $H |- $E, P(THE x.P(x)), $F" +end + +ML + +(*Abstract over "sobj" -- representation of a sequence of formulae *) +fun abs_sobj t = Abs("sobj", Type("sobj",[]), t); + +(*Representation of empty sequence*) +val Sempty = abs_sobj (Bound 0); + +fun seq_obj_tr(Const("@SeqId",_)$id) = id | + seq_obj_tr(Const("@SeqVar",_)$id) = id | + seq_obj_tr(fm) = Const("Seqof",dummyT)$fm; + +fun seq_tr(_$obj$seq) = seq_obj_tr(obj)$seq_tr(seq) | + seq_tr(_) = Bound 0; + +fun seq_tr1(Const("@MtSeq",_)) = Sempty | + seq_tr1(seq) = abs_sobj(seq_tr seq); + +fun true_tr[s1,s2] = Const("Trueprop",dummyT)$seq_tr1 s1$seq_tr1 s2; + +fun seq_obj_tr'(Const("Seqof",_)$fm) = fm | + seq_obj_tr'(id) = Const("@SeqId",dummyT)$id; + +fun seq_tr'(obj$sq,C) = + let val sq' = case sq of + Bound 0 => Const("@MtSeqCont",dummyT) | + _ => seq_tr'(sq,Const("@SeqCont",dummyT)) + in C $ seq_obj_tr' obj $ sq' end; + +fun seq_tr1'(Bound 0) = Const("@MtSeq",dummyT) | + seq_tr1' s = seq_tr'(s,Const("@NmtSeq",dummyT)); + +fun true_tr'[Abs(_,_,s1),Abs(_,_,s2)] = + Const("@Trueprop",dummyT)$seq_tr1' s1$seq_tr1' s2; + +val parse_translation = [("@Trueprop",true_tr)]; +val print_translation = [("Trueprop",true_tr')]; diff -r 000000000000 -r a5a9c433f639 src/LK/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LK/Makefile Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,42 @@ +######################################################################### +# # +# Makefile for Isabelle (LK) # +# # +######################################################################### + +#To make the system, cd to this directory and type +# make -f Makefile +#To make the system and test it on standard examples, type +# make -f Makefile test + +#Environment variable ISABELLECOMP specifies the compiler. +#Environment variable ISABELLEBIN specifies the destination directory. +#For Poly/ML, ISABELLEBIN must begin with a / + +#Makes pure Isabelle (Pure) if this file is ABSENT -- but not +#if it is out of date, since this Makefile does not know its dependencies! + +BIN = $(ISABELLEBIN) +COMP = $(ISABELLECOMP) +FILES = ROOT.ML lk.thy lk.ML + +$(BIN)/LK: $(BIN)/Pure $(FILES) + case "$(COMP)" in \ + poly*) echo 'make_database"$(BIN)/LK"; quit();' \ + | $(COMP) $(BIN)/Pure;\ + echo 'open PolyML; use"ROOT";' | $(COMP) $(BIN)/LK ;;\ + sml*) echo 'use"ROOT.ML"; xML"$(BIN)/LK" banner;' | $(BIN)/Pure ;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +$(BIN)/Pure: + cd ../Pure; $(MAKE) + +test: ex/ROOT.ML $(BIN)/LK + case "$(COMP)" in \ + poly*) echo 'use"ex/ROOT.ML"; quit();' | $(COMP) $(BIN)/LK ;;\ + sml*) echo 'use"ex/ROOT.ML";' | $(BIN)/LK;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +.PRECIOUS: $(BIN)/Pure $(BIN)/LK diff -r 000000000000 -r a5a9c433f639 src/LK/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LK/README Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,20 @@ + LK: Classical First-Order Sequent Calculus + +This directory contains the Standard ML sources of the Isabelle system for +First-Order Logic as a classical sequent calculus. (For a Natural Deduction +version, see FOL.) Important files include + +ROOT.ML -- loads all source files. Enter an ML image containing Pure +Isabelle and type: use "ROOT.ML"; + +Makefile -- compiles the files under Poly/ML or SML of New Jersey + +ex -- subdirectory containing examples. To execute them, enter an ML image +containing LK and type: use "ex/ROOT.ML"; + + +Useful references on sequent calculi: + + Jean Gallier, Logic for Computer Science (Harper&Row, 1986) + + G. Takeuti, Proof Theory (North Holland, 1987) diff -r 000000000000 -r a5a9c433f639 src/LK/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LK/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,20 @@ +(* Title: LK/ROOT + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Adds Classical Sequent Calculus to a database containing pure Isabelle. +Should be executed in the subdirectory LK. +*) + +val banner = "Classical First-Order Sequent Calculus"; +writeln banner; + +print_depth 1; + +use_thy "lk"; + +use "../Pure/install_pp.ML"; +print_depth 8; + +val LK_build_completed = (); (*indicate successful build*) diff -r 000000000000 -r a5a9c433f639 src/LK/ex/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LK/ex/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,18 @@ +(* Title: LK/ex/ROOT + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Executes all examples for Classical Logic. +*) + +LK_build_completed; (*Cause examples to fail if LK did*) + +writeln"Root file for LK examples"; + +proof_timing := true; +time_use "ex/prop.ML"; +time_use "ex/quant.ML"; +time_use "ex/hard-quant.ML"; + +maketest"END: Root file for LK examples"; diff -r 000000000000 -r a5a9c433f639 src/LK/ex/hard-quant.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LK/ex/hard-quant.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,277 @@ +(* Title: LK/ex/hard-quant + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Hard examples with quantifiers. Can be read to test the LK system. +From F. J. Pelletier, + Seventy-Five Problems for Testing Automatic Theorem Provers, + J. Automated Reasoning 2 (1986), 191-216. + Errata, JAR 4 (1988), 236-236. + +Uses pc_tac rather than fast_tac when the former is significantly faster. +*) + +writeln"File LK/ex/hard-quant."; + +goal LK.thy "|- (ALL x. P(x) & Q(x)) <-> (ALL x. P(x)) & (ALL x. Q(x))"; +by (fast_tac LK_pack 1); +result(); + +goal LK.thy "|- (EX x. P-->Q(x)) <-> (P --> (EX x.Q(x)))"; +by (fast_tac LK_pack 1); +result(); + +goal LK.thy "|- (EX x.P(x)-->Q) <-> (ALL x.P(x)) --> Q"; +by (fast_tac LK_pack 1); +result(); + +goal LK.thy "|- (ALL x.P(x)) | Q <-> (ALL x. P(x) | Q)"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problems requiring quantifier duplication"; + +(*Not provable by fast_tac LK_pack: needs multiple instantiation of ALL*) +goal LK.thy "|- (ALL x. P(x)-->P(f(x))) & P(d)-->P(f(f(f(d))))"; +by (best_tac LK_dup_pack 1); +result(); + +(*Needs double instantiation of the quantifier*) +goal LK.thy "|- EX x. P(x) --> P(a) & P(b)"; +by (fast_tac LK_dup_pack 1); +result(); + +goal LK.thy "|- EX z. P(z) --> (ALL x. P(x))"; +by (best_tac LK_dup_pack 1); +result(); + +writeln"Hard examples with quantifiers"; + +writeln"Problem 18"; +goal LK.thy "|- EX y. ALL x. P(y)-->P(x)"; +by (best_tac LK_dup_pack 1); +result(); + +writeln"Problem 19"; +goal LK.thy "|- EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))"; +by (best_tac LK_dup_pack 1); +result(); + +writeln"Problem 20"; +goal LK.thy "|- (ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w))) \ +\ --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 21"; +goal LK.thy "|- (EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> (EX x. P<->Q(x))"; +by (best_tac LK_dup_pack 1); +result(); + +writeln"Problem 22"; +goal LK.thy "|- (ALL x. P <-> Q(x)) --> (P <-> (ALL x. Q(x)))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 23"; +goal LK.thy "|- (ALL x. P | Q(x)) <-> (P | (ALL x. Q(x)))"; +by (best_tac LK_pack 1); +result(); + +writeln"Problem 24"; +goal LK.thy "|- ~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) & \ +\ ~(EX x.P(x)) --> (EX x.Q(x)) & (ALL x. Q(x)|R(x) --> S(x)) \ +\ --> (EX x. P(x)&R(x))"; +by (pc_tac LK_pack 1); +result(); + +writeln"Problem 25"; +goal LK.thy "|- (EX x. P(x)) & \ +\ (ALL x. L(x) --> ~ (M(x) & R(x))) & \ +\ (ALL x. P(x) --> (M(x) & L(x))) & \ +\ ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x))) \ +\ --> (EX x. Q(x)&P(x))"; +by (best_tac LK_pack 1); +result(); + +writeln"Problem 26"; +goal LK.thy "|- ((EX x. p(x)) <-> (EX x. q(x))) & \ +\ (ALL x. ALL y. p(x) & q(y) --> (r(x) <-> s(y))) \ +\ --> ((ALL x. p(x)-->r(x)) <-> (ALL x. q(x)-->s(x)))"; +by (pc_tac LK_pack 1); +result(); + +writeln"Problem 27"; +goal LK.thy "|- (EX x. P(x) & ~Q(x)) & \ +\ (ALL x. P(x) --> R(x)) & \ +\ (ALL x. M(x) & L(x) --> P(x)) & \ +\ ((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x))) \ +\ --> (ALL x. M(x) --> ~L(x))"; +by (pc_tac LK_pack 1); +result(); + +writeln"Problem 28. AMENDED"; +goal LK.thy "|- (ALL x. P(x) --> (ALL x. Q(x))) & \ +\ ((ALL x. Q(x)|R(x)) --> (EX x. Q(x)&S(x))) & \ +\ ((EX x.S(x)) --> (ALL x. L(x) --> M(x))) \ +\ --> (ALL x. P(x) & L(x) --> M(x))"; +by (pc_tac LK_pack 1); +result(); + +writeln"Problem 29. Essentially the same as Principia Mathematica *11.71"; +goal LK.thy "|- (EX x. P(x)) & (EX y. Q(y)) \ +\ --> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y)) <-> \ +\ (ALL x y. P(x) & Q(y) --> R(x) & S(y)))"; +by (pc_tac LK_pack 1); +result(); + +writeln"Problem 30"; +goal LK.thy "|- (ALL x. P(x) | Q(x) --> ~ R(x)) & \ +\ (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x)) \ +\ --> (ALL x. S(x))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 31"; +goal LK.thy "|- ~(EX x.P(x) & (Q(x) | R(x))) & \ +\ (EX x. L(x) & P(x)) & \ +\ (ALL x. ~ R(x) --> M(x)) \ +\ --> (EX x. L(x) & M(x))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 32"; +goal LK.thy "|- (ALL x. P(x) & (Q(x)|R(x))-->S(x)) & \ +\ (ALL x. S(x) & R(x) --> L(x)) & \ +\ (ALL x. M(x) --> R(x)) \ +\ --> (ALL x. P(x) & M(x) --> L(x))"; +by (best_tac LK_pack 1); +result(); + +writeln"Problem 33"; +goal LK.thy "|- (ALL x. P(a) & (P(x)-->P(b))-->P(c)) <-> \ +\ (ALL x. (~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c)))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 34 AMENDED (TWICE!!) NOT PROVED AUTOMATICALLY"; +(*Andrews's challenge*) +goal LK.thy "|- ((EX x. ALL y. p(x) <-> p(y)) <-> \ +\ ((EX x. q(x)) <-> (ALL y. p(y)))) <-> \ +\ ((EX x. ALL y. q(x) <-> q(y)) <-> \ +\ ((EX x. p(x)) <-> (ALL y. q(y))))"; +by (safe_goal_tac LK_pack 1); (*53 secs*) (*13 secs*) +by (TRYALL (fast_tac LK_pack)); (*165 secs*) (*117 secs*) (*138 secs*) +(*for some reason, pc_tac leaves 14 subgoals instead of 6*) +by (TRYALL (best_tac LK_dup_pack)); (*55 secs*) (*29 secs*) (*54 secs*) +result(); + + +writeln"Problem 35"; +goal LK.thy "|- EX x y. P(x,y) --> (ALL u v. P(u,v))"; +by (best_tac LK_dup_pack 1); (*27 secs??*) +result(); + +writeln"Problem 36"; +goal LK.thy "|- (ALL x. EX y. J(x,y)) & \ +\ (ALL x. EX y. G(x,y)) & \ +\ (ALL x y. J(x,y) | G(x,y) --> \ +\ (ALL z. J(y,z) | G(y,z) --> H(x,z))) \ +\ --> (ALL x. EX y. H(x,y))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 37"; +goal LK.thy "|- (ALL z. EX w. ALL x. EX y. \ +\ (P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (EX u.Q(u,w)))) & \ +\ (ALL x z. ~P(x,z) --> (EX y. Q(y,z))) & \ +\ ((EX x y. Q(x,y)) --> (ALL x. R(x,x))) \ +\ --> (ALL x. EX y. R(x,y))"; +by (pc_tac LK_pack 1); (*slow*) +by flexflex_tac; +result(); + +writeln"Problem 38. NOT PROVED"; +goal LK.thy + "|- (ALL x. p(a) & (p(x) --> (EX y. p(y) & r(x,y))) --> \ +\ (EX z. EX w. p(z) & r(x,w) & r(w,z))) <-> \ +\ (ALL x. (~p(a) | p(x) | (EX z. EX w. p(z) & r(x,w) & r(w,z))) & \ +\ (~p(a) | ~(EX y. p(y) & r(x,y)) | \ +\ (EX z. EX w. p(z) & r(x,w) & r(w,z))))"; + +writeln"Problem 39"; +goal LK.thy "|- ~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 40. AMENDED"; +goal LK.thy "|- (EX y. ALL x. F(x,y) <-> F(x,x)) --> \ +\ ~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 41"; +goal LK.thy "|- (ALL z. EX y. ALL x. f(x,y) <-> f(x,z) & ~ f(x,x)) \ +\ --> ~ (EX z. ALL x. f(x,z))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 42"; +goal LK.thy "|- ~ (EX y. ALL x. p(x,y) <-> ~ (EX z. p(x,z) & p(z,x)))"; + +writeln"Problem 43 NOT PROVED AUTOMATICALLY"; +goal LK.thy "|- (ALL x. ALL y. q(x,y) <-> (ALL z. p(z,x) <-> p(z,y))) \ +\ --> (ALL x. (ALL y. q(x,y) <-> q(y,x)))"; + +writeln"Problem 44"; +goal LK.thy "|- (ALL x. f(x) --> \ +\ (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y)))) & \ +\ (EX x. j(x) & (ALL y. g(y) --> h(x,y))) \ +\ --> (EX x. j(x) & ~f(x))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 45"; +goal LK.thy "|- (ALL x. f(x) & (ALL y. g(y) & h(x,y) --> j(x,y)) \ +\ --> (ALL y. g(y) & h(x,y) --> k(y))) & \ +\ ~ (EX y. l(y) & k(y)) & \ +\ (EX x. f(x) & (ALL y. h(x,y) --> l(y)) \ +\ & (ALL y. g(y) & h(x,y) --> j(x,y))) \ +\ --> (EX x. f(x) & ~ (EX y. g(y) & h(x,y)))"; +by (best_tac LK_pack 1); +result(); + + +writeln"Problem 50"; +goal LK.thy + "|- (ALL x. P(a,x) | (ALL y.P(x,y))) --> (EX x. ALL y.P(x,y))"; +by (best_tac LK_dup_pack 1); +result(); + +writeln"Problem 57"; +goal LK.thy + "|- P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) & \ +\ (ALL x y z. P(x,y) & P(y,z) --> P(x,z)) --> P(f(a,b), f(a,c))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 59"; +(*Unification works poorly here -- the abstraction %sobj prevents efficient + operation of the occurs check*) +Unify.trace_bound := !Unify.search_bound - 1; +goal LK.thy "|- (ALL x. P(x) <-> ~P(f(x))) --> (EX x. P(x) & ~P(f(x)))"; +by (best_tac LK_dup_pack 1); +result(); + +writeln"Problem 60"; +goal LK.thy + "|- ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Reached end of file."; + +(*18 June 92: loaded in 372 secs*) +(*19 June 92: loaded in 166 secs except #34, using repeat_goal_tac*) +(*29 June 92: loaded in 370 secs*) diff -r 000000000000 -r a5a9c433f639 src/LK/ex/hardquant.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LK/ex/hardquant.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,277 @@ +(* Title: LK/ex/hard-quant + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Hard examples with quantifiers. Can be read to test the LK system. +From F. J. Pelletier, + Seventy-Five Problems for Testing Automatic Theorem Provers, + J. Automated Reasoning 2 (1986), 191-216. + Errata, JAR 4 (1988), 236-236. + +Uses pc_tac rather than fast_tac when the former is significantly faster. +*) + +writeln"File LK/ex/hard-quant."; + +goal LK.thy "|- (ALL x. P(x) & Q(x)) <-> (ALL x. P(x)) & (ALL x. Q(x))"; +by (fast_tac LK_pack 1); +result(); + +goal LK.thy "|- (EX x. P-->Q(x)) <-> (P --> (EX x.Q(x)))"; +by (fast_tac LK_pack 1); +result(); + +goal LK.thy "|- (EX x.P(x)-->Q) <-> (ALL x.P(x)) --> Q"; +by (fast_tac LK_pack 1); +result(); + +goal LK.thy "|- (ALL x.P(x)) | Q <-> (ALL x. P(x) | Q)"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problems requiring quantifier duplication"; + +(*Not provable by fast_tac LK_pack: needs multiple instantiation of ALL*) +goal LK.thy "|- (ALL x. P(x)-->P(f(x))) & P(d)-->P(f(f(f(d))))"; +by (best_tac LK_dup_pack 1); +result(); + +(*Needs double instantiation of the quantifier*) +goal LK.thy "|- EX x. P(x) --> P(a) & P(b)"; +by (fast_tac LK_dup_pack 1); +result(); + +goal LK.thy "|- EX z. P(z) --> (ALL x. P(x))"; +by (best_tac LK_dup_pack 1); +result(); + +writeln"Hard examples with quantifiers"; + +writeln"Problem 18"; +goal LK.thy "|- EX y. ALL x. P(y)-->P(x)"; +by (best_tac LK_dup_pack 1); +result(); + +writeln"Problem 19"; +goal LK.thy "|- EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))"; +by (best_tac LK_dup_pack 1); +result(); + +writeln"Problem 20"; +goal LK.thy "|- (ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w))) \ +\ --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 21"; +goal LK.thy "|- (EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> (EX x. P<->Q(x))"; +by (best_tac LK_dup_pack 1); +result(); + +writeln"Problem 22"; +goal LK.thy "|- (ALL x. P <-> Q(x)) --> (P <-> (ALL x. Q(x)))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 23"; +goal LK.thy "|- (ALL x. P | Q(x)) <-> (P | (ALL x. Q(x)))"; +by (best_tac LK_pack 1); +result(); + +writeln"Problem 24"; +goal LK.thy "|- ~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) & \ +\ ~(EX x.P(x)) --> (EX x.Q(x)) & (ALL x. Q(x)|R(x) --> S(x)) \ +\ --> (EX x. P(x)&R(x))"; +by (pc_tac LK_pack 1); +result(); + +writeln"Problem 25"; +goal LK.thy "|- (EX x. P(x)) & \ +\ (ALL x. L(x) --> ~ (M(x) & R(x))) & \ +\ (ALL x. P(x) --> (M(x) & L(x))) & \ +\ ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x))) \ +\ --> (EX x. Q(x)&P(x))"; +by (best_tac LK_pack 1); +result(); + +writeln"Problem 26"; +goal LK.thy "|- ((EX x. p(x)) <-> (EX x. q(x))) & \ +\ (ALL x. ALL y. p(x) & q(y) --> (r(x) <-> s(y))) \ +\ --> ((ALL x. p(x)-->r(x)) <-> (ALL x. q(x)-->s(x)))"; +by (pc_tac LK_pack 1); +result(); + +writeln"Problem 27"; +goal LK.thy "|- (EX x. P(x) & ~Q(x)) & \ +\ (ALL x. P(x) --> R(x)) & \ +\ (ALL x. M(x) & L(x) --> P(x)) & \ +\ ((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x))) \ +\ --> (ALL x. M(x) --> ~L(x))"; +by (pc_tac LK_pack 1); +result(); + +writeln"Problem 28. AMENDED"; +goal LK.thy "|- (ALL x. P(x) --> (ALL x. Q(x))) & \ +\ ((ALL x. Q(x)|R(x)) --> (EX x. Q(x)&S(x))) & \ +\ ((EX x.S(x)) --> (ALL x. L(x) --> M(x))) \ +\ --> (ALL x. P(x) & L(x) --> M(x))"; +by (pc_tac LK_pack 1); +result(); + +writeln"Problem 29. Essentially the same as Principia Mathematica *11.71"; +goal LK.thy "|- (EX x. P(x)) & (EX y. Q(y)) \ +\ --> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y)) <-> \ +\ (ALL x y. P(x) & Q(y) --> R(x) & S(y)))"; +by (pc_tac LK_pack 1); +result(); + +writeln"Problem 30"; +goal LK.thy "|- (ALL x. P(x) | Q(x) --> ~ R(x)) & \ +\ (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x)) \ +\ --> (ALL x. S(x))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 31"; +goal LK.thy "|- ~(EX x.P(x) & (Q(x) | R(x))) & \ +\ (EX x. L(x) & P(x)) & \ +\ (ALL x. ~ R(x) --> M(x)) \ +\ --> (EX x. L(x) & M(x))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 32"; +goal LK.thy "|- (ALL x. P(x) & (Q(x)|R(x))-->S(x)) & \ +\ (ALL x. S(x) & R(x) --> L(x)) & \ +\ (ALL x. M(x) --> R(x)) \ +\ --> (ALL x. P(x) & M(x) --> L(x))"; +by (best_tac LK_pack 1); +result(); + +writeln"Problem 33"; +goal LK.thy "|- (ALL x. P(a) & (P(x)-->P(b))-->P(c)) <-> \ +\ (ALL x. (~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c)))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 34 AMENDED (TWICE!!) NOT PROVED AUTOMATICALLY"; +(*Andrews's challenge*) +goal LK.thy "|- ((EX x. ALL y. p(x) <-> p(y)) <-> \ +\ ((EX x. q(x)) <-> (ALL y. p(y)))) <-> \ +\ ((EX x. ALL y. q(x) <-> q(y)) <-> \ +\ ((EX x. p(x)) <-> (ALL y. q(y))))"; +by (safe_goal_tac LK_pack 1); (*53 secs*) (*13 secs*) +by (TRYALL (fast_tac LK_pack)); (*165 secs*) (*117 secs*) (*138 secs*) +(*for some reason, pc_tac leaves 14 subgoals instead of 6*) +by (TRYALL (best_tac LK_dup_pack)); (*55 secs*) (*29 secs*) (*54 secs*) +result(); + + +writeln"Problem 35"; +goal LK.thy "|- EX x y. P(x,y) --> (ALL u v. P(u,v))"; +by (best_tac LK_dup_pack 1); (*27 secs??*) +result(); + +writeln"Problem 36"; +goal LK.thy "|- (ALL x. EX y. J(x,y)) & \ +\ (ALL x. EX y. G(x,y)) & \ +\ (ALL x y. J(x,y) | G(x,y) --> \ +\ (ALL z. J(y,z) | G(y,z) --> H(x,z))) \ +\ --> (ALL x. EX y. H(x,y))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 37"; +goal LK.thy "|- (ALL z. EX w. ALL x. EX y. \ +\ (P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (EX u.Q(u,w)))) & \ +\ (ALL x z. ~P(x,z) --> (EX y. Q(y,z))) & \ +\ ((EX x y. Q(x,y)) --> (ALL x. R(x,x))) \ +\ --> (ALL x. EX y. R(x,y))"; +by (pc_tac LK_pack 1); (*slow*) +by flexflex_tac; +result(); + +writeln"Problem 38. NOT PROVED"; +goal LK.thy + "|- (ALL x. p(a) & (p(x) --> (EX y. p(y) & r(x,y))) --> \ +\ (EX z. EX w. p(z) & r(x,w) & r(w,z))) <-> \ +\ (ALL x. (~p(a) | p(x) | (EX z. EX w. p(z) & r(x,w) & r(w,z))) & \ +\ (~p(a) | ~(EX y. p(y) & r(x,y)) | \ +\ (EX z. EX w. p(z) & r(x,w) & r(w,z))))"; + +writeln"Problem 39"; +goal LK.thy "|- ~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 40. AMENDED"; +goal LK.thy "|- (EX y. ALL x. F(x,y) <-> F(x,x)) --> \ +\ ~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 41"; +goal LK.thy "|- (ALL z. EX y. ALL x. f(x,y) <-> f(x,z) & ~ f(x,x)) \ +\ --> ~ (EX z. ALL x. f(x,z))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 42"; +goal LK.thy "|- ~ (EX y. ALL x. p(x,y) <-> ~ (EX z. p(x,z) & p(z,x)))"; + +writeln"Problem 43 NOT PROVED AUTOMATICALLY"; +goal LK.thy "|- (ALL x. ALL y. q(x,y) <-> (ALL z. p(z,x) <-> p(z,y))) \ +\ --> (ALL x. (ALL y. q(x,y) <-> q(y,x)))"; + +writeln"Problem 44"; +goal LK.thy "|- (ALL x. f(x) --> \ +\ (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y)))) & \ +\ (EX x. j(x) & (ALL y. g(y) --> h(x,y))) \ +\ --> (EX x. j(x) & ~f(x))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 45"; +goal LK.thy "|- (ALL x. f(x) & (ALL y. g(y) & h(x,y) --> j(x,y)) \ +\ --> (ALL y. g(y) & h(x,y) --> k(y))) & \ +\ ~ (EX y. l(y) & k(y)) & \ +\ (EX x. f(x) & (ALL y. h(x,y) --> l(y)) \ +\ & (ALL y. g(y) & h(x,y) --> j(x,y))) \ +\ --> (EX x. f(x) & ~ (EX y. g(y) & h(x,y)))"; +by (best_tac LK_pack 1); +result(); + + +writeln"Problem 50"; +goal LK.thy + "|- (ALL x. P(a,x) | (ALL y.P(x,y))) --> (EX x. ALL y.P(x,y))"; +by (best_tac LK_dup_pack 1); +result(); + +writeln"Problem 57"; +goal LK.thy + "|- P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) & \ +\ (ALL x y z. P(x,y) & P(y,z) --> P(x,z)) --> P(f(a,b), f(a,c))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Problem 59"; +(*Unification works poorly here -- the abstraction %sobj prevents efficient + operation of the occurs check*) +Unify.trace_bound := !Unify.search_bound - 1; +goal LK.thy "|- (ALL x. P(x) <-> ~P(f(x))) --> (EX x. P(x) & ~P(f(x)))"; +by (best_tac LK_dup_pack 1); +result(); + +writeln"Problem 60"; +goal LK.thy + "|- ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Reached end of file."; + +(*18 June 92: loaded in 372 secs*) +(*19 June 92: loaded in 166 secs except #34, using repeat_goal_tac*) +(*29 June 92: loaded in 370 secs*) diff -r 000000000000 -r a5a9c433f639 src/LK/ex/prop.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LK/ex/prop.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,194 @@ +(* Title: LK/ex/prop + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Classical sequent calculus: examples with propositional connectives +Can be read to test the LK system. +*) + +writeln"LK/ex/prop: propositional examples"; + +writeln"absorptive laws of & and | "; + +goal LK.thy "|- P & P <-> P"; +by (fast_tac prop_pack 1); +result(); + +goal LK.thy "|- P | P <-> P"; +by (fast_tac prop_pack 1); +result(); + +writeln"commutative laws of & and | "; + +goal LK.thy "|- P & Q <-> Q & P"; +by (fast_tac prop_pack 1); +result(); + +goal LK.thy "|- P | Q <-> Q | P"; +by (fast_tac prop_pack 1); +result(); + + +writeln"associative laws of & and | "; + +goal LK.thy "|- (P & Q) & R <-> P & (Q & R)"; +by (fast_tac prop_pack 1); +result(); + +goal LK.thy "|- (P | Q) | R <-> P | (Q | R)"; +by (fast_tac prop_pack 1); +result(); + +writeln"distributive laws of & and | "; +goal LK.thy "|- (P & Q) | R <-> (P | R) & (Q | R)"; +by (fast_tac prop_pack 1); +result(); + +goal LK.thy "|- (P | Q) & R <-> (P & R) | (Q & R)"; +by (fast_tac prop_pack 1); +result(); + +writeln"Laws involving implication"; + +goal LK.thy "|- (P|Q --> R) <-> (P-->R) & (Q-->R)"; +by (fast_tac prop_pack 1); +result(); + +goal LK.thy "|- (P & Q --> R) <-> (P--> (Q-->R))"; +by (fast_tac prop_pack 1); +result(); + + +goal LK.thy "|- (P --> Q & R) <-> (P-->Q) & (P-->R)"; +by (fast_tac prop_pack 1); +result(); + + +writeln"Classical theorems"; + +goal LK.thy "|- P|Q --> P| ~P&Q"; +by (fast_tac prop_pack 1); +result(); + + +goal LK.thy "|- (P-->Q)&(~P-->R) --> (P&Q | R)"; +by (fast_tac prop_pack 1); +result(); + + +goal LK.thy "|- P&Q | ~P&R <-> (P-->Q)&(~P-->R)"; +by (fast_tac prop_pack 1); +result(); + + +goal LK.thy "|- (P-->Q) | (P-->R) <-> (P --> Q | R)"; +by (fast_tac prop_pack 1); +result(); + + +(*If and only if*) + +goal LK.thy "|- (P<->Q) <-> (Q<->P)"; +by (fast_tac prop_pack 1); +result(); + +goal LK.thy "|- ~ (P <-> ~P)"; +by (fast_tac prop_pack 1); +result(); + + +(*Sample problems from + F. J. Pelletier, + Seventy-Five Problems for Testing Automatic Theorem Provers, + J. Automated Reasoning 2 (1986), 191-216. + Errata, JAR 4 (1988), 236-236. +*) + +(*1*) +goal LK.thy "|- (P-->Q) <-> (~Q --> ~P)"; +by (fast_tac prop_pack 1); +result(); + +(*2*) +goal LK.thy "|- ~ ~ P <-> P"; +by (fast_tac prop_pack 1); +result(); + +(*3*) +goal LK.thy "|- ~(P-->Q) --> (Q-->P)"; +by (fast_tac prop_pack 1); +result(); + +(*4*) +goal LK.thy "|- (~P-->Q) <-> (~Q --> P)"; +by (fast_tac prop_pack 1); +result(); + +(*5*) +goal LK.thy "|- ((P|Q)-->(P|R)) --> (P|(Q-->R))"; +by (fast_tac prop_pack 1); +result(); + +(*6*) +goal LK.thy "|- P | ~ P"; +by (fast_tac prop_pack 1); +result(); + +(*7*) +goal LK.thy "|- P | ~ ~ ~ P"; +by (fast_tac prop_pack 1); +result(); + +(*8. Peirce's law*) +goal LK.thy "|- ((P-->Q) --> P) --> P"; +by (fast_tac prop_pack 1); +result(); + +(*9*) +goal LK.thy "|- ((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"; +by (fast_tac prop_pack 1); +result(); + +(*10*) +goal LK.thy "Q-->R, R-->P&Q, P-->(Q|R) |- P<->Q"; +by (fast_tac prop_pack 1); +result(); + +(*11. Proved in each direction (incorrectly, says Pelletier!!) *) +goal LK.thy "|- P<->P"; +by (fast_tac prop_pack 1); +result(); + +(*12. "Dijkstra's law"*) +goal LK.thy "|- ((P <-> Q) <-> R) <-> (P <-> (Q <-> R))"; +by (fast_tac prop_pack 1); +result(); + +(*13. Distributive law*) +goal LK.thy "|- P | (Q & R) <-> (P | Q) & (P | R)"; +by (fast_tac prop_pack 1); +result(); + +(*14*) +goal LK.thy "|- (P <-> Q) <-> ((Q | ~P) & (~Q|P))"; +by (fast_tac prop_pack 1); +result(); + + +(*15*) +goal LK.thy "|- (P --> Q) <-> (~P | Q)"; +by (fast_tac prop_pack 1); +result(); + +(*16*) +goal LK.thy "|- (P-->Q) | (Q-->P)"; +by (fast_tac prop_pack 1); +result(); + +(*17*) +goal LK.thy "|- ((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S))"; +by (fast_tac prop_pack 1); +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/LK/ex/quant.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LK/ex/quant.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,160 @@ +(* Title: LK/ex/quant + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Classical sequent calculus: examples with quantifiers. +*) + + +writeln"LK/ex/quant: Examples with quantifiers"; + +goal LK.thy "|- (ALL x. P) <-> P"; +by (fast_tac LK_pack 1); +result(); + +goal LK.thy "|- (ALL x y.P(x,y)) <-> (ALL y x.P(x,y))"; +by (fast_tac LK_pack 1); +result(); + +goal LK.thy "ALL u.P(u), ALL v.Q(v) |- ALL u v. P(u) & Q(v)"; +by (fast_tac LK_pack 1); +result(); + +writeln"Permutation of existential quantifier."; +goal LK.thy "|- (EX x y.P(x,y)) <-> (EX y x.P(x,y))"; +by (fast_tac LK_pack 1); +result(); + +goal LK.thy "|- (ALL x. P(x) & Q(x)) <-> (ALL x.P(x)) & (ALL x.Q(x))"; +by (fast_tac LK_pack 1); +result(); + + +(*Converse is invalid*) +goal LK.thy "|- (ALL x. P(x)) | (ALL x. Q(x)) --> (ALL x. P(x)|Q(x))"; +by (fast_tac LK_pack 1); +result(); + + +writeln"Pushing ALL into an implication."; +goal LK.thy "|- (ALL x. P --> Q(x)) <-> (P --> (ALL x. Q(x)))"; +by (fast_tac LK_pack 1); +result(); + + +goal LK.thy "|- (ALL x.P(x)-->Q) <-> ((EX x.P(x)) --> Q)"; +by (fast_tac LK_pack 1); +result(); + + +goal LK.thy "|- (EX x. P) <-> P"; +by (fast_tac LK_pack 1); +result(); + + +writeln"Distribution of EX over disjunction."; +goal LK.thy "|- (EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))"; +by (fast_tac LK_pack 1); +result(); +(*5 secs*) + +(*Converse is invalid*) +goal LK.thy "|- (EX x. P(x) & Q(x)) --> (EX x. P(x)) & (EX x. Q(x))"; +by (fast_tac LK_pack 1); +result(); + + +writeln"Harder examples: classical theorems."; + +goal LK.thy "|- (EX x. P-->Q(x)) <-> (P --> (EX x.Q(x)))"; +by (fast_tac LK_pack 1); +result(); +(*3 secs*) + + +goal LK.thy "|- (EX x.P(x)-->Q) <-> (ALL x.P(x)) --> Q"; +by (fast_tac LK_pack 1); +result(); +(*5 secs*) + + +goal LK.thy "|- (ALL x.P(x)) | Q <-> (ALL x. P(x) | Q)"; +by (fast_tac LK_pack 1); +result(); + + +writeln"Basic test of quantifier reasoning"; +goal LK.thy + "|- (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"; +by (fast_tac LK_pack 1); +result(); + + +goal LK.thy "|- (ALL x. Q(x)) --> (EX x. Q(x))"; +by (fast_tac LK_pack 1); +result(); + + +writeln"The following are invalid!"; + +(*INVALID*) +goal LK.thy "|- (ALL x. EX y. Q(x,y)) --> (EX y. ALL x. Q(x,y))"; +by (fast_tac LK_pack 1) handle ERROR => writeln"Failed, as expected"; +(*Check that subgoals remain: proof failed.*) +getgoal 1; + +(*INVALID*) +goal LK.thy "|- (EX x. Q(x)) --> (ALL x. Q(x))"; +by (fast_tac LK_pack 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + +goal LK.thy "|- P(?a) --> (ALL x.P(x))"; +by (fast_tac LK_pack 1) handle ERROR => writeln"Failed, as expected"; +(*Check that subgoals remain: proof failed.*) +getgoal 1; + +goal LK.thy "|- (P(?a) --> (ALL x.Q(x))) --> (ALL x. P(x) --> Q(x))"; +by (fast_tac LK_pack 1) handle ERROR => writeln"Failed, as expected"; +getgoal 1; + + +writeln"Back to things that are provable..."; + +goal LK.thy "|- (ALL x. P(x)-->Q(x)) & (EX x.P(x)) --> (EX x.Q(x))"; +by (fast_tac LK_pack 1); +result(); + +(*An example of why exR should be delayed as long as possible*) +goal LK.thy "|- (P--> (EX x.Q(x))) & P--> (EX x.Q(x))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Solving for a Var"; +goal LK.thy + "|- (ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"; +by (fast_tac LK_pack 1); +result(); + + +writeln"Principia Mathematica *11.53"; +goal LK.thy + "|- (ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))"; +by (fast_tac LK_pack 1); +result(); + + +writeln"Principia Mathematica *11.55"; +goal LK.thy "|- (EX x y. P(x) & Q(x,y)) <-> (EX x. P(x) & (EX y.Q(x,y)))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Principia Mathematica *11.61"; +goal LK.thy + "|- (EX y. ALL x. P(x) --> Q(x,y)) --> (ALL x. P(x) --> (EX y. Q(x,y)))"; +by (fast_tac LK_pack 1); +result(); + +writeln"Reached end of file."; + +(*21 August 88: loaded in 45.7 secs*) diff -r 000000000000 -r a5a9c433f639 src/LK/lk.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LK/lk.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,237 @@ +(* Title: LK/lk.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Tactics and lemmas for lk.thy (thanks also to Philippe de Groote) +*) + +open LK; + +(*Higher precedence than := facilitates use of references*) +infix 4 add_safes add_unsafes; + +signature LK_RESOLVE = + sig + datatype pack = Pack of thm list * thm list + val add_safes: pack * thm list -> pack + val add_unsafes: pack * thm list -> pack + val allL_thin: thm + val best_tac: pack -> int -> tactic + val could_res: term * term -> bool + val could_resolve_seq: term * term -> bool + val cutL_tac: string -> int -> tactic + val cutR_tac: string -> int -> tactic + val conL: thm + val conR: thm + val empty_pack: pack + val exR_thin: thm + val fast_tac: pack -> int -> tactic + val filseq_resolve_tac: thm list -> int -> int -> tactic + val forms_of_seq: term -> term list + val has_prems: int -> thm -> bool + val iffL: thm + val iffR: thm + val less: thm * thm -> bool + val LK_dup_pack: pack + val LK_pack: pack + val pc_tac: pack -> int -> tactic + val prop_pack: pack + val repeat_goal_tac: pack -> int -> tactic + val reresolve_tac: thm list -> int -> tactic + val RESOLVE_THEN: thm list -> (int -> tactic) -> int -> tactic + val safe_goal_tac: pack -> int -> tactic + val step_tac: pack -> int -> tactic + val symL: thm + val TrueR: thm + end; + + +structure LK_Resolve : LK_RESOLVE = +struct + +(*Cut and thin, replacing the right-side formula*) +fun cutR_tac (sP: string) i = + res_inst_tac [ ("P",sP) ] cut i THEN rtac thinR i; + +(*Cut and thin, replacing the left-side formula*) +fun cutL_tac (sP: string) i = + res_inst_tac [ ("P",sP) ] cut i THEN rtac thinL (i+1); + + +(** If-and-only-if rules **) +val iffR = prove_goalw LK.thy [iff_def] + "[| $H,P |- $E,Q,$F; $H,Q |- $E,P,$F |] ==> $H |- $E, P <-> Q, $F" + (fn prems=> [ (REPEAT (resolve_tac (prems@[conjR,impR]) 1)) ]); + +val iffL = prove_goalw LK.thy [iff_def] + "[| $H,$G |- $E,P,Q; $H,Q,P,$G |- $E |] ==> $H, P <-> Q, $G |- $E" + (fn prems=> [ (REPEAT (resolve_tac (prems@[conjL,impL,basic]) 1)) ]); + +val TrueR = prove_goalw LK.thy [True_def] + "$H |- $E, True, $F" + (fn _=> [ rtac impR 1, rtac basic 1 ]); + + +(** Weakened quantifier rules. Incomplete, they let the search terminate.**) + +val allL_thin = prove_goal LK.thy + "$H, P(x), $G |- $E ==> $H, ALL x.P(x), $G |- $E" + (fn prems=> [ (rtac allL 1), (rtac thinL 1), (resolve_tac prems 1) ]); + +val exR_thin = prove_goal LK.thy + "$H |- $E, P(x), $F ==> $H |- $E, EX x.P(x), $F" + (fn prems=> [ (rtac exR 1), (rtac thinR 1), (resolve_tac prems 1) ]); + +(* Symmetry of equality in hypotheses *) +val symL = prove_goal LK.thy + "$H, $G, B = A |- $E ==> $H, A = B, $G |- $E" + (fn prems=> + [ (rtac cut 1), + (rtac thinL 2), + (resolve_tac prems 2), + (resolve_tac [basic RS sym] 1) ]); + + +(**** Theorem Packs ****) + +datatype pack = Pack of thm list * thm list; + +(*A theorem pack has the form (safe rules, unsafe rules) + An unsafe rule is incomplete or introduces variables in subgoals, + and is tried only when the safe rules are not applicable. *) + +fun less (rl1,rl2) = (nprems_of rl1) < (nprems_of rl2); + +val empty_pack = Pack([],[]); + +fun (Pack(safes,unsafes)) add_safes ths = + Pack(sort less (ths@safes), unsafes); + +fun (Pack(safes,unsafes)) add_unsafes ths = + Pack(safes, sort less (ths@unsafes)); + +(*The rules of LK*) +val prop_pack = empty_pack add_safes + [basic, refl, conjL, conjR, disjL, disjR, impL, impR, + notL, notR, iffL, iffR]; + +val LK_pack = prop_pack add_safes [allR, exL] + add_unsafes [allL_thin, exR_thin]; + +val LK_dup_pack = prop_pack add_safes [allR, exL] + add_unsafes [allL, exR]; + + + +(*Returns the list of all formulas in the sequent*) +fun forms_of_seq (Const("Seqof",_) $ P $ u) = P :: forms_of_seq u + | forms_of_seq (H $ u) = forms_of_seq u + | forms_of_seq _ = []; + +(*Tests whether two sequences (left or right sides) could be resolved. + seqp is a premise (subgoal), seqc is a conclusion of an object-rule. + Assumes each formula in seqc is surrounded by sequence variables + -- checks that each concl formula looks like some subgoal formula. + It SHOULD check order as well, using recursion rather than forall/exists*) +fun could_res (seqp,seqc) = + forall (fn Qc => exists (fn Qp => could_unify (Qp,Qc)) + (forms_of_seq seqp)) + (forms_of_seq seqc); + +(*Tests whether two sequents G|-H could be resolved, comparing each side.*) +fun could_resolve_seq (prem,conc) = + case (prem,conc) of + (_ $ Abs(_,_,leftp) $ Abs(_,_,rightp), + _ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) => + could_res (leftp,leftc) andalso could_res (rightp,rightc) + | _ => false; + + +(*Like filt_resolve_tac, using could_resolve_seq + Much faster than resolve_tac when there are many rules. + Resolve subgoal i using the rules, unless more than maxr are compatible. *) +fun filseq_resolve_tac rules maxr = SUBGOAL(fn (prem,i) => + let val rls = filter_thms could_resolve_seq (maxr+1, prem, rules) + in if length rls > maxr then no_tac else resolve_tac rls i + end); + + +(*Predicate: does the rule have n premises? *) +fun has_prems n rule = (nprems_of rule = n); + +(*Continuation-style tactical for resolution. + The list of rules is partitioned into 0, 1, 2 premises. + The resulting tactic, gtac, tries to resolve with rules. + If successful, it recursively applies nextac to the new subgoals only. + Else fails. (Treatment of goals due to Ph. de Groote) + Bind (RESOLVE_THEN rules) to a variable: it preprocesses the rules. *) + +(*Takes rule lists separated in to 0, 1, 2, >2 premises. + The abstraction over state prevents needless divergence in recursion. + The 9999 should be a parameter, to delay treatment of flexible goals. *) +fun RESOLVE_THEN rules = + let val [rls0,rls1,rls2] = partition_list has_prems 0 2 rules; + fun tac nextac i = STATE (fn state => + filseq_resolve_tac rls0 9999 i + ORELSE + (DETERM(filseq_resolve_tac rls1 9999 i) THEN TRY(nextac i)) + ORELSE + (DETERM(filseq_resolve_tac rls2 9999 i) THEN TRY(nextac(i+1)) + THEN TRY(nextac i)) ) + in tac end; + + +(*repeated resolution applied to the designated goal*) +fun reresolve_tac rules = + let val restac = RESOLVE_THEN rules; (*preprocessing done now*) + fun gtac i = restac gtac i + in gtac end; + +(*tries the safe rules repeatedly before the unsafe rules. *) +fun repeat_goal_tac (Pack(safes,unsafes)) = + let val restac = RESOLVE_THEN safes + and lastrestac = RESOLVE_THEN unsafes; + fun gtac i = restac gtac i ORELSE lastrestac gtac i + in gtac end; + + +(*Tries safe rules only*) +fun safe_goal_tac (Pack(safes,unsafes)) = reresolve_tac safes; + +(*Tries a safe rule or else a unsafe rule. Single-step for tracing. *) +fun step_tac (thm_pack as Pack(safes,unsafes)) = + safe_goal_tac thm_pack ORELSE' + filseq_resolve_tac unsafes 9999; + + +(* Tactic for reducing a goal, using Predicate Calculus rules. + A decision procedure for Propositional Calculus, it is incomplete + for Predicate-Calculus because of allL_thin and exR_thin. + Fails if it can do nothing. *) +fun pc_tac thm_pack = SELECT_GOAL (DEPTH_SOLVE (repeat_goal_tac thm_pack 1)); + +(*The following two tactics are analogous to those provided by + Provers/classical. In fact, pc_tac is usually FASTER than fast_tac!*) +fun fast_tac thm_pack = + SELECT_GOAL (DEPTH_SOLVE (step_tac thm_pack 1)); + +fun best_tac thm_pack = + SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) + (step_tac thm_pack 1)); + +(** Contraction. Useful since some rules are not complete. **) + +val conR = prove_goal LK.thy + "$H |- $E, P, $F, P ==> $H |- $E, P, $F" + (fn prems=> + [ (rtac cut 1), (REPEAT (resolve_tac (prems@[basic]) 1)) ]); + +val conL = prove_goal LK.thy + "$H, P, $G, P |- $E ==> $H, P, $G |- $E" + (fn prems=> + [ (rtac cut 1), (REPEAT (resolve_tac (prems@[basic]) 1)) ]); + +end; + +open LK_Resolve; diff -r 000000000000 -r a5a9c433f639 src/LK/lk.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LK/lk.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,124 @@ +(* Title: LK/lk.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Classical First-Order Sequent Calculus +*) + +LK = Pure + +classes term < logic +default term +types o 0 + sequence,seqobj,seqcont,sequ,sobj 0 +arities o :: logic +consts + True,False :: "o" + "=" :: "['a,'a] => o" (infixl 50) + "Not" :: "o => o" ("~ _" [40] 40) + "&" :: "[o,o] => o" (infixr 35) + "|" :: "[o,o] => o" (infixr 30) + "-->","<->" :: "[o,o] => o" (infixr 25) + The :: "('a => o) => 'a" (binder "THE " 10) + All :: "('a => o) => o" (binder "ALL " 10) + Ex :: "('a => o) => o" (binder "EX " 10) + + (*Representation of sequents*) + Trueprop :: "[sobj=>sobj,sobj=>sobj] => prop" + Seqof :: "o => sobj=>sobj" + "@Trueprop" :: "[sequence,sequence] => prop" ("((_)/ |- (_))" [6,6] 5) + "@MtSeq" :: "sequence" ("" [] 1000) + "@NmtSeq" :: "[seqobj,seqcont] => sequence" ("__" [] 1000) + "@MtSeqCont" :: "seqcont" ("" [] 1000) + "@SeqCont" :: "[seqobj,seqcont] => seqcont" (",/ __" [] 1000) + "" :: "o => seqobj" ("_" [] 1000) + "@SeqId" :: "id => seqobj" ("$_" [] 1000) + "@SeqVar" :: "var => seqobj" ("$_") + +rules + (*Structural rules*) + + basic "$H, P, $G |- $E, P, $F" + + thinR "$H |- $E, $F ==> $H |- $E, P, $F" + thinL "$H, $G |- $E ==> $H, P, $G |- $E" + + cut "[| $H |- $E, P; $H, P |- $E |] ==> $H |- $E" + + (*Propositional rules*) + + conjR "[| $H|- $E, P, $F; $H|- $E, Q, $F |] ==> $H|- $E, P&Q, $F" + conjL "$H, P, Q, $G |- $E ==> $H, P & Q, $G |- $E" + + disjR "$H |- $E, P, Q, $F ==> $H |- $E, P|Q, $F" + disjL "[| $H, P, $G |- $E; $H, Q, $G |- $E |] ==> $H, P|Q, $G |- $E" + + impR "$H, P |- $E, Q, $F ==> $H |- $E, P-->Q, $F" + impL "[| $H,$G |- $E,P; $H, Q, $G |- $E |] ==> $H, P-->Q, $G |- $E" + + notR "$H, P |- $E, $F ==> $H |- $E, ~P, $F" + notL "$H, $G |- $E, P ==> $H, ~P, $G |- $E" + + FalseL "$H, False, $G |- $E" + + True_def "True == False-->False" + iff_def "P<->Q == (P-->Q) & (Q-->P)" + + (*Quantifiers*) + + allR "(!!x.$H |- $E, P(x), $F) ==> $H |- $E, ALL x.P(x), $F" + allL "$H, P(x), $G, ALL x.P(x) |- $E ==> $H, ALL x.P(x), $G |- $E" + + exR "$H |- $E, P(x), $F, EX x.P(x) ==> $H |- $E, EX x.P(x), $F" + exL "(!!x.$H, P(x), $G |- $E) ==> $H, EX x.P(x), $G |- $E" + + (*Equality*) + + refl "$H |- $E, a=a, $F" + sym "$H |- $E, a=b, $F ==> $H |- $E, b=a, $F" + trans "[| $H|- $E, a=b, $F; $H|- $E, b=c, $F |] ==> $H|- $E, a=c, $F" + + + (*Descriptions*) + + The "[| $H |- $E, P(a), $F; !!x.$H, P(x) |- $E, x=a, $F |] ==> \ +\ $H |- $E, P(THE x.P(x)), $F" +end + +ML + +(*Abstract over "sobj" -- representation of a sequence of formulae *) +fun abs_sobj t = Abs("sobj", Type("sobj",[]), t); + +(*Representation of empty sequence*) +val Sempty = abs_sobj (Bound 0); + +fun seq_obj_tr(Const("@SeqId",_)$id) = id | + seq_obj_tr(Const("@SeqVar",_)$id) = id | + seq_obj_tr(fm) = Const("Seqof",dummyT)$fm; + +fun seq_tr(_$obj$seq) = seq_obj_tr(obj)$seq_tr(seq) | + seq_tr(_) = Bound 0; + +fun seq_tr1(Const("@MtSeq",_)) = Sempty | + seq_tr1(seq) = abs_sobj(seq_tr seq); + +fun true_tr[s1,s2] = Const("Trueprop",dummyT)$seq_tr1 s1$seq_tr1 s2; + +fun seq_obj_tr'(Const("Seqof",_)$fm) = fm | + seq_obj_tr'(id) = Const("@SeqId",dummyT)$id; + +fun seq_tr'(obj$sq,C) = + let val sq' = case sq of + Bound 0 => Const("@MtSeqCont",dummyT) | + _ => seq_tr'(sq,Const("@SeqCont",dummyT)) + in C $ seq_obj_tr' obj $ sq' end; + +fun seq_tr1'(Bound 0) = Const("@MtSeq",dummyT) | + seq_tr1' s = seq_tr'(s,Const("@NmtSeq",dummyT)); + +fun true_tr'[Abs(_,_,s1),Abs(_,_,s2)] = + Const("@Trueprop",dummyT)$seq_tr1' s1$seq_tr1' s2; + +val parse_translation = [("@Trueprop",true_tr)]; +val print_translation = [("Trueprop",true_tr')]; diff -r 000000000000 -r a5a9c433f639 src/Modal/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/Makefile Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,42 @@ +######################################################################### +# # +# Makefile for Isabelle (Modal) # +# # +######################################################################### + +#To make the system, cd to this directory and type +# make -f Makefile +#To make the system and test it on standard examples, type +# make -f Makefile test + +#Environment variable ISABELLECOMP specifies the compiler. +#Environment variable ISABELLEBIN specifies the destination directory. +#For Poly/ML, ISABELLEBIN must begin with a / + +#Makes LK if this file is ABSENT -- but not +#if it is out of date, since this Makefile does not know its dependencies! + +BIN = $(ISABELLEBIN) +COMP = $(ISABELLECOMP) +FILES = ROOT.ML modal0.thy prover.ML T.thy S4.thy S43.thy + +#Uses cp rather than make_database because Poly/ML allows only 3 levels +$(BIN)/Modal: $(BIN)/LK $(FILES) + case "$(COMP)" in \ + poly*) cp $(BIN)/LK $(BIN)/Modal;\ + echo 'open PolyML; use"ROOT";' | $(COMP) $(BIN)/Modal ;;\ + sml*) echo 'use"ROOT.ML"; xML"$(BIN)/Modal" banner;' | $(BIN)/LK;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +$(BIN)/LK: + cd ../LK; $(MAKE) + +test: ex/ROOT.ML $(BIN)/Modal + case "$(COMP)" in \ + poly*) echo 'use"ex/ROOT.ML"; quit();' | $(COMP) $(BIN)/Modal ;;\ + sml*) echo 'use"ex/ROOT.ML";' | $(BIN)/Modal;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +.PRECIOUS: $(BIN)/LK $(BIN)/Modal diff -r 000000000000 -r a5a9c433f639 src/Modal/Modal0.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/Modal0.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,39 @@ +(* Title: 91/Modal/modal0 + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge +*) + +Modal0 = LK + + +consts + box :: "o=>o" ("[]_" [50] 50) + dia :: "o=>o" ("<>_" [50] 50) + "--<",">-<" :: "[o,o]=>o" (infixr 25) + "@Lstar" :: "[sequence,sequence]=>prop" ("(_)|L>(_)" [6,6] 5) + "@Rstar" :: "[sequence,sequence]=>prop" ("(_)|R>(_)" [6,6] 5) + Lstar,Rstar :: "[sobj=>sobj,sobj=>sobj]=>prop" + +rules + (* Definitions *) + + strimp_def "P --< Q == [](P --> Q)" + streqv_def "P >-< Q == (P --< Q) & (Q --< P)" +end + +ML + +local + + val Lstar = "Lstar"; + val Rstar = "Rstar"; + val SLstar = "@Lstar"; + val SRstar = "@Rstar"; + + fun star_tr c [s1,s2] = Const(c,dummyT)$LK.seq_tr1 s1$LK.seq_tr1 s2; + fun star_tr' c [Abs(_,_,s1),Abs(_,_,s2)] = + Const(c,dummyT) $ LK.seq_tr1' s1 $ LK.seq_tr1' s2; +in +val parse_translation = [(SLstar,star_tr Lstar), (SRstar,star_tr Rstar)]; +val print_translation = [(Lstar,star_tr' SLstar), (Rstar,star_tr' SRstar)] +end; diff -r 000000000000 -r a5a9c433f639 src/Modal/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/README Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,22 @@ + Modal: First-Order Modal Sequent Calculus + +This directory contains the Standard ML sources of the Isabelle system for +Modal Logic. Important files include + +ROOT.ML -- loads all source files. Enter an ML image containing LK + and type: use "ROOT.ML"; + +ex -- subdirectory containing examples. Files, Tthms.ML, S4thms.ML and +S43thms.ML contain the theorems of Modal Logics, so arranged since +T $H |- $E, P <-> Q, $F" + (fn prems=> + [ (rewrite_goals_tac [iff_def]), + (REPEAT (resolve_tac (prems@[conjR,impR]) 1)) ]); + + val iffL = prove_goal thy + "[| $H,$G |- $E,P,Q; $H,Q,P,$G |- $E |] ==> $H, P <-> Q, $G |- $E" + (fn prems=> + [ rewrite_goals_tac [iff_def], + (REPEAT (resolve_tac (prems@[conjL,impL,basic]) 1)) ]) +in +val safe_rls = [basic,conjL,conjR,disjL,disjR,impL,impR,notL,notR,iffL,iffR]; +val unsafe_rls = [allR,exL]; +val bound_rls = [allL,exR]; +end + +end; + +use "prover.ML"; + +use_thy"T"; + +local open Modal0_rls T +in structure MP_Rule : MODAL_PROVER_RULE = + struct + val rewrite_rls = rewrite_rls + val safe_rls = safe_rls + val unsafe_rls = unsafe_rls @ [boxR,diaL] + val bound_rls = bound_rls @ [boxL,diaR] + val aside_rls = [lstar0,lstar1,lstar2,rstar0,rstar1,rstar2] + end +end; +structure T_Prover = Modal_ProverFun(MP_Rule); + +use_thy"S4"; + +local open Modal0_rls S4 +in structure MP_Rule : MODAL_PROVER_RULE = + struct + val rewrite_rls = rewrite_rls + val safe_rls = safe_rls + val unsafe_rls = unsafe_rls @ [boxR,diaL] + val bound_rls = bound_rls @ [boxL,diaR] + val aside_rls = [lstar0,lstar1,lstar2,rstar0,rstar1,rstar2] + end; +end; +structure S4_Prover = Modal_ProverFun(MP_Rule); + +use_thy"S43"; + +local open Modal0_rls S43 +in structure MP_Rule : MODAL_PROVER_RULE = + struct + val rewrite_rls = rewrite_rls + val safe_rls = safe_rls + val unsafe_rls = unsafe_rls @ [pi1,pi2] + val bound_rls = bound_rls @ [boxL,diaR] + val aside_rls = [lstar0,lstar1,lstar2,rstar0,rstar1,rstar2,S43pi0,S43pi1,S43pi2] + end; +end; +structure S43_Prover = Modal_ProverFun(MP_Rule); + +val Modal_build_completed = (); (*indicate successful build*) + +(*printing functions are inherited from LK*) diff -r 000000000000 -r a5a9c433f639 src/Modal/S4.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/S4.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,30 @@ +(* Title: 91/Modal/S4 + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge +*) + +S4 = Modal0 + +rules +(* Definition of the star operation using a set of Horn clauses *) +(* For system S4: gamma * == {[]P | []P : gamma} *) +(* delta * == {<>P | <>P : delta} *) + + lstar0 "|L>" + lstar1 "$G |L> $H ==> []P, $G |L> []P, $H" + lstar2 "$G |L> $H ==> P, $G |L> $H" + rstar0 "|R>" + rstar1 "$G |R> $H ==> <>P, $G |R> <>P, $H" + rstar2 "$G |R> $H ==> P, $G |R> $H" + +(* Rules for [] and <> *) + + boxR + "[| $E |L> $E'; $F |R> $F'; $G |R> $G'; \ +\ $E' |- $F', P, $G'|] ==> $E |- $F, []P, $G" + boxL "$E,P,$F,[]P |- $G ==> $E, []P, $F |- $G" + diaR "$E |- $F,P,$G,<>P ==> $E |- $F, <>P, $G" + diaL + "[| $E |L> $E'; $F |L> $F'; $G |R> $G'; \ +\ $E', P, $F' |- $G'|] ==> $E, <>P, $F |- $G" +end diff -r 000000000000 -r a5a9c433f639 src/Modal/S43.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/S43.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,79 @@ +(* Title: 91/Modal/S43 + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge + +This implements Rajeev Gore's sequent calculus for S43. +*) + +S43 = Modal0 + + +consts + S43pi :: "[sobj=>sobj, sobj=>sobj, sobj=>sobj,\ +\ sobj=>sobj, sobj=>sobj, sobj=>sobj] => prop" + "@S43pi" :: "[sequence, sequence, sequence, sequence, sequence,\ +\ sequence] => prop" ("S43pi((_);(_);(_);(_);(_);(_))" [] 5) + +rules +(* Definition of the star operation using a set of Horn clauses *) +(* For system S43: gamma * == {[]P | []P : gamma} *) +(* delta * == {<>P | <>P : delta} *) + + lstar0 "|L>" + lstar1 "$G |L> $H ==> []P, $G |L> []P, $H" + lstar2 "$G |L> $H ==> P, $G |L> $H" + rstar0 "|R>" + rstar1 "$G |R> $H ==> <>P, $G |R> <>P, $H" + rstar2 "$G |R> $H ==> P, $G |R> $H" + +(* Set of Horn clauses to generate the antecedents for the S43 pi rule *) +(* ie *) +(* S1...Sk,Sk+1...Sk+m *) +(* ---------------------------------- *) +(* <>P1...<>Pk, $G |- $H, []Q1...[]Qm *) +(* *) +(* where Si == <>P1...<>Pi-1,<>Pi+1,..<>Pk,Pi, $G * |- $H *, []Q1...[]Qm *) +(* and Sj == <>P1...<>Pk, $G * |- $H *, []Q1...[]Qj-1,[]Qj+1...[]Qm,Qj *) +(* and 1<=i<=k and kP,$L'; $L;; $R; $Lbox;$Rdia); $L',P,$L,$Lbox |- $R,$Rdia |] ==> \ +\ S43pi $L'; <>P,$L;; $R; $Lbox;$Rdia" + S43pi2 + "[| (S43pi $L';; []P,$R'; $R; $Lbox;$Rdia); $L',$Lbox |- $R',P,$R,$Rdia |] ==> \ +\ S43pi $L';; $R'; []P,$R; $Lbox;$Rdia" + +(* Rules for [] and <> for S43 *) + + boxL "$E, P, $F, []P |- $G ==> $E, []P, $F |- $G" + diaR "$E |- $F, P, $G, <>P ==> $E |- $F, <>P, $G" + pi1 + "[| $L1,<>P,$L2 |L> $Lbox; $L1,<>P,$L2 |R> $Ldia; $R |L> $Rbox; $R |R> $Rdia; \ +\ S43pi ; $Ldia;; $Rbox; $Lbox; $Rdia |] ==> \ +\ $L1, <>P, $L2 |- $R" + pi2 + "[| $L |L> $Lbox; $L |R> $Ldia; $R1,[]P,$R2 |L> $Rbox; $R1,[]P,$R2 |R> $Rdia; \ +\ S43pi ; $Ldia;; $Rbox; $Lbox; $Rdia |] ==> \ +\ $L |- $R1, []P, $R2" +end + +ML + +local + + val S43pi = "S43pi"; + val SS43pi = "@S43pi"; + + val tr = LK.seq_tr1; + val tr' = LK.seq_tr1'; + + fun s43pi_tr[s1,s2,s3,s4,s5,s6]= + Const(S43pi,dummyT)$tr s1$tr s2$tr s3$tr s4$tr s5$tr s6; + fun s43pi_tr'[Abs(_,_,s1),Abs(_,_,s2),Abs(_,_,s3), + Abs(_,_,s4),Abs(_,_,s5),Abs(_,_,s6)] = + Const(SS43pi,dummyT)$tr' s1$tr' s2$tr' s3$tr' s4$tr' s5$tr' s6; +in +val parse_translation = [(SS43pi,s43pi_tr)]; +val print_translation = [(S43pi,s43pi_tr')] +end diff -r 000000000000 -r a5a9c433f639 src/Modal/T.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/T.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,30 @@ +(* Title: 91/Modal/T + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge +*) + +T = Modal0 + +rules +(* Definition of the star operation using a set of Horn clauses *) +(* For system T: gamma * == {P | []P : gamma} *) +(* delta * == {P | <>P : delta} *) + + lstar0 "|L>" + lstar1 "$G |L> $H ==> []P, $G |L> P, $H" + lstar2 "$G |L> $H ==> P, $G |L> $H" + rstar0 "|R>" + rstar1 "$G |R> $H ==> <>P, $G |R> P, $H" + rstar2 "$G |R> $H ==> P, $G |R> $H" + +(* Rules for [] and <> *) + + boxR + "[| $E |L> $E'; $F |R> $F'; $G |R> $G'; \ +\ $E' |- $F', P, $G'|] ==> $E |- $F, []P, $G" + boxL "$E, P, $F |- $G ==> $E, []P, $F |- $G" + diaR "$E |- $F, P, $G ==> $E |- $F, <>P, $G" + diaL + "[| $E |L> $E'; $F |L> $F'; $G |R> $G'; \ +\ $E', P, $F'|- $G'|] ==> $E, <>P, $F |- $G" +end diff -r 000000000000 -r a5a9c433f639 src/Modal/ex/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/ex/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,28 @@ +(* Title: Modal/ex/ROOT + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge +*) + +Modal_build_completed; (*Cause examples to fail if Modal did*) + +proof_timing := true; + +writeln "\nTheorems of T\n"; +fun try s = (writeln s; prove_goal T.thy s (fn _=> [T_Prover.solve_tac 2])); +time_use "ex/Tthms.ML"; + +writeln "\nTheorems of S4\n"; +fun try s = (writeln s; prove_goal S4.thy s (fn _=> [S4_Prover.solve_tac 2])); +time_use "ex/Tthms.ML"; +time_use "ex/S4thms.ML"; + +writeln "\nTheorems of S43\n"; +fun try s = (writeln s; + prove_goal S43.thy s (fn _=> [S43_Prover.solve_tac 2 ORELSE + S43_Prover.solve_tac 3])); +time_use "ex/Tthms.ML"; +time_use "ex/S4thms.ML"; +time_use "ex/S43thms.ML"; + +maketest"END: Root file for Modal examples"; diff -r 000000000000 -r a5a9c433f639 src/Modal/ex/S43thms.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/ex/S43thms.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,49 @@ +(* Title: 91/Modal/ex/S43thms + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge +*) + +(* Theorems of system S43 *) + +try "|- <>[]P --> []<>P"; +try "|- <>[]P --> [][]<>P"; +try "|- [](<>P | <>Q) --> []<>P | []<>Q"; +try "|- <>[]P & <>[]Q --> <>([]P & []Q)"; +try "|- []([]P --> []Q) | []([]Q --> []P)"; +try "|- [](<>P --> <>Q) | [](<>Q --> <>P)"; +try "|- []([]P --> Q) | []([]Q --> P)"; +try "|- [](P --> <>Q) | [](Q --> <>P)"; +try "|- [](P --> []Q-->R) | [](P | ([]R --> Q))"; +try "|- [](P | (Q --> <>C)) | [](P --> C --> <>Q)"; +try "|- []([]P | Q) & [](P | []Q) --> []P | []Q"; +try "|- <>P & <>Q --> <>(<>P & Q) | <>(P & <>Q)"; +try "|- [](P | Q) & []([]P | Q) & [](P | []Q) --> []P | []Q"; +try "|- <>P & <>Q --> <>(P & Q) | <>(<>P & Q) | <>(P & <>Q)"; +try "|- <>[]<>P <-> []<>P"; +try "|- []<>[]P <-> <>[]P"; + +(* These are from Hailpern, LNCS 129 *) + +try "|- [](P & Q) <-> []P & []Q"; +try "|- <>(P | Q) <-> <>P | <>Q"; +try "|- <>(P --> Q) <-> []P --> <>Q"; + +try "|- [](P --> Q) --> <>P --> <>Q"; +try "|- []P --> []<>P"; +try "|- <>[]P --> <>P"; +try "|- []<>[]P --> []<>P"; +try "|- <>[]P --> <>[]<>P"; +try "|- <>[]P --> []<>P"; +try "|- []<>[]P <-> <>[]P"; +try "|- <>[]<>P <-> []<>P"; + +try "|- []P | []Q --> [](P | Q)"; +try "|- <>(P & Q) --> <>P & <>Q"; +try "|- [](P | Q) --> []P | <>Q"; +try "|- <>P & []Q --> <>(P & Q)"; +try "|- [](P | Q) --> <>P | []Q"; +try "|- [](P | Q) --> []<>P | []<>Q"; +try "|- <>[]P & <>[]Q --> <>(P & Q)"; +try "|- <>[](P & Q) <-> <>[]P & <>[]Q"; +try "|- []<>(P | Q) <-> []<>P | []<>Q"; diff -r 000000000000 -r a5a9c433f639 src/Modal/ex/S4thms.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/ex/S4thms.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,40 @@ +(* Title: 91/Modal/ex/S4thms + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge +*) + +(* Theorems of system S4 from Hughes and Cresswell, p.46 *) + +try "|- []A --> A"; (* refexivity *) +try "|- []A --> [][]A"; (* transitivity *) +try "|- []A --> <>A"; (* seriality *) +try "|- <>[](<>A --> []<>A)"; +try "|- <>[](<>[]A --> []A)"; +try "|- []P <-> [][]P"; +try "|- <>P <-> <><>P"; +try "|- <>[]<>P --> <>P"; +try "|- []<>P <-> []<>[]<>P"; +try "|- <>[]P <-> <>[]<>[]P"; + +(* Theorems for system S4 from Hughes and Cresswell, p.60 *) + +try "|- []P | []Q <-> []([]P | []Q)"; +try "|- ((P>- ((P>- []P & []Q"; +try "|- <>(P | Q) <-> <>P | <>Q"; +try "|- <>(P --> Q) <-> ([]P --> <>Q)"; + +try "|- [](P --> Q) --> (<>P --> <>Q)"; +try "|- []P --> []<>P"; +try "|- <>[]P --> <>P"; + +try "|- []P | []Q --> [](P | Q)"; +try "|- <>(P & Q) --> <>P & <>Q"; +try "|- [](P | Q) --> []P | <>Q"; +try "|- <>P & []Q --> <>(P & Q)"; +try "|- [](P | Q) --> <>P | []Q"; + diff -r 000000000000 -r a5a9c433f639 src/Modal/ex/Tthms.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/ex/Tthms.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,31 @@ +(* Title: 91/Modal/ex/Tthms + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge +*) + +(* Theorems of system T from Hughes and Cresswell and Hailpern, LNCS 129 *) + +try "|- []P --> P"; +try "|- [](P-->Q) --> ([]P-->[]Q)"; (* normality*) +try "|- (P-- []P --> []Q"; +try "|- P --> <>P"; + +try "|- [](P & Q) <-> []P & []Q"; +try "|- <>(P | Q) <-> <>P | <>Q"; +try "|- [](P<->Q) <-> (P>-(P-->Q) <-> ([]P--><>Q)"; +try "|- []P <-> ~<>(~P)"; +try "|- [](~P) <-> ~<>P"; +try "|- ~[]P <-> <>(~P)"; +try "|- [][]P <-> ~<><>(~P)"; +try "|- ~<>(P | Q) <-> ~<>P & ~<>Q"; + +try "|- []P | []Q --> [](P | Q)"; +try "|- <>(P & Q) --> <>P & <>Q"; +try "|- [](P | Q) --> []P | <>Q"; +try "|- <>P & []Q --> <>(P & Q)"; +try "|- [](P | Q) --> <>P | []Q"; +try "|- <>(P-->(Q & R)) --> ([]P --> <>Q) & ([]P--><>R)"; +try "|- (P-- (P-- <>Q --> <>(P & Q)"; diff -r 000000000000 -r a5a9c433f639 src/Modal/modal0.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/modal0.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,39 @@ +(* Title: 91/Modal/modal0 + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge +*) + +Modal0 = LK + + +consts + box :: "o=>o" ("[]_" [50] 50) + dia :: "o=>o" ("<>_" [50] 50) + "--<",">-<" :: "[o,o]=>o" (infixr 25) + "@Lstar" :: "[sequence,sequence]=>prop" ("(_)|L>(_)" [6,6] 5) + "@Rstar" :: "[sequence,sequence]=>prop" ("(_)|R>(_)" [6,6] 5) + Lstar,Rstar :: "[sobj=>sobj,sobj=>sobj]=>prop" + +rules + (* Definitions *) + + strimp_def "P --< Q == [](P --> Q)" + streqv_def "P >-< Q == (P --< Q) & (Q --< P)" +end + +ML + +local + + val Lstar = "Lstar"; + val Rstar = "Rstar"; + val SLstar = "@Lstar"; + val SRstar = "@Rstar"; + + fun star_tr c [s1,s2] = Const(c,dummyT)$LK.seq_tr1 s1$LK.seq_tr1 s2; + fun star_tr' c [Abs(_,_,s1),Abs(_,_,s2)] = + Const(c,dummyT) $ LK.seq_tr1' s1 $ LK.seq_tr1' s2; +in +val parse_translation = [(SLstar,star_tr Lstar), (SRstar,star_tr Rstar)]; +val print_translation = [(Lstar,star_tr' SLstar), (Rstar,star_tr' SRstar)] +end; diff -r 000000000000 -r a5a9c433f639 src/Modal/prover.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/prover.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,90 @@ +(* Title: 91/Modal/prover + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge +*) + +signature MODAL_PROVER_RULE = +sig + val rewrite_rls : thm list + val safe_rls : thm list + val unsafe_rls : thm list + val bound_rls : thm list + val aside_rls : thm list +end; + +signature MODAL_PROVER = +sig + val rule_tac : thm list -> int ->tactic + val step_tac : int -> tactic + val solven_tac : int -> int -> tactic + val solve_tac : int -> tactic +end; + +functor Modal_ProverFun (Modal_Rule: MODAL_PROVER_RULE) : MODAL_PROVER = +struct +local open Modal_Rule +in + +(*Returns the list of all formulas in the sequent*) +fun forms_of_seq (Const("Seqof",_) $ P $ u) = P :: forms_of_seq u + | forms_of_seq (H $ u) = forms_of_seq u + | forms_of_seq _ = []; + +(*Tests whether two sequences (left or right sides) could be resolved. + seqp is a premise (subgoal), seqc is a conclusion of an object-rule. + Assumes each formula in seqc is surrounded by sequence variables + -- checks that each concl formula looks like some subgoal formula.*) +fun could_res (seqp,seqc) = + forall (fn Qc => exists (fn Qp => could_unify (Qp,Qc)) + (forms_of_seq seqp)) + (forms_of_seq seqc); + +(*Tests whether two sequents G|-H could be resolved, comparing each side.*) +fun could_resolve_seq (prem,conc) = + case (prem,conc) of + (_ $ Abs(_,_,leftp) $ Abs(_,_,rightp), + _ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) => + could_res (leftp,leftc) andalso could_res (rightp,rightc) + | _ => false; + +(*Like filt_resolve_tac, using could_resolve_seq + Much faster than resolve_tac when there are many rules. + Resolve subgoal i using the rules, unless more than maxr are compatible. *) +fun filseq_resolve_tac rules maxr = SUBGOAL(fn (prem,i) => + let val rls = filter_thms could_resolve_seq (maxr+1, prem, rules) + in if length rls > maxr then no_tac else resolve_tac rls i + end); + +fun fresolve_tac rls n = filseq_resolve_tac rls 999 n; + +(* NB No back tracking possible with aside rules *) + +fun aside_tac n = DETERM(REPEAT (filt_resolve_tac aside_rls 999 n)); +fun rule_tac rls n = fresolve_tac rls n THEN aside_tac n; + +val fres_safe_tac = fresolve_tac safe_rls; +val fres_unsafe_tac = fresolve_tac unsafe_rls THEN' aside_tac; +val fres_bound_tac = fresolve_tac bound_rls; + +fun UPTOGOAL n tf = let fun tac i = if i tac(nprems_of state)) end; + +(* Depth first search bounded by d *) +fun solven_tac d n = STATE (fn state => + if d<0 then no_tac + else if (nprems_of state = 0) then all_tac + else (DETERM(fres_safe_tac n) THEN UPTOGOAL n (solven_tac d)) ORELSE + ((fres_unsafe_tac n THEN UPTOGOAL n (solven_tac d)) APPEND + (fres_bound_tac n THEN UPTOGOAL n (solven_tac (d-1))))); + +fun solve_tac d = rewrite_goals_tac rewrite_rls THEN solven_tac d 1; + +fun step_tac n = STATE (fn state => + if (nprems_of state = 0) then all_tac + else (DETERM(fres_safe_tac n)) ORELSE + (fres_unsafe_tac n APPEND fres_bound_tac n)); + +end; +end; diff -r 000000000000 -r a5a9c433f639 src/Modal/s4.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/s4.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,30 @@ +(* Title: 91/Modal/S4 + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge +*) + +S4 = Modal0 + +rules +(* Definition of the star operation using a set of Horn clauses *) +(* For system S4: gamma * == {[]P | []P : gamma} *) +(* delta * == {<>P | <>P : delta} *) + + lstar0 "|L>" + lstar1 "$G |L> $H ==> []P, $G |L> []P, $H" + lstar2 "$G |L> $H ==> P, $G |L> $H" + rstar0 "|R>" + rstar1 "$G |R> $H ==> <>P, $G |R> <>P, $H" + rstar2 "$G |R> $H ==> P, $G |R> $H" + +(* Rules for [] and <> *) + + boxR + "[| $E |L> $E'; $F |R> $F'; $G |R> $G'; \ +\ $E' |- $F', P, $G'|] ==> $E |- $F, []P, $G" + boxL "$E,P,$F,[]P |- $G ==> $E, []P, $F |- $G" + diaR "$E |- $F,P,$G,<>P ==> $E |- $F, <>P, $G" + diaL + "[| $E |L> $E'; $F |L> $F'; $G |R> $G'; \ +\ $E', P, $F' |- $G'|] ==> $E, <>P, $F |- $G" +end diff -r 000000000000 -r a5a9c433f639 src/Modal/s43.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/s43.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,79 @@ +(* Title: 91/Modal/S43 + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge + +This implements Rajeev Gore's sequent calculus for S43. +*) + +S43 = Modal0 + + +consts + S43pi :: "[sobj=>sobj, sobj=>sobj, sobj=>sobj,\ +\ sobj=>sobj, sobj=>sobj, sobj=>sobj] => prop" + "@S43pi" :: "[sequence, sequence, sequence, sequence, sequence,\ +\ sequence] => prop" ("S43pi((_);(_);(_);(_);(_);(_))" [] 5) + +rules +(* Definition of the star operation using a set of Horn clauses *) +(* For system S43: gamma * == {[]P | []P : gamma} *) +(* delta * == {<>P | <>P : delta} *) + + lstar0 "|L>" + lstar1 "$G |L> $H ==> []P, $G |L> []P, $H" + lstar2 "$G |L> $H ==> P, $G |L> $H" + rstar0 "|R>" + rstar1 "$G |R> $H ==> <>P, $G |R> <>P, $H" + rstar2 "$G |R> $H ==> P, $G |R> $H" + +(* Set of Horn clauses to generate the antecedents for the S43 pi rule *) +(* ie *) +(* S1...Sk,Sk+1...Sk+m *) +(* ---------------------------------- *) +(* <>P1...<>Pk, $G |- $H, []Q1...[]Qm *) +(* *) +(* where Si == <>P1...<>Pi-1,<>Pi+1,..<>Pk,Pi, $G * |- $H *, []Q1...[]Qm *) +(* and Sj == <>P1...<>Pk, $G * |- $H *, []Q1...[]Qj-1,[]Qj+1...[]Qm,Qj *) +(* and 1<=i<=k and kP,$L'; $L;; $R; $Lbox;$Rdia); $L',P,$L,$Lbox |- $R,$Rdia |] ==> \ +\ S43pi $L'; <>P,$L;; $R; $Lbox;$Rdia" + S43pi2 + "[| (S43pi $L';; []P,$R'; $R; $Lbox;$Rdia); $L',$Lbox |- $R',P,$R,$Rdia |] ==> \ +\ S43pi $L';; $R'; []P,$R; $Lbox;$Rdia" + +(* Rules for [] and <> for S43 *) + + boxL "$E, P, $F, []P |- $G ==> $E, []P, $F |- $G" + diaR "$E |- $F, P, $G, <>P ==> $E |- $F, <>P, $G" + pi1 + "[| $L1,<>P,$L2 |L> $Lbox; $L1,<>P,$L2 |R> $Ldia; $R |L> $Rbox; $R |R> $Rdia; \ +\ S43pi ; $Ldia;; $Rbox; $Lbox; $Rdia |] ==> \ +\ $L1, <>P, $L2 |- $R" + pi2 + "[| $L |L> $Lbox; $L |R> $Ldia; $R1,[]P,$R2 |L> $Rbox; $R1,[]P,$R2 |R> $Rdia; \ +\ S43pi ; $Ldia;; $Rbox; $Lbox; $Rdia |] ==> \ +\ $L |- $R1, []P, $R2" +end + +ML + +local + + val S43pi = "S43pi"; + val SS43pi = "@S43pi"; + + val tr = LK.seq_tr1; + val tr' = LK.seq_tr1'; + + fun s43pi_tr[s1,s2,s3,s4,s5,s6]= + Const(S43pi,dummyT)$tr s1$tr s2$tr s3$tr s4$tr s5$tr s6; + fun s43pi_tr'[Abs(_,_,s1),Abs(_,_,s2),Abs(_,_,s3), + Abs(_,_,s4),Abs(_,_,s5),Abs(_,_,s6)] = + Const(SS43pi,dummyT)$tr' s1$tr' s2$tr' s3$tr' s4$tr' s5$tr' s6; +in +val parse_translation = [(SS43pi,s43pi_tr)]; +val print_translation = [(S43pi,s43pi_tr')] +end diff -r 000000000000 -r a5a9c433f639 src/Modal/t.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Modal/t.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,30 @@ +(* Title: 91/Modal/T + ID: $Id$ + Author: Martin Coen + Copyright 1991 University of Cambridge +*) + +T = Modal0 + +rules +(* Definition of the star operation using a set of Horn clauses *) +(* For system T: gamma * == {P | []P : gamma} *) +(* delta * == {P | <>P : delta} *) + + lstar0 "|L>" + lstar1 "$G |L> $H ==> []P, $G |L> P, $H" + lstar2 "$G |L> $H ==> P, $G |L> $H" + rstar0 "|R>" + rstar1 "$G |R> $H ==> <>P, $G |R> P, $H" + rstar2 "$G |R> $H ==> P, $G |R> $H" + +(* Rules for [] and <> *) + + boxR + "[| $E |L> $E'; $F |R> $F'; $G |R> $G'; \ +\ $E' |- $F', P, $G'|] ==> $E |- $F, []P, $G" + boxL "$E, P, $F |- $G ==> $E, []P, $F |- $G" + diaR "$E |- $F, P, $G ==> $E |- $F, <>P, $G" + diaL + "[| $E |L> $E'; $F |L> $F'; $G |R> $G'; \ +\ $E', P, $F'|- $G'|] ==> $E, <>P, $F |- $G" +end diff -r 000000000000 -r a5a9c433f639 src/Provers/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Provers/README Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,12 @@ + Provers + +This directory contains ML sources of generic theorem proving tools. +Typically, they can be applied to various logics, provided rules of a +certain form are derivable. Unfortunately, little documentation is +available. + +classical.ML -- theorem prover for classical logics +genelim.ML -- bits and pieces for deriving elimination rules +ind.ML -- a simple induction package +simp.ML -- a powerful simplifier +typedsimp.ML -- a rather basic simplifier for explicitly typed logics diff -r 000000000000 -r a5a9c433f639 src/Provers/classical.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Provers/classical.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,206 @@ +(* Title: Provers/classical + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Theorem prover for classical reasoning, including predicate calculus, set +theory, etc. + +Rules must be classified as intr, elim, safe, hazardous. + +A rule is unsafe unless it can be applied blindly without harmful results. +For a rule to be safe, its premises and conclusion should be logically +equivalent. There should be no variables in the premises that are not in +the conclusion. +*) + +signature CLASSICAL_DATA = + sig + val mp: thm (* [| P-->Q; P |] ==> Q *) + val not_elim: thm (* [| ~P; P |] ==> R *) + val swap: thm (* ~P ==> (~Q ==> P) ==> Q *) + val sizef : thm -> int (* size function for BEST_FIRST *) + val hyp_subst_tacs: (int -> tactic) list + end; + +(*Higher precedence than := facilitates use of references*) +infix 4 addSIs addSEs addSDs addIs addEs addDs; + + +signature CLASSICAL = + sig + type claset + val empty_cs: claset + val addDs : claset * thm list -> claset + val addEs : claset * thm list -> claset + val addIs : claset * thm list -> claset + val addSDs: claset * thm list -> claset + val addSEs: claset * thm list -> claset + val addSIs: claset * thm list -> claset + val print_cs: claset -> unit + val rep_claset: claset -> + {safeIs: thm list, safeEs: thm list, hazIs: thm list, hazEs: thm list} + val best_tac : claset -> int -> tactic + val chain_tac : int -> tactic + val contr_tac : int -> tactic + val eq_mp_tac: int -> tactic + val fast_tac : claset -> int -> tactic + val joinrules : thm list * thm list -> (bool * thm) list + val mp_tac: int -> tactic + val safe_tac : claset -> tactic + val safe_step_tac : claset -> int -> tactic + val slow_step_tac : claset -> int -> tactic + val slow_best_tac : claset -> int -> tactic + val slow_tac : claset -> int -> tactic + val step_tac : claset -> int -> tactic + val swapify : thm list -> thm list + val swap_res_tac : thm list -> int -> tactic + val inst_step_tac : claset -> int -> tactic + end; + + +functor ClassicalFun(Data: CLASSICAL_DATA): CLASSICAL = +struct + +local open Data in + +(** Useful tactics for classical reasoning **) + +val imp_elim = make_elim mp; + +(*Solve goal that assumes both P and ~P. *) +val contr_tac = eresolve_tac [not_elim] THEN' assume_tac; + +(*Finds P-->Q and P in the assumptions, replaces implication by Q *) +fun mp_tac i = eresolve_tac ([not_elim,imp_elim]) i THEN assume_tac i; + +(*Like mp_tac but instantiates no variables*) +fun eq_mp_tac i = ematch_tac ([not_elim,imp_elim]) i THEN eq_assume_tac i; + +(*Creates rules to eliminate ~A, from rules to introduce A*) +fun swapify intrs = intrs RLN (2, [swap]); + +(*Uses introduction rules in the normal way, or on negated assumptions, + trying rules in order. *) +fun swap_res_tac rls = + let fun tacf rl = rtac rl ORELSE' etac (rl RSN (2,swap)) + in assume_tac ORELSE' contr_tac ORELSE' FIRST' (map tacf rls) + end; + +(*Given assumption P-->Q, reduces subgoal Q to P [deletes the implication!] *) +fun chain_tac i = + eresolve_tac [imp_elim] i THEN + (assume_tac (i+1) ORELSE contr_tac (i+1)); + +(*** Classical rule sets ***) + +type netpair = (int*(bool*thm)) Net.net * (int*(bool*thm)) Net.net; + +datatype claset = + CS of {safeIs : thm list, + safeEs : thm list, + hazIs : thm list, + hazEs : thm list, + safe0_netpair : netpair, + safep_netpair : netpair, + haz_netpair : netpair}; + +fun rep_claset (CS{safeIs,safeEs,hazIs,hazEs,...}) = + {safeIs=safeIs, safeEs=safeEs, hazIs=hazIs, hazEs=hazEs}; + +(*For use with biresolve_tac. Combines intrs with swap to catch negated + assumptions; pairs elims with true; sorts. *) +fun joinrules (intrs,elims) = + sort lessb + (map (pair true) (elims @ swapify intrs) @ + map (pair false) intrs); + +(*Make a claset from the four kinds of rules*) +fun make_cs {safeIs,safeEs,hazIs,hazEs} = + let val (safe0_brls, safep_brls) = (*0 subgoals vs 1 or more*) + take_prefix (fn brl => subgoals_of_brl brl=0) + (joinrules(safeIs, safeEs)) + in CS{safeIs = safeIs, + safeEs = safeEs, + hazIs = hazIs, + hazEs = hazEs, + safe0_netpair = build_netpair safe0_brls, + safep_netpair = build_netpair safep_brls, + haz_netpair = build_netpair (joinrules(hazIs, hazEs))} + end; + +(*** Manipulation of clasets ***) + +val empty_cs = make_cs{safeIs=[], safeEs=[], hazIs=[], hazEs=[]}; + +fun print_cs (CS{safeIs,safeEs,hazIs,hazEs,...}) = + (writeln"Introduction rules"; prths hazIs; + writeln"Safe introduction rules"; prths safeIs; + writeln"Elimination rules"; prths hazEs; + writeln"Safe elimination rules"; prths safeEs; + ()); + +fun (CS{safeIs,safeEs,hazIs,hazEs,...}) addSIs ths = + make_cs {safeIs=ths@safeIs, safeEs=safeEs, hazIs=hazIs, hazEs=hazEs}; + +fun (CS{safeIs,safeEs,hazIs,hazEs,...}) addSEs ths = + make_cs {safeIs=safeIs, safeEs=ths@safeEs, hazIs=hazIs, hazEs=hazEs}; + +fun cs addSDs ths = cs addSEs (map make_elim ths); + +fun (CS{safeIs,safeEs,hazIs,hazEs,...}) addIs ths = + make_cs {safeIs=safeIs, safeEs=safeEs, hazIs=ths@hazIs, hazEs=hazEs}; + +fun (CS{safeIs,safeEs,hazIs,hazEs,...}) addEs ths = + make_cs {safeIs=safeIs, safeEs=safeEs, hazIs=hazIs, hazEs=ths@hazEs}; + +fun cs addDs ths = cs addEs (map make_elim ths); + +(*** Simple tactics for theorem proving ***) + +(*Attack subgoals using safe inferences -- matching, not resolution*) +fun safe_step_tac (CS{safe0_netpair,safep_netpair,...}) = + FIRST' [eq_assume_tac, + eq_mp_tac, + bimatch_from_nets_tac safe0_netpair, + FIRST' hyp_subst_tacs, + bimatch_from_nets_tac safep_netpair] ; + +(*Repeatedly attack subgoals using safe inferences -- it's deterministic!*) +fun safe_tac cs = DETERM (REPEAT_FIRST (safe_step_tac cs)); + +(*These steps could instantiate variables and are therefore unsafe.*) +fun inst_step_tac (CS{safe0_netpair,safep_netpair,...}) = + assume_tac APPEND' + contr_tac APPEND' + biresolve_from_nets_tac safe0_netpair APPEND' + biresolve_from_nets_tac safep_netpair; + +(*Single step for the prover. FAILS unless it makes progress. *) +fun step_tac (cs as (CS{haz_netpair,...})) i = + FIRST [safe_tac cs, + inst_step_tac cs i, + biresolve_from_nets_tac haz_netpair i]; + +(*Using a "safe" rule to instantiate variables is unsafe. This tactic + allows backtracking from "safe" rules to "unsafe" rules here.*) +fun slow_step_tac (cs as (CS{haz_netpair,...})) i = + safe_tac cs ORELSE + (inst_step_tac cs i APPEND biresolve_from_nets_tac haz_netpair i); + +(*** The following tactics all fail unless they solve one goal ***) + +(*Dumb but fast*) +fun fast_tac cs = SELECT_GOAL (DEPTH_SOLVE (step_tac cs 1)); + +(*Slower but smarter than fast_tac*) +fun best_tac cs = + SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, sizef) (step_tac cs 1)); + +fun slow_tac cs = SELECT_GOAL (DEPTH_SOLVE (slow_step_tac cs 1)); + +fun slow_best_tac cs = + SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, sizef) (slow_step_tac cs 1)); + +end; +end; diff -r 000000000000 -r a5a9c433f639 src/Provers/genelim.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Provers/genelim.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,18 @@ +(** Generalized elimination rules **) + +(*Generalized elimination for two conclusions*) +val prems = goal pure_thy + "[| PROP U ==> PROP VA; \ +\ PROP U ==> PROP VB; \ +\ PROP U; \ +\ [| PROP VA; PROP VB |] ==> PROP W \ +\ |] ==> PROP W"; +by (REPEAT (resolve_tac prems 1)); +val general_elim2_rl = result(); + +fun make_elim2 (rl1,rl2) = standard (rl2 COMP rl1 COMP general_elim2_rl); +fun elim2_tac (rl1,rl2) = eresolve_tac [rl2 COMP rl1 COMP general_elim2_rl]; + + +(*For example, make_elim2(conjunct1,conjunct2) + yields conjunction elimination *) diff -r 000000000000 -r a5a9c433f639 src/Provers/hypsubst.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Provers/hypsubst.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,112 @@ +(* Title: Provers/hypsubst + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Martin Coen's tactic for substitution in the hypotheses +*) + +signature HYPSUBST_DATA = + sig + val dest_eq: term -> term*term + val imp_intr: thm (* (P ==> Q) ==> P-->Q *) + val rev_cut_eq: thm (* [| a=b; a=b ==> R |] ==> R *) + val rev_mp: thm (* [| P; P-->Q |] ==> Q *) + val subst: thm (* [| a=b; P(a) |] ==> P(b) *) + val sym: thm (* a=b ==> b=a *) + end; + +signature HYPSUBST = + sig + val bound_hyp_subst_tac : int -> tactic + val hyp_subst_tac : int -> tactic + (*exported purely for debugging purposes*) + val eq_var : bool -> term -> term * thm + val inspect_pair : bool -> term * term -> term * thm + val liftvar : int -> typ list -> term + end; + +functor HypsubstFun(Data: HYPSUBST_DATA): HYPSUBST = +struct + +local open Data in + +fun REPEATN 0 tac = all_tac + | REPEATN n tac = Tactic(fn state => + tapply(tac THEN REPEATN (n-1) tac, state)); + +local + val T = case #1 (types_sorts rev_cut_eq) ("a",0) of + Some T => T + | None => error"No such variable in rev_cut_eq" +in + fun liftvar inc paramTs = Var(("a",inc), paramTs ---> incr_tvar inc T); +end; + +exception EQ_VAR; + +fun loose (i,t) = 0 mem add_loose_bnos(t,i,[]); + +(*It's not safe to substitute for a constant; consider 0=1. + It's not safe to substitute for x=t[x] since x is not eliminated. + It's not safe to substitute for a variable free in the premises, + but how could we check for this?*) +fun inspect_pair bnd (t,u) = + case (Pattern.eta_contract t, Pattern.eta_contract u) of + (Bound i, _) => if loose(i,u) then raise Match + else (t, asm_rl) + | (_, Bound i) => if loose(i,t) then raise Match + else (u, sym) + | (Free _, _) => if bnd orelse Logic.occs(t,u) then raise Match + else (t, asm_rl) + | (_, Free _) => if bnd orelse Logic.occs(u,t) then raise Match + else (u, sym) + | _ => raise Match; + + (* Extracts the name of the variable on the left (resp. right) of an equality + assumption. Returns the rule asm_rl (resp. sym). *) +fun eq_var bnd (Const("all",_) $ Abs(_,_,t)) = eq_var bnd t + | eq_var bnd (Const("==>",_) $ A $ B) = + (inspect_pair bnd (dest_eq A) + (*Match comes from inspect_pair or dest_eq*) + handle Match => eq_var bnd B) + | eq_var bnd _ = raise EQ_VAR; + +(*Lift and instantiate a rule wrt the given state and subgoal number *) +fun lift_instpair (state, i, t, rule) = + let val {maxidx,sign,...} = rep_thm state + val (_, _, Bi, _) = dest_state(state,i) + val params = Logic.strip_params Bi + val var = liftvar (maxidx+1) (map #2 params) + and u = Unify.rlist_abs(rev params, t) + and cterm = Sign.cterm_of sign + in cterm_instantiate [(cterm var, cterm u)] (lift_rule (state,i) rule) + end; + +fun eres_instpair_tac t rule i = STATE (fn state => + compose_tac (true, lift_instpair (state, i, t, rule), + length(prems_of rule)) i); + +val ssubst = sym RS subst; + +(*Select a suitable equality assumption and substitute throughout the subgoal + Replaces only Bound variables if bnd is true*) +fun gen_hyp_subst_tac bnd i = DETERM (STATE(fn state => + let val (_,_,Bi,_) = dest_state(state,i) + val n = length(Logic.strip_assums_hyp Bi) - 1 + val (t,symopt) = eq_var bnd Bi + in eres_instpair_tac t (symopt RS rev_cut_eq) i THEN + REPEATN n (etac rev_mp i) THEN + etac ssubst i THEN REPEATN n (rtac imp_intr i) + end + handle THM _ => no_tac | EQ_VAR => no_tac)); + +(*Substitutes for Free or Bound variables*) +val hyp_subst_tac = gen_hyp_subst_tac false; + +(*Substitutes for Bound variables only -- this is always safe*) +val bound_hyp_subst_tac = gen_hyp_subst_tac true; + +end +end; + diff -r 000000000000 -r a5a9c433f639 src/Provers/ind.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Provers/ind.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,65 @@ +(* Title: Provers/ind + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +Generic induction package -- for use with simplifier +*) + +signature IND_DATA = + sig + val spec: thm (* All(?P) ==> ?P(?a) *) + end; + + +signature IND = + sig + val all_frees_tac: string -> int -> tactic + val ALL_IND_TAC: thm -> (int -> tactic) -> (int -> tactic) + val IND_TAC: thm -> (int -> tactic) -> string -> (int -> tactic) + end; + + +functor InductionFun(Ind_Data: IND_DATA):IND = +struct +local open Ind_Data in + +val _$(_$Var(a_ixname,aT)) = concl_of spec; +val a_name = implode(tl(explode(Syntax.string_of_vname a_ixname))); + +fun add_term_frees tsig = +let fun add(tm, vars) = case tm of + Free(v,T) => if Type.typ_instance(tsig,T,aT) then v ins vars + else vars + | Abs (_,_,body) => add(body,vars) + | rator$rand => add(rator, add(rand, vars)) + | _ => vars +in add end; + + +fun qnt_tac i = fn (tac,var) => tac THEN res_inst_tac [(a_name,var)] spec i; + +(*Generalizes over all free variables, with the named var outermost.*) +fun all_frees_tac (var:string) i = Tactic(fn thm => + let val tsig = #tsig(Sign.rep_sg(#sign(rep_thm thm))); + val frees = add_term_frees tsig (nth_elem(i-1,prems_of thm),[var]); + val frees' = sort(op>)(frees \ var) @ [var] + in tapply(foldl (qnt_tac i) (all_tac,frees'), thm) end); + +fun REPEAT_SIMP_TAC simp_tac n i = +let fun repeat thm = tapply(COND (has_fewer_prems n) all_tac + let val k = length(prems_of thm) + in simp_tac i THEN COND (has_fewer_prems k) (Tactic repeat) all_tac + end, thm) +in Tactic repeat end; + +fun ALL_IND_TAC sch simp_tac i = Tactic(fn thm => tapply( + resolve_tac [sch] i THEN + REPEAT_SIMP_TAC simp_tac (length(prems_of thm)) i, thm)); + +fun IND_TAC sch simp_tac var = + all_frees_tac var THEN' ALL_IND_TAC sch simp_tac; + + +end +end; diff -r 000000000000 -r a5a9c433f639 src/Provers/simp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Provers/simp.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,650 @@ +(* Title: Provers/simp + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 University of Cambridge + +Generic simplifier, suitable for most logics. The only known exception is +Constructive Type Theory. The reflexivity axiom must be unconditional, +namely a=a not a:A ==> a=a:A. Used typedsimp.ML otherwise. +*) + +signature SIMP_DATA = +sig + val dest_red : term -> term * term * term + val mk_rew_rules : thm -> thm list + val norm_thms : (thm*thm) list (* [(?x>>norm(?x), norm(?x)>>?x), ...] *) + val red1 : thm (* ?P>>?Q ==> ?P ==> ?Q *) + val red2 : thm (* ?P>>?Q ==> ?Q ==> ?P *) + val refl_thms : thm list + val subst_thms : thm list (* [ ?a>>?b ==> ?P(?a) ==> ?P(?b), ...] *) + val trans_thms : thm list +end; + + +infix 4 addrews addcongs addsplits delrews delcongs setauto; + +signature SIMP = +sig + type simpset + val empty_ss : simpset + val addcongs : simpset * thm list -> simpset + val addrews : simpset * thm list -> simpset + val addsplits : simpset * thm list -> simpset + val delcongs : simpset * thm list -> simpset + val delrews : simpset * thm list -> simpset + val dest_ss : simpset -> thm list * thm list + val print_ss : simpset -> unit + val setauto : simpset * (thm list -> int -> tactic) -> simpset + val ASM_SIMP_TAC : simpset -> int -> tactic + val SPLIT_TAC : simpset -> int -> tactic + val SIMP_SPLIT2_TAC : simpset -> int -> tactic + val SIMP_THM : simpset -> thm -> thm + val SIMP_TAC : simpset -> int -> tactic + val mk_congs : theory -> string list -> thm list + val mk_typed_congs : theory -> (string * string) list -> thm list +(* temporarily disabled: + val extract_free_congs : unit -> thm list +*) + val tracing : bool ref +end; + +functor SimpFun (Simp_data: SIMP_DATA) : SIMP = +struct + +local open Simp_data Logic in + +(*For taking apart reductions into left, right hand sides*) +val lhs_of = #2 o dest_red; +val rhs_of = #3 o dest_red; + +(*** Indexing and filtering of theorems ***) + +fun eq_brl ((b1,th1),(b2,th2)) = b1=b2 andalso eq_thm(th1,th2); + +(*insert a thm in a discrimination net by its lhs*) +fun lhs_insert_thm (th,net) = + Net.insert_term((lhs_of (concl_of th), (false,th)), net, eq_brl) + handle Net.INSERT => net; + +(*match subgoal i against possible theorems in the net. + Similar to match_from_nat_tac, but the net does not contain numbers; + rewrite rules are not ordered.*) +fun net_tac net = + SUBGOAL(fn (prem,i) => + match_tac (Net.match_term net (strip_assums_concl prem)) i); + +(*match subgoal i against possible theorems indexed by lhs in the net*) +fun lhs_net_tac net = + SUBGOAL(fn (prem,i) => + bimatch_tac (Net.match_term net + (lhs_of (strip_assums_concl prem))) i); + +fun nth_subgoal i thm = nth_elem(i-1,prems_of thm); + +fun goal_concl i thm = strip_assums_concl(nth_subgoal i thm); + +fun lhs_of_eq i thm = lhs_of(goal_concl i thm) +and rhs_of_eq i thm = rhs_of(goal_concl i thm); + +fun var_lhs(thm,i) = +let fun var(Var _) = true + | var(Abs(_,_,t)) = var t + | var(f$_) = var f + | var _ = false; +in var(lhs_of_eq i thm) end; + +fun contains_op opns = + let fun contains(Const(s,_)) = s mem opns | + contains(s$t) = contains s orelse contains t | + contains(Abs(_,_,t)) = contains t | + contains _ = false; + in contains end; + +fun may_match(match_ops,i) = contains_op match_ops o lhs_of_eq i; + +val (normI_thms,normE_thms) = split_list norm_thms; + +(*Get the norm constants from norm_thms*) +val norms = + let fun norm thm = + case lhs_of(concl_of thm) of + Const(n,_)$_ => n + | _ => (prths normE_thms; error"No constant in lhs of a norm_thm") + in map norm normE_thms end; + +fun lhs_is_NORM(thm,i) = case lhs_of_eq i thm of + Const(s,_)$_ => s mem norms | _ => false; + +val refl_tac = resolve_tac refl_thms; + +fun find_res thms thm = + let fun find [] = (prths thms; error"Check Simp_Data") + | find(th::thms) = thm RS th handle _ => find thms + in find thms end; + +val mk_trans = find_res trans_thms; + +fun mk_trans2 thm = +let fun mk[] = error"Check transitivity" + | mk(t::ts) = (thm RSN (2,t)) handle _ => mk ts +in mk trans_thms end; + +(*Applies tactic and returns the first resulting state, FAILS if none!*) +fun one_result(tac,thm) = case Sequence.pull(tapply(tac,thm)) of + Some(thm',_) => thm' + | None => raise THM("Simplifier: could not continue", 0, [thm]); + +fun res1(thm,thms,i) = one_result(resolve_tac thms i,thm); + + +(**** Adding "NORM" tags ****) + +(*get name of the constant from conclusion of a congruence rule*) +fun cong_const cong = + case head_of (lhs_of (concl_of cong)) of + Const(c,_) => c + | _ => "" (*a placeholder distinct from const names*); + +(*true if the term is an atomic proposition (no ==> signs) *) +val atomic = null o strip_assums_hyp; + +(*ccs contains the names of the constants possessing congruence rules*) +fun add_hidden_vars ccs = + let fun add_hvars(tm,hvars) = case tm of + Abs(_,_,body) => add_term_vars(body,hvars) + | _$_ => let val (f,args) = strip_comb tm + in case f of + Const(c,T) => + if c mem ccs + then foldr add_hvars (args,hvars) + else add_term_vars(tm,hvars) + | _ => add_term_vars(tm,hvars) + end + | _ => hvars; + in add_hvars end; + +fun add_new_asm_vars new_asms = + let fun itf((tm,at),vars) = + if at then vars else add_term_vars(tm,vars) + fun add_list(tm,al,vars) = let val (_,tml) = strip_comb tm + in if length(tml)=length(al) + then foldr itf (tml~~al,vars) + else vars + end + fun add_vars (tm,vars) = case tm of + Abs (_,_,body) => add_vars(body,vars) + | r$s => (case head_of tm of + Const(c,T) => (case assoc(new_asms,c) of + None => add_vars(r,add_vars(s,vars)) + | Some(al) => add_list(tm,al,vars)) + | _ => add_vars(r,add_vars(s,vars))) + | _ => vars + in add_vars end; + + +fun add_norms(congs,ccs,new_asms) thm = +let val thm' = mk_trans2 thm; +(* thm': [?z -> l; Prems; r -> ?t] ==> ?z -> ?t *) + val nops = nprems_of thm' + val lhs = rhs_of_eq 1 thm' + val rhs = lhs_of_eq nops thm' + val asms = tl(rev(tl(prems_of thm'))) + val hvars = foldr (add_hidden_vars ccs) (lhs::rhs::asms,[]) + val hvars = add_new_asm_vars new_asms (rhs,hvars) + fun it_asms (asm,hvars) = + if atomic asm then add_new_asm_vars new_asms (asm,hvars) + else add_term_frees(asm,hvars) + val hvars = foldr it_asms (asms,hvars) + val hvs = map (#1 o dest_Var) hvars + val refl1_tac = refl_tac 1 + val add_norm_tac = DEPTH_FIRST (has_fewer_prems nops) + (STATE(fn thm => + case head_of(rhs_of_eq 1 thm) of + Var(ixn,_) => if ixn mem hvs then refl1_tac + else resolve_tac normI_thms 1 ORELSE refl1_tac + | Const _ => resolve_tac normI_thms 1 ORELSE + resolve_tac congs 1 ORELSE refl1_tac + | Free _ => resolve_tac congs 1 ORELSE refl1_tac + | _ => refl1_tac)) + val Some(thm'',_) = Sequence.pull(tapply(add_norm_tac,thm')) +in thm'' end; + +fun add_norm_tags congs = + let val ccs = map cong_const congs + val new_asms = filter (exists not o #2) + (ccs ~~ (map (map atomic o prems_of) congs)); + in add_norms(congs,ccs,new_asms) end; + +fun normed_rews congs = + let val add_norms = add_norm_tags congs; + in fn thm => map (varifyT o add_norms o mk_trans) (mk_rew_rules(freezeT thm)) + end; + +fun NORM norm_lhs_tac = EVERY'[resolve_tac [red2], norm_lhs_tac, refl_tac]; + +val trans_norms = map mk_trans normE_thms; + + +(* SIMPSET *) + +datatype simpset = + SS of {auto_tac: thm list -> int -> tactic, + congs: thm list, + cong_net: thm Net.net, + mk_simps: thm -> thm list, + simps: (thm * thm list) list, + simp_net: thm Net.net, + splits: thm list, + split_consts: string list} + +val empty_ss = SS{auto_tac= K (K no_tac), congs=[], cong_net=Net.empty, + mk_simps=normed_rews[], simps=[], simp_net=Net.empty, + splits=[], split_consts=[]}; + +(** Insertion of congruences, rewrites and case splits **) + +(*insert a thm in a thm net*) +fun insert_thm_warn (th,net) = + Net.insert_term((concl_of th, th), net, eq_thm) + handle Net.INSERT => + (writeln"\nDuplicate rewrite or congruence rule:"; print_thm th; + net); + +val insert_thms = foldr insert_thm_warn; + +fun addrew(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net, + splits,split_consts}, thm) = +let val thms = mk_simps thm +in SS{auto_tac=auto_tac,congs=congs, cong_net=cong_net, mk_simps=mk_simps, + simps = (thm,thms)::simps, simp_net = insert_thms(thms,simp_net), + splits=splits,split_consts=split_consts} +end; + +val op addrews = foldl addrew; + +fun op addcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net, + splits,split_consts}, thms) = +let val congs' = thms @ congs; +in SS{auto_tac=auto_tac, congs= congs', + cong_net= insert_thms (map mk_trans thms,cong_net), + mk_simps= normed_rews congs', simps=simps, simp_net=simp_net, + splits=splits,split_consts=split_consts} +end; + +fun split_err() = error("split rule not of the form ?P(c(...)) = ..."); + +fun split_const(_ $ t) = + (case head_of t of Const(a,_) => a | _ => split_err()) + | split_const _ = split_err(); + +fun addsplit(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net, + splits,split_consts}, thm) = +let val a = split_const(lhs_of(concl_of thm)) +in SS{auto_tac=auto_tac,congs=congs,cong_net=cong_net, + mk_simps=mk_simps,simps=simps,simp_net=simp_net, + splits=splits@[mk_trans thm],split_consts=split_consts@[a]} end; + +val op addsplits = foldl addsplit; + +(** Deletion of congruences and rewrites **) + +(*delete a thm from a thm net*) +fun delete_thm_warn (th,net) = + Net.delete_term((concl_of th, th), net, eq_thm) + handle Net.DELETE => + (writeln"\nNo such rewrite or congruence rule:"; print_thm th; + net); + +val delete_thms = foldr delete_thm_warn; + +fun op delcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net, + splits,split_consts}, thms) = +let val congs' = foldl (gen_rem eq_thm) (congs,thms) +in SS{auto_tac=auto_tac, congs= congs', + cong_net= delete_thms(map mk_trans thms,cong_net), + mk_simps= normed_rews congs', simps=simps, simp_net=simp_net, + splits=splits,split_consts=split_consts} +end; + +fun delrew(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net, + splits,split_consts}, thm) = +let fun find((p as (th,ths))::ps',ps) = + if eq_thm(thm,th) then (ths,ps@ps') else find(ps',p::ps) + | find([],simps') = (writeln"\nNo such rewrite or congruence rule:"; + print_thm thm; + ([],simps')) + val (thms,simps') = find(simps,[]) +in SS{auto_tac=auto_tac, congs=congs, cong_net=cong_net, mk_simps=mk_simps, + simps = simps', simp_net = delete_thms(thms,simp_net), + splits=splits,split_consts=split_consts} +end; + +val op delrews = foldl delrew; + + +fun op setauto(SS{congs,cong_net,mk_simps,simps,simp_net, + splits,split_consts,...}, auto_tac) = + SS{auto_tac=auto_tac, congs=congs, cong_net=cong_net, mk_simps=mk_simps, + simps=simps, simp_net=simp_net,splits=splits,split_consts=split_consts}; + + +(** Inspection of a simpset **) + +fun dest_ss(SS{congs,simps,...}) = (congs, map #1 simps); + +fun print_ss(SS{congs,simps,splits,...}) = + (writeln"Congruences:"; prths congs; + writeln"Case Splits"; prths splits; + writeln"Rewrite Rules:"; prths (map #1 simps); ()); + + +(* Rewriting with case splits *) + +fun splittable a i thm = + let val tm = goal_concl i thm + fun nobound(Abs(_,_,tm),j,k) = nobound(tm,j,k+1) + | nobound(s$t,j,k) = nobound(s,j,k) andalso nobound(t,j,k) + | nobound(Bound n,j,k) = n < k orelse k+j <= n + | nobound(_) = true; + fun check_args(al,j) = forall (fn t => nobound(t,j,0)) al + fun find_if(Abs(_,_,tm),j) = find_if(tm,j+1) + | find_if(tm as s$t,j) = let val (f,al) = strip_comb tm in + case f of Const(c,_) => if c=a then check_args(al,j) + else find_if(s,j) orelse find_if(t,j) + | _ => find_if(s,j) orelse find_if(t,j) end + | find_if(_) = false; + in find_if(tm,0) end; + +fun split_tac (cong_tac,splits,split_consts) i = + let fun seq_try (split::splits,a::bs) thm = tapply( + COND (splittable a i) (DETERM(resolve_tac[split]i)) + (Tactic(seq_try(splits,bs))), thm) + | seq_try([],_) thm = tapply(no_tac,thm) + and try_rew thm = tapply(Tactic(seq_try(splits,split_consts)) + ORELSE Tactic one_subt, thm) + and one_subt thm = + let val test = has_fewer_prems (nprems_of thm + 1) + fun loop thm = tapply(COND test no_tac + ((Tactic try_rew THEN DEPTH_FIRST test (refl_tac i)) + ORELSE (refl_tac i THEN Tactic loop)), thm) + in tapply(cong_tac THEN Tactic loop, thm) end + in if null splits then no_tac + else COND (may_match(split_consts,i)) (Tactic try_rew) no_tac + end; + +fun SPLIT_TAC (SS{cong_net,splits,split_consts,...}) i = +let val cong_tac = net_tac cong_net i +in NORM (split_tac (cong_tac,splits,split_consts)) i end; + +(* Rewriting Automaton *) + +datatype cntrl = STOP | MK_EQ | ASMS of int | SIMP_LHS | REW | REFL | TRUE + | PROVE | POP_CS | POP_ARTR | SPLIT; +(* +fun pr_cntrl c = case c of STOP => prs("STOP") | MK_EQ => prs("MK_EQ") | +ASMS i => print_int i | POP_ARTR => prs("POP_ARTR") | +SIMP_LHS => prs("SIMP_LHS") | REW => prs("REW") | REFL => prs("REFL") | +TRUE => prs("TRUE") | PROVE => prs("PROVE") | POP_CS => prs("POP_CS") | SPLIT +=> prs("SPLIT"); +*) +fun simp_refl([],_,ss) = ss + | simp_refl(a'::ns,a,ss) = if a'=a then simp_refl(ns,a,SIMP_LHS::REFL::ss) + else simp_refl(ns,a,ASMS(a)::SIMP_LHS::REFL::POP_ARTR::ss); + +(** Tracing **) + +val tracing = ref false; + +(*Replace parameters by Free variables in P*) +fun variants_abs ([],P) = P + | variants_abs ((a,T)::aTs, P) = + variants_abs (aTs, #2 (variant_abs(a,T,P))); + +(*Select subgoal i from proof state; substitute parameters, for printing*) +fun prepare_goal i st = + let val subgi = nth_subgoal i st + val params = rev(strip_params subgi) + in variants_abs (params, strip_assums_concl subgi) end; + +(*print lhs of conclusion of subgoal i*) +fun pr_goal_lhs i st = + writeln (Sign.string_of_term (#sign(rep_thm st)) + (lhs_of (prepare_goal i st))); + +(*print conclusion of subgoal i*) +fun pr_goal_concl i st = + writeln (Sign.string_of_term (#sign(rep_thm st)) (prepare_goal i st)) + +(*print subgoals i to j (inclusive)*) +fun pr_goals (i,j) st = + if i>j then () + else (pr_goal_concl i st; pr_goals (i+1,j) st); + +(*Print rewrite for tracing; i=subgoal#, n=number of new subgoals, + thm=old state, thm'=new state *) +fun pr_rew (i,n,thm,thm',not_asms) = + if !tracing + then (if not_asms then () else writeln"Assumption used in"; + pr_goal_lhs i thm; writeln"->"; pr_goal_lhs (i+n) thm'; + if n>0 then (writeln"Conditions:"; pr_goals (i, i+n-1) thm') + else (); + writeln"" ) + else (); + +(* Skip the first n hyps of a goal, and return the rest in generalized form *) +fun strip_varify(Const("==>", _) $ H $ B, n, vs) = + if n=0 then subst_bounds(vs,H)::strip_varify(B,0,vs) + else strip_varify(B,n-1,vs) + | strip_varify(Const("all",_)$Abs(_,T,t), n, vs) = + strip_varify(t,n,Var(("?",length vs),T)::vs) + | strip_varify _ = []; + +fun execute(ss,if_fl,auto_tac,cong_tac,splits,split_consts,net,i) thm = let + +fun simp_lhs(thm,ss,anet,ats,cs) = + if var_lhs(thm,i) then (ss,thm,anet,ats,cs) else + if lhs_is_NORM(thm,i) then (ss, res1(thm,trans_norms,i), anet,ats,cs) + else case Sequence.pull(tapply(cong_tac i,thm)) of + Some(thm',_) => + let val ps = prems_of thm and ps' = prems_of thm'; + val n = length(ps')-length(ps); + val a = length(strip_assums_hyp(nth_elem(i-1,ps))) + val l = map (fn p => length(strip_assums_hyp(p))) + (take(n,drop(i-1,ps'))); + in (simp_refl(rev(l),a,REW::ss),thm',anet,ats,cs) end + | None => (REW::ss,thm,anet,ats,cs); + +(*NB: the "Adding rewrites:" trace will look strange because assumptions + are represented by rules, generalized over their parameters*) +fun add_asms(ss,thm,a,anet,ats,cs) = + let val As = strip_varify(nth_subgoal i thm, a, []); + val thms = map (trivial o Sign.cterm_of(#sign(rep_thm(thm))))As; + val new_rws = flat(map mk_rew_rules thms); + val rwrls = map mk_trans (flat(map mk_rew_rules thms)); + val anet' = foldr lhs_insert_thm (rwrls,anet) + in if !tracing andalso not(null new_rws) + then (writeln"Adding rewrites:"; prths new_rws; ()) + else (); + (ss,thm,anet',anet::ats,cs) + end; + +fun rew(seq,thm,ss,anet,ats,cs, more) = case Sequence.pull seq of + Some(thm',seq') => + let val n = (nprems_of thm') - (nprems_of thm) + in pr_rew(i,n,thm,thm',more); + if n=0 then (SIMP_LHS::ss, thm', anet, ats, cs) + else ((replicate n PROVE) @ (POP_CS::SIMP_LHS::ss), + thm', anet, ats, (ss,thm,anet,ats,seq',more)::cs) + end + | None => if more + then rew(tapply(lhs_net_tac anet i THEN assume_tac i,thm), + thm,ss,anet,ats,cs,false) + else (ss,thm,anet,ats,cs); + +fun try_true(thm,ss,anet,ats,cs) = + case Sequence.pull(tapply(auto_tac i,thm)) of + Some(thm',_) => (ss,thm',anet,ats,cs) + | None => let val (ss0,thm0,anet0,ats0,seq,more)::cs0 = cs + in if !tracing + then (writeln"*** Failed to prove precondition. Normal form:"; + pr_goal_concl i thm; writeln"") + else (); + rew(seq,thm0,ss0,anet0,ats0,cs0,more) + end; + +fun split(thm,ss,anet,ats,cs) = + case Sequence.pull(tapply(split_tac + (cong_tac i,splits,split_consts) i,thm)) of + Some(thm',_) => (SIMP_LHS::SPLIT::ss,thm',anet,ats,cs) + | None => (ss,thm,anet,ats,cs); + +fun step(s::ss, thm, anet, ats, cs) = case s of + MK_EQ => (ss, res1(thm,[red2],i), anet, ats, cs) + | ASMS(a) => add_asms(ss,thm,a,anet,ats,cs) + | SIMP_LHS => simp_lhs(thm,ss,anet,ats,cs) + | REW => rew(tapply(net_tac net i,thm),thm,ss,anet,ats,cs,true) + | REFL => (ss, res1(thm,refl_thms,i), anet, ats, cs) + | TRUE => try_true(res1(thm,refl_thms,i),ss,anet,ats,cs) + | PROVE => (if if_fl then MK_EQ::SIMP_LHS::SPLIT::TRUE::ss + else MK_EQ::SIMP_LHS::TRUE::ss, thm, anet, ats, cs) + | POP_ARTR => (ss,thm,hd ats,tl ats,cs) + | POP_CS => (ss,thm,anet,ats,tl cs) + | SPLIT => split(thm,ss,anet,ats,cs); + +fun exec(state as (s::ss, thm, _, _, _)) = + if s=STOP then thm else exec(step(state)); + +in exec(ss, thm, Net.empty, [], []) end; + + +(*ss = list of commands (not simpset!); + fl = even use case splits to solve conditional rewrite rules; + addhyps = add hyps to simpset*) +fun EXEC_TAC (ss,fl,addhyps) simpset = METAHYPS + (fn hyps => + case (if addhyps then simpset addrews hyps else simpset) of + (SS{auto_tac,cong_net,simp_net,splits,split_consts,...}) => + PRIMITIVE(execute(ss,fl,auto_tac hyps, + net_tac cong_net,splits,split_consts, + simp_net, 1)) + THEN TRY(auto_tac hyps 1)); + +val SIMP_TAC = EXEC_TAC([MK_EQ,SIMP_LHS,SPLIT,REFL,STOP],false,false); + +val ASM_SIMP_TAC = + EXEC_TAC([ASMS(0),MK_EQ,SIMP_LHS,SPLIT,REFL,STOP],false,true); + +val SIMP_SPLIT2_TAC = EXEC_TAC([MK_EQ,SIMP_LHS,SPLIT,REFL,STOP],true,false); + +fun REWRITE (ss,fl) (SS{auto_tac,cong_net,simp_net,splits,split_consts,...}) = +let val cong_tac = net_tac cong_net +in fn thm => + let val state = thm RSN (2,red1) + in execute(ss,fl,auto_tac[],cong_tac,splits,split_consts,simp_net,1)state + end +end; + +val SIMP_THM = REWRITE ([ASMS(0),SIMP_LHS,SPLIT,REFL,STOP],false); + + +(* Compute Congruence rules for individual constants using the substition + rules *) + +val subst_thms = map standard subst_thms; + + +fun exp_app(0,t) = t + | exp_app(i,t) = exp_app(i-1,t $ Bound (i-1)); + +fun exp_abs(Type("fun",[T1,T2]),t,i) = + Abs("x"^string_of_int i,T1,exp_abs(T2,t,i+1)) + | exp_abs(T,t,i) = exp_app(i,t); + +fun eta_Var(ixn,T) = exp_abs(T,Var(ixn,T),0); + + +fun Pinst(f,fT,(eq,eqT),k,i,T,yik,Ts) = +let fun xn_list(x,n) = + let val ixs = map (fn i => (x^(radixstring(26,"a",i)),0)) (0 upto n); + in map eta_Var (ixs ~~ (take(n+1,Ts))) end + val lhs = list_comb(f,xn_list("X",k-1)) + val rhs = list_comb(f,xn_list("X",i-1) @ [Bound 0] @ yik) +in Abs("", T, Const(eq,[fT,fT]--->eqT) $ lhs $ rhs) end; + +fun find_subst tsig T = +let fun find (thm::thms) = + let val (Const(_,cT), va, vb) = dest_red(hd(prems_of thm)); + val [P] = term_vars(concl_of thm) \\ [va,vb] + val eqT::_ = binder_types cT + in if Type.typ_instance(tsig,T,eqT) then Some(thm,va,vb,P) + else find thms + end + | find [] = None +in find subst_thms end; + +fun mk_cong sg (f,aTs,rT) (refl,eq) = +let val tsig = #tsig(Sign.rep_sg sg); + val k = length aTs; + fun ri((subst,va as Var(_,Ta),vb as Var(_,Tb),P),i,si,T,yik) = + let val ca = Sign.cterm_of sg va + and cx = Sign.cterm_of sg (eta_Var(("X"^si,0),T)) + val cb = Sign.cterm_of sg vb + and cy = Sign.cterm_of sg (eta_Var(("Y"^si,0),T)) + val cP = Sign.cterm_of sg P + and cp = Sign.cterm_of sg (Pinst(f,rT,eq,k,i,T,yik,aTs)) + in cterm_instantiate [(ca,cx),(cb,cy),(cP,cp)] subst end; + fun mk(c,T::Ts,i,yik) = + let val si = radixstring(26,"a",i) + in case find_subst tsig T of + None => mk(c,Ts,i-1,eta_Var(("X"^si,0),T)::yik) + | Some s => let val c' = c RSN (2,ri(s,i,si,T,yik)) + in mk(c',Ts,i-1,eta_Var(("Y"^si,0),T)::yik) end + end + | mk(c,[],_,_) = c; +in mk(refl,rev aTs,k-1,[]) end; + +fun mk_cong_type sg (f,T) = +let val (aTs,rT) = strip_type T; + val tsig = #tsig(Sign.rep_sg sg); + fun find_refl(r::rs) = + let val (Const(eq,eqT),_,_) = dest_red(concl_of r) + in if Type.typ_instance(tsig, rT, hd(binder_types eqT)) + then Some(r,(eq,body_type eqT)) else find_refl rs + end + | find_refl([]) = None; +in case find_refl refl_thms of + None => [] | Some(refl) => [mk_cong sg (f,aTs,rT) refl] +end; + +fun mk_cong_thy thy f = +let val sg = sign_of thy; + val T = case Sign.Symtab.lookup(#const_tab(Sign.rep_sg sg),f) of + None => error(f^" not declared") | Some(T) => T; + val T' = incr_tvar 9 T; +in mk_cong_type sg (Const(f,T'),T') end; + +fun mk_congs thy = filter_out is_fact o flat o map (mk_cong_thy thy); + +fun mk_typed_congs thy = +let val sg = sign_of thy; + val S0 = Type.defaultS(#tsig(Sign.rep_sg sg)) + fun readfT(f,s) = + let val T = incr_tvar 9 (Sign.read_typ(sg,K(Some(S0))) s); + val t = case Sign.Symtab.lookup(#const_tab(Sign.rep_sg sg),f) of + Some(_) => Const(f,T) | None => Free(f,T) + in (t,T) end +in flat o map (mk_cong_type sg o readfT) end; + +(* This code is fishy, esp the "let val T' = ..." +fun extract_free_congs() = +let val {prop,sign,...} = rep_thm(topthm()); + val frees = add_term_frees(prop,[]); + fun filter(Free(a,T as Type("fun",_))) = + let val T' = incr_tvar 9 (Type.varifyT T) + in [(Free(a,T),T)] end + | filter _ = [] +in flat(map (mk_cong_type sign) (flat (map filter frees))) end; +*) + +end (* local *) +end (* SIMP *); diff -r 000000000000 -r a5a9c433f639 src/Provers/simplifier.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Provers/simplifier.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,151 @@ +infix addsimps addcongs + setsolver setloop setmksimps setsubgoaler; + +signature SIMPLIFIER = +sig + type simpset + val addcongs: simpset * thm list -> simpset + val addsimps: simpset * thm list -> simpset + val asm_full_simp_tac: simpset -> int -> tactic + val asm_simp_tac: simpset -> int -> tactic + val empty_ss: simpset + val merge_ss: simpset * simpset -> simpset + val prems_of_ss: simpset -> thm list + val rep_ss: simpset -> {simps: thm list, congs: thm list} + val setsolver: simpset * (thm list -> int -> tactic) -> simpset + val setloop: simpset * (int -> tactic) -> simpset + val setmksimps: simpset * (thm -> thm list) -> simpset + val setsubgoaler: simpset * (simpset -> int -> tactic) -> simpset + val simp_tac: simpset -> int -> tactic +end; + +structure Simplifier:SIMPLIFIER = +struct + +datatype simpset = + SS of {mss: meta_simpset, + simps: thm list, + congs: thm list, + subgoal_tac: simpset -> int -> tactic, + finish_tac: thm list -> int -> tactic, + loop_tac: int -> tactic}; + +val empty_ss = + SS{mss=empty_mss, + simps= [], + congs= [], + subgoal_tac= K(K no_tac), + finish_tac= K(K no_tac), + loop_tac= K no_tac}; + + +fun (SS{mss,simps,congs,subgoal_tac,finish_tac,...}) setloop loop_tac = + SS{mss=mss, + simps= simps, + congs= congs, + subgoal_tac= subgoal_tac, + finish_tac=finish_tac, + loop_tac=loop_tac}; + +fun (SS{mss,simps,congs,subgoal_tac,loop_tac,...}) setsolver finish_tac = + SS{mss=mss, + simps= simps, + congs= congs, + subgoal_tac= subgoal_tac, + finish_tac=finish_tac, + loop_tac=loop_tac}; + +fun (SS{mss,simps,congs,finish_tac,loop_tac,...}) setsubgoaler subgoal_tac = + SS{mss=mss, + simps= simps, + congs= congs, + subgoal_tac= subgoal_tac, + finish_tac=finish_tac, + loop_tac=loop_tac}; + +fun (SS{mss,simps,congs,subgoal_tac,finish_tac,loop_tac}) setmksimps mk_simps = + SS{mss=Thm.set_mk_rews(mss,mk_simps), + simps= simps, + congs= congs, + subgoal_tac= subgoal_tac, + finish_tac=finish_tac, + loop_tac=loop_tac}; + +fun (SS{mss,simps,congs,subgoal_tac,finish_tac,loop_tac}) addsimps rews = + let val rews' = flat(map (Thm.mk_rews_of_mss mss) rews) + in + SS{mss= Thm.add_simps(mss,rews'), + simps= rews' @ simps, + congs= congs, + subgoal_tac= subgoal_tac, + finish_tac=finish_tac, + loop_tac=loop_tac} + end; + +fun (SS{mss,simps,congs,subgoal_tac,finish_tac,loop_tac}) addcongs newcongs = + SS{mss= Thm.add_congs(mss,newcongs), + simps= simps, + congs= newcongs @ congs, + subgoal_tac= subgoal_tac, + finish_tac=finish_tac, + loop_tac=loop_tac}; + +fun merge_ss(SS{mss,simps,congs,subgoal_tac,finish_tac,loop_tac}, + SS{simps=simps2,congs=congs2,...}) = + let val simps' = gen_union eq_thm (simps,simps2) + val congs' = gen_union eq_thm (congs,congs2) + val mss' = Thm.set_mk_rews(empty_mss,Thm.mk_rews_of_mss mss) + val mss' = Thm.add_simps(mss',simps') + val mss' = Thm.add_congs(mss',congs') + in SS{mss= mss', + simps= simps, + congs= congs', + subgoal_tac= subgoal_tac, + finish_tac= finish_tac, + loop_tac= loop_tac} + end; + +fun prems_of_ss(SS{mss,...}) = prems_of_mss mss; + +fun rep_ss(SS{simps,congs,...}) = {simps=simps,congs=congs}; + +fun add_asms (SS{mss,simps,congs,subgoal_tac,finish_tac,loop_tac}) prems = + let val rews = flat(map (mk_rews_of_mss mss) prems) + in SS{mss= add_prems(add_simps(mss,rews),prems), simps=simps, congs=congs, + subgoal_tac=subgoal_tac,finish_tac=finish_tac, + loop_tac=loop_tac} + end; + +fun asm_full_simp_tac(SS{mss,simps,congs,subgoal_tac,finish_tac,loop_tac}) = + let fun solve_all_tac mss = + let val ss = SS{mss=mss,simps=simps,congs=congs, + subgoal_tac=subgoal_tac, + finish_tac=finish_tac, loop_tac=loop_tac} + fun solve1 thm = tapply( + STATE(fn state => let val n = nprems_of state + in if n=0 then all_tac + else subgoal_tac ss 1 THEN + COND (has_fewer_prems n) (Tactic solve1) no_tac + end), + thm) + in DEPTH_SOLVE(Tactic solve1) end + + fun simp_loop i thm = + tapply(asm_rewrite_goal_tac solve_all_tac mss i THEN + (finish_tac (prems_of_mss mss) i ORELSE STATE(looper i)), + thm) + and allsimp i m state = + let val n = nprems_of state + in EVERY(map (fn j => simp_loop_tac (i+j)) ((n-m) downto 0)) end + and looper i state = + TRY(loop_tac i THEN STATE(allsimp i (nprems_of state))) + and simp_loop_tac i = Tactic(simp_loop i) + + in simp_loop_tac end; + +fun asm_simp_tac ss = + METAHYPS(fn prems => asm_full_simp_tac (add_asms ss prems) 1); + +fun simp_tac ss = METAHYPS(fn _ => asm_full_simp_tac ss 1); + +end; diff -r 000000000000 -r a5a9c433f639 src/Provers/splitter.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Provers/splitter.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,112 @@ +(*** case splitting ***) +(* +Use: + +val split_tac = mk_case_split_tac iffD; + +by(case_split_tac splits i); + +where splits = [P(elim(...)) == rhs, ...] + iffD = [| P <-> Q; Q |] ==> P (* is called iffD2 in HOL *) + +*) + +fun mk_case_split_tac iffD = +let + +val lift = prove_goal Pure.thy + "[| !!x::'b::logic. Q(x) == R(x) |] ==> P(%x.Q(x)) == P(%x.R(x))::'a::logic" + (fn [prem] => [rewtac prem, rtac reflexive_thm 1]); +(* +val iffD = prove_goal Pure.thy "[| PROP P == PROP Q; PROP Q |] ==> PROP P" + (fn [p1,p2] => [rtac (equal_elim (symmetric p1) p2) 1]); +*) +val trlift = lift RS transitive_thm; +val _ $ (Var(P,PT)$_) $ _ = concl_of trlift; + +fun contains cmap = + let fun isin i (Abs(_,_,b)) = isin 0 b + | isin i (s$t) = isin (i+1) s orelse isin 0 t + | isin i (Const(c,_)) = (case assoc(cmap,c) of + Some(n,_) => n <= i + | None => false) + | isin _ _ = false + in isin 0 end; + +fun mk_context cmap Ts maxi t = + let fun var (t,i) = Var(("X",i),type_of1(Ts,t)) + + fun mka((None,i,ts),t) = if contains cmap t + then let val (u,T,j) = mk(t,i) in (Some(T),j,ts@[u]) end + else (None,i+1,ts@[var(t,i)]) + | mka((some,i,ts),t) = (some,i+1,ts@[var(t,i)]) + and mk(t as Abs _, i) = (Bound 0,type_of1(Ts,t),i) + | mk(t,i) = + let val (f,ts) = strip_comb t + val (Some(T),j,us) = foldl mka ((None,i,[]),ts) + in (list_comb(f,us),T,j) end + + val (u,T,_) = mk(t,maxi+1) + in Abs("",T,u) end; + +fun nth_subgoal i thm = nth_elem(i-1,prems_of thm); + +fun goal_concl i thm = Logic.strip_assums_concl(nth_subgoal i thm); + +fun inst_lift cmap state lift i = + let val sg = #sign(rep_thm state) + val tsig = #tsig(Sign.rep_sg sg) + val goali = nth_subgoal i state + val Ts = map #2 (Logic.strip_params goali) + val _ $ t $ _ = Logic.strip_assums_concl goali; + val cntxt = mk_context cmap (rev Ts) (#maxidx(rep_thm lift)) t + val cu = Sign.cterm_of sg cntxt + val uT = #T(Sign.rep_cterm cu) + val cP' = Sign.cterm_of sg (Var(P,uT)) + val ixnTs = Type.typ_match tsig ([],(PT,uT)); + val ixncTs = map (fn (x,y) => (x,Sign.ctyp_of sg y)) ixnTs; + in instantiate (ixncTs, [(cP',cu)]) lift end; + +fun splittable al i thm = + let val tm = goal_concl i thm + fun nobound j k (Abs(_,_,t)) = nobound j (k+1) t + | nobound j k (s$t) = nobound j k s andalso nobound j k t + | nobound j k (Bound n) = n < k orelse k+j <= n + | nobound _ _ _ = true; + fun find j (None,t) = (case t of + Abs(_,_,t) => find (j+1) (None,t) + | _ => let val (h,args) = strip_comb t + fun no() = foldl (find j) (None,args) + in case h of + Const(c,_) => + (case assoc(al,c) of + Some(n,thm) => + if length args < n then no() + else if forall (nobound j 0) (take(n,args)) + then Some(thm) + else no() + | None => no()) + | _ => no() + end) + | find _ (some,_) = some + in find 0 (None,tm) end; + +fun split_tac [] i = no_tac + | split_tac splits i = + let fun const(thm) = let val _$(_$lhs)$_ = concl_of thm + val (Const(a,_),args) = strip_comb lhs + in (a,(length args,thm)) end + val cmap = map const splits; + fun lift thm = rtac (inst_lift cmap thm trlift i) i + fun lift_split thm = + case splittable cmap i thm of + Some(split) => rtac split i + | None => EVERY[STATE lift, rtac reflexive_thm (i+1), + STATE lift_split] + in STATE(fn thm => + if (i <= nprems_of thm) andalso contains cmap (goal_concl i thm) + then EVERY[rtac iffD i, STATE lift_split] + else no_tac) + end; + +in split_tac end; diff -r 000000000000 -r a5a9c433f639 src/Provers/typedsimp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Provers/typedsimp.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,127 @@ +(* Title: typedsimp + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Functor for constructing simplifiers. Suitable for Constructive Type +Theory with its typed reflexivity axiom a:A ==> a=a:A. For most logics try +simp.ML. +*) + +signature TSIMP_DATA = + sig + val refl: thm (*Reflexive law*) + val sym: thm (*Symmetric law*) + val trans: thm (*Transitive law*) + val refl_red: thm (* reduce(a,a) *) + val trans_red: thm (* [|a=b; reduce(b,c) |] ==> a=c *) + val red_if_equal: thm (* a=b ==> reduce(a,b) *) + (*Built-in rewrite rules*) + val default_rls: thm list + (*Type checking or similar -- solution of routine conditions*) + val routine_tac: thm list -> int -> tactic + end; + +signature TSIMP = + sig + val asm_res_tac: thm list -> int -> tactic + val cond_norm_tac: ((int->tactic) * thm list * thm list) -> tactic + val cond_step_tac: ((int->tactic) * thm list * thm list) -> int -> tactic + val norm_tac: (thm list * thm list) -> tactic + val process_rules: thm list -> thm list * thm list + val rewrite_res_tac: int -> tactic + val split_eqn: thm + val step_tac: (thm list * thm list) -> int -> tactic + val subconv_res_tac: thm list -> int -> tactic + end; + + +functor TSimpFun (TSimp_data: TSIMP_DATA) : TSIMP = +struct +local open TSimp_data +in + +(*For simplifying both sides of an equation: + [| a=c; b=c |] ==> b=a + Can use resolve_tac [split_eqn] to prepare an equation for simplification. *) +val split_eqn = standard (sym RSN (2,trans) RS sym); + + +(* [| a=b; b=c |] ==> reduce(a,c) *) +val red_trans = standard (trans RS red_if_equal); + +(*For REWRITE rule: Make a reduction rule for simplification, e.g. + [| a: C(0); ... ; a=c: C(0) |] ==> rec(0,a,b) = c: C(0) *) +fun simp_rule rl = rl RS trans; + +(*For REWRITE rule: Make rule for resimplifying if possible, e.g. + [| a: C(0); ...; a=c: C(0) |] ==> reduce(rec(0,a,b), c) *) +fun resimp_rule rl = rl RS red_trans; + +(*For CONGRUENCE rule, like a=b ==> succ(a) = succ(b) + Make rule for simplifying subterms, e.g. + [| a=b: N; reduce(succ(b), c) |] ==> succ(a)=c: N *) +fun subconv_rule rl = rl RS trans_red; + +(*If the rule proves an equality then add both forms to simp_rls + else add the rule to other_rls*) +fun add_rule (rl, (simp_rls, other_rls)) = + (simp_rule rl :: resimp_rule rl :: simp_rls, other_rls) + handle THM _ => (simp_rls, rl :: other_rls); + +(*Given the list rls, return the pair (simp_rls, other_rls).*) +fun process_rules rls = foldr add_rule (rls, ([],[])); + +(*Given list of rewrite rules, return list of both forms, reject others*) +fun process_rewrites rls = + case process_rules rls of + (simp_rls,[]) => simp_rls + | (_,others) => raise THM + ("process_rewrites: Ill-formed rewrite", 0, others); + +(*Process the default rewrite rules*) +val simp_rls = process_rewrites default_rls; + +(*If subgoal is too flexible (e.g. ?a=?b or just ?P) then filt_resolve_tac + will fail! The filter will pass all the rules, and the bound permits + no ambiguity.*) + +(*Resolution with rewrite/sub rules. Builds the tree for filt_resolve_tac.*) +val rewrite_res_tac = filt_resolve_tac simp_rls 2; + +(*The congruence rules for simplifying subterms. If subgoal is too flexible + then only refl,refl_red will be used (if even them!). *) +fun subconv_res_tac congr_rls = + filt_resolve_tac (map subconv_rule congr_rls) 2 + ORELSE' filt_resolve_tac [refl,refl_red] 1; + +(*Resolve with asms, whether rewrites or not*) +fun asm_res_tac asms = + let val (xsimp_rls,xother_rls) = process_rules asms + in routine_tac xother_rls ORELSE' + filt_resolve_tac xsimp_rls 2 + end; + +(*Single step for simple rewriting*) +fun step_tac (congr_rls,asms) = + asm_res_tac asms ORELSE' rewrite_res_tac ORELSE' + subconv_res_tac congr_rls; + +(*Single step for conditional rewriting: prove_cond_tac handles new subgoals.*) +fun cond_step_tac (prove_cond_tac, congr_rls, asms) = + asm_res_tac asms ORELSE' rewrite_res_tac ORELSE' + (resolve_tac [trans, red_trans] THEN' prove_cond_tac) ORELSE' + subconv_res_tac congr_rls; + +(*Unconditional normalization tactic*) +fun norm_tac arg = REPEAT_FIRST (step_tac arg) THEN + TRYALL (resolve_tac [red_if_equal]); + +(*Conditional normalization tactic*) +fun cond_norm_tac arg = REPEAT_FIRST (cond_step_tac arg) THEN + TRYALL (resolve_tac [red_if_equal]); + +end; +end; + + diff -r 000000000000 -r a5a9c433f639 src/Pure/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Makefile Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,50 @@ +######################################################################### +# # +# Makefile for Isabelle (Pure) # +# # +######################################################################### + +#The pure part is common to all systems. +#Object-logics (like FOL) are loaded on top of it. + +#To make the system, cd to this directory and type +# make -f Makefile + +#Environment variable ML_DBASE specifies the initial Poly/ML database, from +# the Poly/ML distribution directory. +#WARNING: Poly/ML parent databases should not be moved! + +#Environment variable ISABELLECOMP specifies the compiler. +#Environment variable ISABELLEBIN specifies the destination directory. +#For Poly/ML, ISABELLEBIN must begin with a / + +BIN = $(ISABELLEBIN) +COMP = $(ISABELLECOMP) +FILES = POLY.ML NJ.ML ROOT.ML library.ML term.ML symtab.ML type.ML sign.ML\ + sequence.ML envir.ML pattern.ML unify.ML logic.ML thm.ML net.ML\ + drule.ML tctical.ML tactic.ML goals.ML + +SYNTAX_FILES = Syntax/ROOT.ML Syntax/ast.ML Syntax/xgram.ML\ + Syntax/extension.ML Syntax/lexicon.ML Syntax/parse_tree.ML\ + Syntax/earley0A.ML Syntax/type_ext.ML Syntax/sextension.ML\ + Syntax/pretty.ML Syntax/printer.ML Syntax/syntax.ML + +THY_FILES = Thy/ROOT.ML Thy/scan.ML Thy/parse.ML Thy/syntax.ML Thy/read.ML + +#Uses cp rather than make_database because Poly/ML allows only 3 levels +$(BIN)/Pure: $(FILES) $(SYNTAX_FILES) $(THY_FILES) $(ML_DBASE) + case "$(COMP)" in \ + poly*) echo database=$${ML_DBASE:?'No Poly/ML database specified'};\ + cp $(ML_DBASE) $(BIN)/Pure; chmod u+w $(BIN)/Pure;\ + echo 'PolyML.use"POLY";use"ROOT";' | $(COMP) $(BIN)/Pure;;\ + sml*) if [ ! '(' -d $${ISABELLEBIN:?} -a -w $${ISABELLEBIN:?} ')' ];\ + then echo Bad value for ISABELLEBIN : \ + $(BIN) is not a writable directory; \ + exit 1; \ + fi;\ + echo 'use"NJ.ML"; use"ROOT.ML"; xML"$(BIN)/Pure" banner;'\ + | $(COMP);;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +.PRECIOUS: $(BIN)/Pure diff -r 000000000000 -r a5a9c433f639 src/Pure/NJ.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/NJ.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,91 @@ +(* Title: Pure/NJ + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Compatibility file for Standard ML of New Jersey. +*) + +(*** Poly/ML emulation ***) + +(*To exit the system -- an alternative to ^D *) +fun quit () = System.Unsafe.CInterface.exit 0; + +(*To change the current directory*) +val cd = System.Directory.cd; + +(*To limit the printing depth [divided by 2 for comparibility with Poly/ML]*) +fun print_depth n = (System.Control.Print.printDepth := n div 2; + System.Control.Print.printLength := n); + + +(*Interface for toplevel pretty printers, see also Pure/install_pp.ML*) + +fun make_pp path pprint = + let + open System.PrettyPrint; + + fun pp pps obj = + pprint obj + (add_string pps, begin_block pps INCONSISTENT, + fn wd => add_break pps (wd, 0), fn () => end_block pps); + in + (path, pp) + end; + +fun install_pp (path, pp) = System.PrettyPrint.install_pp path pp; + + +(*** New Jersey ML parameters ***) + +(* Suppresses Garbage Collection messages; default was 2 *) +System.Control.Runtime.gcmessages := 0; + +(*redefine to flush its output immediately -- temporary patch suggested + by Kim Dam Petersen*) +val output = fn(s, t) => (output(s, t); flush_out s); + +(*Dummy version of the Poly/ML function*) +fun commit() = (); + +(*For use in Makefiles -- saves space*) +fun xML filename banner = (exportML filename; print(banner^"\n")); + +(*** Timing functions ***) + +(*Print microseconds, suppressing trailing zeroes*) +fun umakestring 0 = "" + | umakestring n = chr(ord"0" + n div 100000) ^ + umakestring(10 * (n mod 100000)); + +(*A conditional timing function: applies f to () and, if the flag is true, + prints its runtime. *) +fun cond_timeit flag f = + if flag then + let fun string_of_time (System.Timer.TIME {sec, usec}) = + makestring sec ^ "." ^ (if usec=0 then "0" else umakestring usec) + open System.Timer; + val start = start_timer() + val result = f(); + val nongc = check_timer(start) + and gc = check_timer_gc(start); + val both = add_time(nongc, gc) + in print("Non GC " ^ string_of_time nongc ^ + " GC " ^ string_of_time gc ^ + " both "^ string_of_time both ^ " secs\n"); + result + end + else f(); + + +(*** File information ***) + +(*Get time of last modification.*) +local + open System.Timer; + open System.Unsafe.SysIO; +in + fun file_info "" = "" + | file_info name = makestring (mtime (PATH name)); +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/POLY.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/POLY.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,48 @@ +(* Title: Pure/POLY + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Compatibility file for Poly/ML (AHL release 1.88) +*) + +open PolyML ExtendedIO; + + +(*A conditional timing function: applies f to () and, if the flag is true, + prints its runtime.*) + +fun cond_timeit flag f = + if flag then + let val before = System.processtime() + val result = f() + val both = real(System.processtime() - before) / 10.0 + in output(std_out, "Process time (incl GC): "^ + makestring both ^ " secs\n"); + result + end + else f(); + + +(*Save function: in Poly/ML, ignores filename and commits to present file*) + +fun xML filename banner = commit(); + + +(*Interface for toplevel pretty printers, see also Pure/install_pp.ML*) + +fun make_pp _ pprint (str, blk, brk, en) obj = + pprint obj (str, fn ind => blk (ind, false), fn wd => brk (wd, 0), en); + + +(*** File information ***) + +(*Get time of last modification.*) +fun file_info "" = "" + | file_info name = + let val (is, os) = ExtendedIO.execute ("ls -l " ^ name) + val result = ExtendedIO.input_line is; + val dummy = close_in is; + val dummy = close_out os; + in result end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/README Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,15 @@ + Pure: The Pure Isabelle System + +This directory contains the ML source files for Pure Isabelle, which is the +basis for all object-logics. Important files include + +Makefile -- compiles the files under Poly/ML or SML of New Jersey + +Syntax -- subdirectory containing the parser/prettyprinter generator + +ROOT.ML -- loads all source files. Enter ML and type: use "ROOT.ML"; + +NJ.ML -- compatibility file for Standard ML of New Jersey. You may wish to + alter the parameter settings. + +POLY.ML -- compatibility file for Poly/ML diff -r 000000000000 -r a5a9c433f639 src/Pure/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,71 @@ +(* Title: ROOT + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Root file for pure Isabelle: all modules in proper order for loading +Loads pure Isabelle into an empty database. + +To build system, use Makefile (Poly/ML) or Makefile.NJ (New Jersey ML) + +TO DO: +instantiation of theorems can lead to inconsistent sorting of type vars if +'a::S is already present and 'a::T is introduced. +*) + +val banner = "Pure Isabelle"; +val version = "February 93"; + +print_depth 1; + +use "library.ML"; +use "term.ML"; +use "symtab.ML"; + +(*Used for building the parser generator*) +structure Symtab = SymtabFun(); +cd "Syntax"; +use "ROOT.ML"; +cd ".."; + +(* Theory parser *) +cd "Thy"; +use "ROOT.ML"; +cd ".."; + +print_depth 1; +use "type.ML"; +use "sign.ML"; +use "sequence.ML"; +use "envir.ML"; +use "pattern.ML"; +use "unify.ML"; +use "net.ML"; +use "logic.ML"; +use "thm.ML"; +use "drule.ML"; +use "tctical.ML"; +use "tactic.ML"; +use "goals.ML"; + +(*Will be visible to all object-logics.*) +structure Type = TypeFun(structure Symtab=Symtab and Syntax=Syntax); +structure Sign = SignFun(structure Type=Type and Syntax=Syntax); +structure Sequence = SequenceFun(); +structure Envir = EnvirFun(); +structure Pattern = PatternFun(structure Sign=Sign and Envir=Envir); +structure Unify = UnifyFun(structure Sign=Sign and Envir=Envir + and Sequence=Sequence and Pattern=Pattern); +structure Net = NetFun(); +structure Logic = LogicFun(structure Unify=Unify and Net=Net); +structure Thm = ThmFun(structure Logic=Logic and Unify=Unify and Net=Net + and Pattern=Pattern); +structure Drule = DruleFun(structure Logic=Logic and Thm=Thm); +structure Tactical = TacticalFun(structure Logic=Logic and Drule=Drule); +structure Tactic = TacticFun(structure Logic=Logic and Drule=Drule + and Tactical=Tactical and Net=Net); +structure Goals = GoalsFun(structure Logic=Logic and Drule=Drule + and Tactic=Tactic and Pattern=Pattern); +open Basic_Syntax Thm Drule Tactical Tactic Goals; + +structure Pure = struct val thy = pure_thy end; diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/README Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,8 @@ + Pure/Syntax + +This directory contains the source files for Isabelle's syntax module. Which +includes a lexer, parser, pretty printer and macro system. + +There is no Makefile to compile these files separately; they are compiled as +part of Pure Isabelle. + diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,40 @@ +(* Title: Pure/Syntax/ROOT + ID: $Id$ + Author: Tobias Nipkow and Markus Wenzel, TU Muenchen + +This file builds the syntax module. It assumes the standard Isabelle +environment. +*) + +use "ast.ML"; +use "xgram.ML"; +use "extension.ML"; +use "lexicon.ML"; +use "parse_tree.ML"; +use "earley0A.ML"; +use "type_ext.ML"; +use "sextension.ML"; +use "pretty.ML"; +use "printer.ML"; +use "syntax.ML"; + +structure Ast = AstFun(); +structure XGram = XGramFun(Ast); +structure Extension = ExtensionFun(XGram); +structure Lexicon = LexiconFun(Extension); +structure ParseTree = ParseTreeFun(structure Lexicon = Lexicon and Ast = Ast); +structure Earley = EarleyFun(structure XGram = XGram and ParseTree = ParseTree); +structure TypeExt = TypeExtFun(structure Extension = Extension + and Lexicon = Lexicon); +structure SExtension = SExtensionFun(structure TypeExt = TypeExt + and Lexicon = Lexicon); +structure Pretty = PrettyFun(); +structure Printer = PrinterFun(structure Symtab = Symtab and Lexicon = Lexicon + and TypeExt = TypeExt and SExtension = SExtension and Pretty = Pretty); +structure Syntax = SyntaxFun(structure Symtab = Symtab and TypeExt = TypeExt + and Parser = Earley and SExtension = SExtension and Printer = Printer); + +(*Basic_Syntax has the most important primitives, which are made pervasive*) +signature BASIC_SYNTAX = sig include SEXTENSION0 include PRINTER0 end; +structure Basic_Syntax: BASIC_SYNTAX = Syntax; + diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/ast.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/ast.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,249 @@ +(* Title: Pure/Syntax/ast + ID: $Id$ + Author: Markus Wenzel, TU Muenchen + +Abstract Syntax Trees, Syntax Rules and translation, matching, normalization +of asts. +*) + +signature AST = +sig + datatype ast = + Constant of string | + Variable of string | + Appl of ast list + val mk_appl: ast -> ast list -> ast + exception AST of string * ast list + val raise_ast: string -> ast list -> 'a + val str_of_ast: ast -> string + val head_of_rule: ast * ast -> string + val rule_error: ast * ast -> string option + val fold_ast: string -> ast list -> ast + val fold_ast_p: string -> ast list * ast -> ast + val unfold_ast: string -> ast -> ast list + val unfold_ast_p: string -> ast -> ast list * ast + val trace_norm: bool ref + val stat_norm: bool ref + val normalize: (string -> (ast * ast) list) option -> ast -> ast +end; + +functor AstFun()(*: AST *) = (* FIXME *) +struct + + +(** Abstract Syntax Trees **) + +(*asts come in two flavours: + - proper asts representing terms and types: Variables are treated like + Constants; + - patterns used as lhs and rhs in rules: Variables are placeholders for + proper asts*) + +datatype ast = + Constant of string | (* "not", "_%", "fun" *) + Variable of string | (* x, ?x, 'a, ?'a *) + Appl of ast list; (* (f x y z), ("fun" 'a 'b) *) + + +(*the list of subasts of an Appl node has to contain at least 2 elements, i.e. + there are no empty asts or nullary applications; use mk_appl for convenience*) + +fun mk_appl ast [] = ast + | mk_appl ast asts = Appl (ast :: asts); + + +(*exception for system errors involving asts*) + +exception AST of string * ast list; + +fun raise_ast msg asts = raise (AST (msg, asts)); + + +(* print asts in a LISP-like style *) + +fun str_of_ast (Constant a) = quote a + | str_of_ast (Variable x) = x + | str_of_ast (Appl asts) = "(" ^ (space_implode " " (map str_of_ast asts)) ^ ")"; + + +(* head_of_ast, head_of_rule *) + +fun head_of_ast (Constant a) = Some a + | head_of_ast (Appl (Constant a :: _)) = Some a + | head_of_ast _ = None; + +fun head_of_rule (lhs, _) = the (head_of_ast lhs); + + + +(** check Syntax Rules **) + +(*a wellformed rule (lhs, rhs): (ast * ast) has the following properties: + - the head of lhs is a constant, + - the lhs has unique vars, + - vars of rhs is subset of vars of lhs*) + +fun rule_error (rule as (lhs, rhs)) = + let + fun vars_of (Constant _) = [] + | vars_of (Variable x) = [x] + | vars_of (Appl asts) = flat (map vars_of asts); + + fun unique (x :: xs) = not (x mem xs) andalso unique xs + | unique [] = true; + + val lvars = vars_of lhs; + val rvars = vars_of rhs; + in + if is_none (head_of_ast lhs) then Some "lhs has no constant head" + else if not (unique lvars) then Some "duplicate vars in lhs" + else if not (rvars subset lvars) then Some "rhs contains extra variables" + else None + end; + + + +(** translation utilities **) + +(* fold asts *) + +fun fold_ast _ [] = raise Match + | fold_ast _ [y] = y + | fold_ast c (x :: xs) = Appl [Constant c, x, fold_ast c xs]; + +fun fold_ast_p c = foldr (fn (x, xs) => Appl [Constant c, x, xs]); + + +(* unfold asts *) + +fun unfold_ast c (y as Appl [Constant c', x, xs]) = + if c = c' then x :: (unfold_ast c xs) else [y] + | unfold_ast _ y = [y]; + +fun cons_fst x (xs, y) = (x :: xs, y); + +fun unfold_ast_p c (y as Appl [Constant c', x, xs]) = + if c = c' then cons_fst x (unfold_ast_p c xs) + else ([], y) + | unfold_ast_p _ y = ([], y); + + + +(** normalization of asts **) + +(* simple env *) + +structure Env = +struct + val empty = []; + val add = op ::; + val get = the o assoc; +end; + + +(* match *) + +fun match ast pat = + let + exception NO_MATCH; + + fun mtch (Constant a, Constant b, env) = + if a = b then env else raise NO_MATCH + | mtch (Variable a, Constant b, env) = + if a = b then env else raise NO_MATCH + | mtch (ast, Variable x, env) = Env.add ((x, ast), env) + | mtch (Appl asts, Appl pats, env) = mtch_lst (asts, pats, env) + | mtch _ = raise NO_MATCH + and mtch_lst (ast :: asts, pat :: pats, env) = + mtch_lst (asts, pats, mtch (ast, pat, env)) + | mtch_lst ([], [], env) = env + | mtch_lst _ = raise NO_MATCH; + in + Some (mtch (ast, pat, Env.empty)) handle NO_MATCH => None + end; + + +(* normalize *) (* FIXME clean *) + +val trace_norm = ref false; +val stat_norm = ref false; + +(*the normalizer works yoyo-like: top-down, bottom-up, top-down, ...*) + +fun normalize get_rules pre_ast = + let + val passes = ref 0; + val lookups = ref 0; + val failed_matches = ref 0; + val changes = ref 0; + + val trace = ! trace_norm; + + fun inc i = i := ! i + 1; + + fun subst _ (ast as (Constant _)) = ast + | subst env (Variable x) = Env.get (env, x) + | subst env (Appl asts) = Appl (map (subst env) asts); + + fun try_rules ast ((lhs, rhs) :: pats) = + (case match ast lhs of + Some env => (inc changes; Some (subst env rhs)) + | None => (inc failed_matches; try_rules ast pats)) + | try_rules ast [] = None; + + fun try ast a = (inc lookups; try_rules ast (the get_rules a)); + + fun rewrite (ast as Constant a) = try ast a + | rewrite (ast as Variable a) = try ast a + | rewrite (ast as Appl (Constant a :: _)) = try ast a + | rewrite (ast as Appl (Variable a :: _)) = try ast a + | rewrite _ = None; + + fun rewrote old_ast new_ast = + if trace then + writeln ("rewrote: " ^ str_of_ast old_ast ^ " -> " ^ str_of_ast new_ast) + else (); + + fun norm_root ast = + (case rewrite ast of + Some new_ast => (rewrote ast new_ast; norm_root new_ast) + | None => ast); + + fun norm ast = + (case norm_root ast of + Appl sub_asts => + let + val old_changes = ! changes; + val new_ast = Appl (map norm sub_asts); + in + if old_changes = ! changes then new_ast else norm_root new_ast + end + | atomic_ast => atomic_ast); + + fun normal ast = + let + val old_changes = ! changes; + val new_ast = norm ast; + in + inc passes; + if old_changes = ! changes then new_ast else normal new_ast + end; + + + val () = if trace then writeln ("pre: " ^ str_of_ast pre_ast) else (); + + val post_ast = if is_some get_rules then normal pre_ast else pre_ast; + in + if trace orelse ! stat_norm then + writeln ("post: " ^ str_of_ast post_ast ^ "\nnormalize: " ^ + string_of_int (! passes) ^ " passes, " ^ + string_of_int (! lookups) ^ " lookups, " ^ + string_of_int (! changes) ^ " changes, " ^ + string_of_int (! failed_matches) ^ " matches failed") + else (); + post_ast + end; + + +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/earley0A.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/earley0A.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,330 @@ +(* Title: Pure/Syntax/earley0A + ID: $Id$ + Author: Tobias Nipkow + +Changes: + PARSER: renamed Prod to prod + renamed mk_early_gram to mk_earley_gram +*) + +signature PARSER = +sig + structure XGram: XGRAM + structure ParseTree: PARSE_TREE + local open XGram ParseTree ParseTree.Lexicon in + type Gram + val compile_xgram: string list * Token prod list -> Gram + val parse: Gram * string * Token list -> ParseTree + val parsable: Gram * string -> bool + exception SYNTAX_ERR of Token list + val print_gram: Gram * Lexicon -> unit + end +end; + +functor EarleyFun(structure XGram:XGRAM and ParseTree:PARSE_TREE): PARSER = +struct + +structure XGram = XGram; +structure ParseTree = ParseTree; + +(* Linked lists: *) +infix 5 &; +datatype 'a LList = nilL | op & of 'a * ('a LListR) +withtype 'a LListR = 'a LList ref; + +(* Apply proc to all elements in a linked list *) +fun seqll (proc: '_a -> unit) : ('_a LListR -> unit) = + let fun seq (ref nilL) = () | + seq (ref((a:'_a)&l)) = (proc a; seq l) + in seq end; + +fun llist_to_list (ref nilL) = [] | + llist_to_list (ref(a&ll)) = a::(llist_to_list ll); + +val len = length; + +local open Array XGram ParseTree ParseTree.Lexicon in +nonfix sub; + +fun forA(p: int -> unit, a: 'a array) : unit = + let val ub = length a + fun step(i) = if i=ub then () else (p(i); step(i+1)); + in step 0 end; + +fun itA(a0:'a, b: 'b array)(f:'a * 'b -> 'a) : 'a = + let val ub = length b + fun acc(a,i) = if i=ub then a else acc(f(a,sub(b,i)),i+1) + in acc(a0,0) end; + +(* +Gram: name of nt -> number, nt number -> productions array, + nt number -> list of nt's reachable via copy ops +*) + +datatype Symbol = T of Token | NT of int * int + and Op = Op of OpSyn * string * int +withtype OpSyn = Symbol array + and OpListA = Op array * int TokenMap + and FastAcc = int TokenMap; + +type Gram = int Symtab.table * OpListA array * int list array; + +fun non_term(Nonterminal(s,_)) = if predef_term(s)=end_token then [s] else [] + | non_term(_) = []; + +fun non_terms(Prod(_,symbs,_,_)) = flat(map non_term symbs); + +(* mk_pre_grammar converts a list of productions in external format into an +internal Gram object. *) +val dummyTM:FastAcc = mkTokenMap[]; + +fun mk_pre_grammar(prods): Gram = +let fun same_res(Prod(t1,_,_,_), Prod(t2,_,_,_)) = t1=t2; + val partitioned0 = partition_eq same_res prods; + val nts0 = map (fn Prod(ty,_,_,_)::_ => ty) partitioned0; + val nts' = distinct(flat(map non_terms prods)) \\ nts0; + val nts = nts' @ nts0; + val partitioned = (replicate (len nts') []) @ partitioned0; + val ntis = nts ~~ (0 upto (len(nts)-1)); + val tab = foldr Symtab.update (ntis,Symtab.null); + + fun nt_or_vt(s,p) = + (case predef_term(s) of + end_token => let val Some(i) = Symtab.lookup(tab,s) in NT(i,p) end + | tk => T(tk)); + + fun mksyn(Terminal(t)) = [T(t)] + | mksyn(Nonterminal(t)) = [nt_or_vt t] + | mksyn(_) = []; + + fun prod2op(Prod(nt,sy,opn,p)) = + let val syA = arrayoflist(flat(map mksyn sy)) in Op(syA,opn,p) end; + + fun mkops prods = (arrayoflist(map prod2op prods),dummyTM); + + val opLA = arrayoflist(map mkops partitioned); + + val subs = array(length opLA,[]) : int list array; + fun newcp v (a,Op(syA,_,p)) = + if p=copy_pri + then let val NT(k,_) = sub(syA,0) + in if k mem v then a else k ins a end + else a; + fun reach v i = + let val new = itA ([],#1(sub(opLA,i))) (newcp v) + val v' = new union v + in flat(map (reach v') new) union v' end; + fun rch(i) = update(subs,i,reach[i]i); + +in forA(rch,subs); (tab,opLA,subs) end; + +val RootPref = "__"; + +(* Lookahead tables for speeding up parsing. Lkhd is a mapping from +nonterminals (in the form of OpList) to sets (lists) of token strings *) + +type Lkhd = Token list list list; + +(* subst_syn(s,k) syn = [ pref k ts | ts is a token string derivable from sy + under substitution s ] *) +(* This is the general version. +fun cross l1 l2 = flat(map (fn e2 => (map (fn e1 => e1@e2) l1)) l2); + +(* pref k [x1,...,xn] is [x1,...,xk] if 0<=k<=n and [x1,...,xn] otherwise *) +fun pref 0 l = [] + | pref _ [] = [] + | pref k (e::l) = e::(pref (k-1) l); + +fun subst_syn(s:Lkhd,k) = + let fun subst(ref(symb & syn)):Token list list = + let val l1 = case symb of T t => [[t]] | + NT(oplr,_) => let val Some l = assoc(s,!oplr) in l end + in distinct(map (pref k) (cross l1 (subst syn))) end | + subst _ = [[]] + in subst end; +*) +(* This one is specialized to lookahead 1 and a bit too generous *) +fun subst_syn(s:Lkhd,1) syA = + let fun subst i = if i = length(syA) then [[]] else + case sub(syA,i) of + NT(j,_) => let val pre = nth_elem(j,s) + in if [] mem pre then (pre \ []) union subst(i+1) + else pre end | + T(tk) => [[tk]]; + in subst 0 end; + +(* mk_lkhd(G,k) returns a table which associates with every nonterminal N in +G the list of pref k s for all token strings s with N -G->* s *) + +fun mk_lkhd(opLA:OpListA array,k:int):Lkhd = + let fun step(s:Lkhd):Lkhd = + let fun subst_op(l,Op(sy,_,_)) = subst_syn(s,k)sy union l; + fun step2(l,(opA,_)) = l@[itA([],opA)subst_op]; + in writeln"."; itA([],opLA)step2 end; + fun iterate(s:Lkhd):Lkhd = let val s' = step s + in if map len s = map len s' then s + else iterate s' end + in writeln"Computing lookahead tables ..."; + iterate (replicate (length opLA) []) end; + +(* create look ahead tables *) +fun mk_earley_gram(g as (tab,opLA,_):Gram):Gram = + let val lkhd = mk_lkhd(opLA,1); + fun mk_fa(i):FastAcc = + let val opA = #1(sub(opLA,i)); + fun start(j) = let val Op(sy,_,_) = sub(opA,j); + val pre = subst_syn(lkhd,1) sy + in (j,if [] mem pre then [] else map hd pre) end; + in mkTokenMap(map start (0 upto(length(opA)-1))) end; + fun updt(i) = update(opLA,i,(#1(sub(opLA,i)),mk_fa(i))); + + in forA(updt,opLA); g end; + +fun compile_xgram(roots,prods) = + let fun mk_root nt = Prod(RootPref^nt, + [Nonterminal(nt,0),Terminal(end_token)],"",0); + val prods' = (map mk_root roots) @ prods + in mk_earley_gram(mk_pre_grammar(prods')) end; + +(* State: nonterminal#, production#, index in production, + index of originating state set, + parse trees generated so far, +*) + +datatype State = St of int * int * int * int * ParseTree list +withtype StateSet = State LListR * (State -> unit) LListR; +type Compl = State -> unit; +type StateSetList = StateSet array; +(* Debugging: +val print_SL = seqll(fn St(nti,pi,ip,fs,ptl)=> +(print_int nti; prs" "; print_int pi; prs" "; print_int ip; prs" "; +print_int fs; prs" "; print_int(len ptl); prs"\n")); + +fun print_SS(s1,delr) = (writeln"================="; print_SL s1); + +fun count_ss(ref nilL) = 0 + | count_ss(ref(_ & ss)) = count_ss(ss)+1; + +fun print_stat(state_sets) = + let fun pr i = let val(s1,_)=sub(state_sets,i) + in prs" "; print_int(count_ss s1) end; + in prs"["; forA(pr,state_sets); prs"]\n" end; +*) +fun mt_stateS():StateSet = (ref nilL, ref nilL); + +fun mt_states(n):StateSetList = array(n,mt_stateS()); + +fun ismt_stateS((ref nilL,_):StateSet) = true | ismt_stateS _ = false; + +fun fst_state((ref(st & _),_): StateSet) = st; + +fun apply_all_states(f,(slr,_):StateSet) = seqll f slr; + +fun add_state(nti,pi,ip,from,ptl,(sllr,delr):StateSet) = + let fun add(ref(St(nti',pi',ip',from',_) & rest)) = + if nti=nti' andalso pi=pi' andalso ip=ip' andalso from=from' + then () + else add rest | + add(last as ref nilL) = + let val newst = St(nti,pi,ip,from,ptl) + in last := newst & ref nilL; + seqll (fn compl => compl newst) delr + end; + in add sllr end; + +fun complete(nti,syA,opn,p,ptl,ss,si as (_,delr):StateSet,opLA,rchA) = + let val pt = mk_pt(opn,ptl) + fun compl(St(ntj,pj,jp,from,ptl)) = + let val Op(syj,_,_) = sub(fst(sub(opLA,ntj)),pj) in + if jp=length(syj) then () else + case sub(syj,jp) of + NT(nt,p') => if p >= p' andalso nti mem sub(rchA,nt) + then add_state(ntj,pj,jp+1,from,ptl@[pt], si) + else () + | _ => () + end + in apply_all_states(compl,ss); + if length(syA)=0 (* delayed completion in case of empty production: *) + then delr := compl & ref(!delr) else () + end; + +fun predict(tk,isi,si,p',opLA) = fn nti => + let val (opA,tm) = sub(opLA,nti); + fun add(pi) = let val opr as Op(syA,_,p) = sub(opA,pi) + in if p < p' then () else add_state(nti,pi,0,isi,[],si) end + in seq add (applyTokenMap(tm,tk)) end; + + +fun parsable((tab,_,_):Gram, root:string) = + not(Symtab.lookup(tab,RootPref^root) = None); + +exception SYNTAX_ERR of Token list; + +fun unknown c = error("System Error - Trying to parse unknown category "^c); + +fun parse((tab,opLA,rchA):Gram, root:string, tl: Token list): ParseTree = + let val tl' = tl@[end_token]; + val state_sets = mt_states(len tl' +1); + val s0 = mt_stateS(); + val rooti = case Symtab.lookup(tab,RootPref^root) of + Some(ri) => ri | None => unknown root; + + fun lr (tl,isi,si,t) = + if ismt_stateS(si) then raise SYNTAX_ERR(t::tl) else + case tl of + [] => () | + t::tl => + let val si1 = mt_stateS(); + fun process(St(nti,pi,ip,from,ptl)) = + let val opA = #1(sub(opLA,nti)) + val Op(syA,opn,p) = sub(opA,pi) in + if ip = length(syA) + then complete(nti,syA,opn,p,ptl, + sub(state_sets,from),si,opLA,rchA) + else case sub(syA,ip) of + NT(ntj,p) => + seq (predict(t,isi,si,p,opLA)) (sub(rchA,ntj)) + | T(t') => + if matching_tokens(t,t') + then add_state(nti,pi,ip+1,from, + if valued_token(t) + then ptl@[Tip(t)] else ptl, + si1) + else () end; + + in apply_all_states(process,si); + update(state_sets,isi+1,si1); + lr(tl,isi+1,si1,t) end + + in update(state_sets,0,s0); + add_state(rooti,0,0,0,[],s0); + lr(tl',0,s0,end_token(*dummy*)); + (*print_stat state_sets;*) + let val St(_,_,_,_,[pt]) = fst_state(sub(state_sets,len tl')) + in pt end + end; + +fun print_gram ((st,opAA,rchA):Gram,lex) = + let val tts = Lexicon.name_of_token; + val al = map (fn (x,y)=>(y,x)) (Symtab.alist_of st); + fun nt i = let val Some(s) = assoc(al,i) in s end; + fun print_sy(T(end_token)) = prs". " | + print_sy(T(tk)) = (prs(tts tk); prs" ") | + print_sy(NT(i,p)) = (prs((nt i)^"[");print_int p;prs"] "); + fun print_opA(i) = + let val lhs = nt i; + val (opA,_)=sub(opAA,i); + fun print_op(j) = + let val Op(sy,n,p) = sub(opA,j) + in prs(lhs^" = "); forA(fn i=>print_sy(sub(sy,i)),sy); + if n="" then () else prs(" => \""^n^"\""); + prs" (";print_int p;prs")\n" + end; + in forA(print_op,opA) end; + fun print_rch(i) = (print_int i; prs" -> "; + print_list("[","]\n",print_int) (sub(rchA,i))) + in forA(print_opA,opAA) (*; forA(print_rch,rchA) *) end; + +end; + +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/extension.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/extension.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,301 @@ +(* Title: Pure/Syntax/extension + ID: $Id$ + Author: Tobias Nipkow and Markus Wenzel, TU Muenchen + +Syntax definition (internal interface) +*) + +signature EXTENSION0 = +sig + val typeT: typ + val constrainC: string +end; + +signature EXTENSION = +sig + include EXTENSION0 + structure XGram: XGRAM + local open XGram XGram.Ast in + datatype mfix = Mfix of string * typ * string * int list * int + datatype ext = + Ext of { + roots: string list, + mfix: mfix list, + extra_consts: string list, + parse_ast_translation: (string * (ast list -> ast)) list, + parse_preproc: (ast -> ast) option, + parse_postproc: (ast -> ast) option, + parse_translation: (string * (term list -> term)) list, + print_translation: (string * (term list -> term)) list, + print_preproc: (ast -> ast) option, + print_postproc: (ast -> ast) option, + print_ast_translation: (string * (ast list -> ast)) list} + datatype synrules = + SynRules of { + parse_rules: (ast * ast) list, + print_rules: (ast * ast) list} + val max_pri: int + val logic: string + val id: string + val idT: typ + val var: string + val varT: typ + val tfree: string + val tfreeT: typ + val tvar: string + val tvarT: typ + val typ_to_nt: typ -> string + val applC: string + val args: string + val empty_synrules: synrules + val empty: string xgram + val extend: string xgram -> (ext * synrules) -> string xgram + end +end; + +functor ExtensionFun(XGram: XGRAM): EXTENSION = +struct + +structure XGram = XGram; +open XGram XGram.Ast; + + +(** datatype ext **) + +(* Mfix (sy, ty, c, pl, p): + sy: production as symbolic string + ty: type description of production + c: corresponding Isabelle Const + pl: priorities of nonterminals in sy + p: priority of production +*) + +datatype mfix = Mfix of string * typ * string * int list * int; + +datatype ext = + Ext of { + roots: string list, + mfix: mfix list, + extra_consts: string list, + parse_ast_translation: (string * (ast list -> ast)) list, + parse_preproc: (ast -> ast) option, + parse_postproc: (ast -> ast) option, + parse_translation: (string * (term list -> term)) list, + print_translation: (string * (term list -> term)) list, + print_preproc: (ast -> ast) option, + print_postproc: (ast -> ast) option, + print_ast_translation: (string * (ast list -> ast)) list}; + +datatype synrules = + SynRules of { + parse_rules: (ast * ast) list, + print_rules: (ast * ast) list}; + + +(* empty_synrules *) + +val empty_synrules = SynRules {parse_rules = [], print_rules = []}; + + +(* empty xgram *) + +val empty = + XGram { + roots = [], prods = [], consts = [], + parse_ast_translation = [], + parse_preproc = None, + parse_rules = [], + parse_postproc = None, + parse_translation = [], + print_translation = [], + print_preproc = None, + print_rules = [], + print_postproc = None, + print_ast_translation = []}; + + + +(** syntactic constants etc. **) + +val max_pri = 1000; (*maximum legal priority*) + +val logic = "logic"; +val logicT = Type (logic, []); + +val logic1 = "logic1"; +val logic1T = Type (logic1, []); + +val funT = Type ("fun", []); + + +(* terminals *) + +val id = "id"; +val idT = Type (id, []); + +val var = "var"; +val varT = Type (var, []); + +val tfree = "tfree"; +val tfreeT = Type (tfree, []); + +val tvar = "tvar"; +val tvarT = Type (tvar, []); + +val terminalTs = [idT, varT, tfreeT, tvarT]; + + +val args = "args"; +val argsT = Type (args, []); + +val typeT = Type ("type", []); + +val applC = "_appl"; +val constrainC = "_constrain"; + + +fun typ_to_nt (Type (c, _)) = c + | typ_to_nt _ = logic; + + + +(** extend xgram **) (* FIXME clean *) + +fun nonts syn = foldl (fn (i, "_") => i + 1 | (i, _) => i) (0, explode syn); + +val meta_chs = ["(", ")", "/", "_"]; + +fun mk_term(pref, []) = (pref, []) + | mk_term(pref, "'"::c::cl) = mk_term(pref^c, cl) + | mk_term(pref, l as c::cl) = if is_blank(c) orelse c mem meta_chs + then (pref, l) else mk_term(pref^c, cl); + +fun mk_space(sp, []) = (sp, []) | + mk_space(sp, cl as c::cl') = + if is_blank(c) then mk_space(sp^c, cl') else (sp, cl); + +exception ARG_EXN; +exception BLOCK_EXN; + +fun mk_syntax([], ar, _, b, sy) = if b=0 then (sy, ar) else raise BLOCK_EXN + | mk_syntax("_"::cs, Type("fun", [ar, ar']), [], b, sy) = + mk_syntax(cs, ar', [], b, sy@[Nonterminal(typ_to_nt ar, 0)]) + | mk_syntax("_"::cs, Type("fun", [ar, ar']), p::pl, b, sy) = + mk_syntax(cs, ar', pl, b, sy@[Nonterminal(typ_to_nt ar, p)]) + | mk_syntax("_"::cs, _, _, _, _) = raise ARG_EXN + | mk_syntax("("::cs, ar, pl, b, sy) = let val (i, cs') = scan_int cs + in mk_syntax(cs', ar, pl, b+1, sy@[Bg(i)]) end + | mk_syntax(")"::cs, ar, pl, b, sy) = + if b>0 then mk_syntax(cs, ar, pl, b-1, sy@[En]) else raise BLOCK_EXN + | mk_syntax("/"::cs, ar, pl, b, sy) = let val (sp, cs') = take_prefix is_blank cs + in mk_syntax(cs', ar, pl, b, sy@[Brk(length sp)]) end + | mk_syntax(c::cs, ar, pl, b, sy) = + let val (term, rest) = + if is_blank(c) + then let val (sp, cs') = mk_space(c, cs) in (Space(sp), cs') end + else let val (tk, cs') = mk_term("", c::cs) in(Terminal(tk), cs') end + in mk_syntax(rest, ar, pl, b, sy@[term]) end; + +fun pri_test1 p = if 0 <= p andalso p <= max_pri then () + else error("Priority out of range: " ^ string_of_int p) +fun pri_test(pl, p) = (pri_test1 p; seq pri_test1 pl); + +fun mk_prod2(sy, T, opn, pl, p) = +let val (syn, T') = mk_syntax(explode sy, T, pl, 0, []) handle + ARG_EXN => + error("More arguments in "^sy^" than in corresponding type") | + BLOCK_EXN => error("Unbalanced block parantheses in "^sy); + val nt = case T' of Type(c, _) => c | _ => logic1; +in Prod(nt, syn, opn, p) end; + +fun mk_prod1(sy, T, opn, pl, p) = (pri_test(pl, p); mk_prod2(sy, T, opn, pl, p)); + + +fun terminal1(T as Type("fun", _)) = hd(binder_types T) mem terminalTs + | terminal1 _ = false; + +fun mk_prod(Mfix(sy, T, "", pl, p)) = if nonts sy <> 1 + then error"Copy op must have exactly one argument" else + if filter_out is_blank (explode sy) = ["_"] andalso + not(terminal1 T) + then mk_prod2(sy, T, "", [copy_pri], copy_pri) + else mk_prod1(sy, T, "", pl, p) + | mk_prod(Mfix(sy, T, const, pl, p)) = mk_prod1(sy, T, const, pl, p) + + + +fun extend (XGram xgram) (Ext ext, SynRules rules) = + let + infix oo; + + fun None oo None = None + | (Some f) oo None = Some f + | None oo (Some g) = Some g + | (Some f) oo (Some g) = Some (f o g); + + fun descend (from, to) = Mfix ("_", to --> from, "", [0], 0); + + fun parents T = Mfix ("(1'(_'))", T --> T, "", [0], max_pri); + + fun mkappl T = + Mfix ("_(1'(_'))", [funT, argsT] ---> T, applC, [max_pri, 0], max_pri); + + fun mkid T = Mfix ("_", idT --> T, "", [], max_pri); + + fun mkvar T = Mfix ("_", varT --> T, "", [], max_pri); + + fun constrain T = + Mfix ("_::_", [T, typeT]--->T, constrainC, [max_pri, 0], max_pri - 1); + + + val {roots = roots1, prods, consts, + parse_ast_translation = parse_ast_translation1, + parse_preproc = parse_preproc1, + parse_rules = parse_rules1, + parse_postproc = parse_postproc1, + parse_translation = parse_translation1, + print_translation = print_translation1, + print_preproc = print_preproc1, + print_rules = print_rules1, + print_postproc = print_postproc1, + print_ast_translation = print_ast_translation1} = xgram; + + val {roots = roots2, mfix, extra_consts, + parse_ast_translation = parse_ast_translation2, + parse_preproc = parse_preproc2, + parse_postproc = parse_postproc2, + parse_translation = parse_translation2, + print_translation = print_translation2, + print_preproc = print_preproc2, + print_postproc = print_postproc2, + print_ast_translation = print_ast_translation2} = ext; + + val {parse_rules = parse_rules2, print_rules = print_rules2} = rules; + + val Troots = map (apr (Type, [])) (roots2 \\ roots1); + val Troots' = Troots \\ [typeT, propT, logicT]; + val mfix' = mfix @ map parents (Troots \ logicT) @ map mkappl Troots' @ + map mkid Troots' @ map mkvar Troots' @ map constrain Troots' @ + map (apl (logicT, descend)) (Troots \\ [typeT, logicT]) @ + map (apr (descend, logic1T)) Troots'; + in + XGram { + roots = distinct (roots1 @ roots2), +(* roots = roots1 union roots2, *) (* FIXME remove *) + prods = prods @ map mk_prod mfix', + consts = consts union extra_consts, + parse_ast_translation = parse_ast_translation1 @ parse_ast_translation2, + parse_preproc = parse_preproc1 oo parse_preproc2, + parse_rules = parse_rules1 @ parse_rules2, + parse_postproc = parse_postproc2 oo parse_postproc1, + parse_translation = parse_translation1 @ parse_translation2, + print_translation = print_translation1 @ print_translation2, + print_preproc = print_preproc1 oo print_preproc2, + print_rules = print_rules1 @ print_rules2, + print_postproc = print_postproc2 oo print_postproc1, + print_ast_translation = print_ast_translation1 @ print_ast_translation2} + end; + + +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/lexicon.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/lexicon.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,252 @@ +(* Title: Lexicon + ID: $Id$ + Author: Tobias Nipkow, TU Muenchen + Copyright 1993 TU Muenchen + +The Isabelle Lexer + +Changes: +TODO: + Lexicon -> lexicon, Token -> token + end_token -> EndToken ? +*) + +signature LEXICON0 = +sig + val is_identifier: string -> bool + val scan_varname: string list -> indexname * string list + val string_of_vname: indexname -> string +end; + +signature LEXICON = +sig + type Lexicon + datatype Token = Token of string + | IdentSy of string + | VarSy of string * int + | TFreeSy of string + | TVarSy of string * int + | end_token; + val empty: Lexicon + val extend: Lexicon * string list -> Lexicon + val matching_tokens: Token * Token -> bool + val mk_lexicon: string list -> Lexicon + val name_of_token: Token -> string + val predef_term: string -> Token + val is_terminal: string -> bool + val tokenize: Lexicon -> string -> Token list + val token_to_string: Token -> string + val valued_token: Token -> bool + type 'b TokenMap + val mkTokenMap: ('b * Token list) list -> 'b TokenMap + val applyTokenMap: 'b TokenMap * Token -> 'b list + include LEXICON0 +end; + +functor LexiconFun(Extension: EXTENSION): LEXICON = +struct + +open Extension; + + +datatype Token = Token of string + | IdentSy of string + | VarSy of string * int + | TFreeSy of string + | TVarSy of string * int + | end_token; +val no_token = ""; + +datatype dfa = Tip | Branch of string * string * dfa * dfa * dfa; + +type Lexicon = dfa; + +fun is_id(c::cs) = is_letter(c) andalso forall is_letdig cs + | is_id([]) = false; + +fun is_identifier s = is_id(explode s); + +val empty = Tip; + +fun extend1(dfa, s) = + let fun add(Branch(c, a, less, eq, gr), cs as (d::ds)) = + if c>d then Branch(c, a, add(less, cs), eq, gr) else + if cd then next_dfa(gr, c) else Some(a, eq); + +exception ID of string * string list; + +val eof_id = "End of input - identifier expected.\n"; + +(*A string of letters, digits, or ' _ *) +fun xscan_ident exn = +let fun scan [] = raise exn(eof_id, []) + | scan(c::cs) = + if is_letter c + then let val (ds, tail) = take_prefix is_letdig cs + in (implode(c::ds), tail) end + else raise exn("Identifier expected: ", c::cs) +in scan end; + +(*Scan the offset of a Var, if present; otherwise ~1 *) +fun scan_offset cs = case cs of + ("."::[]) => (~1, cs) + | ("."::(ds as c::cs')) => if is_digit c then scan_int ds else (~1, cs) + | _ => (~1, cs); + +fun split_varname s = + let val (rpost, rpref) = take_prefix is_digit (rev(explode s)) + val (i, _) = scan_int(rev rpost) + in (implode(rev rpref), i) end; + +fun xscan_varname exn cs : (string*int) * string list = +let val (a, ds) = xscan_ident exn cs; + val (i, es) = scan_offset ds +in if i = ~1 then (split_varname a, es) else ((a, i), es) end; + +fun scan_varname s = xscan_varname ID s + handle ID(err, cs) => error(err^(implode cs)); + +fun tokenize (dfa) (s:string) : Token list = +let exception LEX_ERR; + exception FAIL of string * string list; + val lexerr = "Lexical error: "; + + fun tokenize1 (_:dfa, []:string list) : Token * string list = + raise LEX_ERR + | tokenize1(dfa, c::cs) = + case next_dfa(dfa, c) of + None => raise LEX_ERR + | Some(a, dfa') => + if a=no_token then tokenize1(dfa', cs) + else (tokenize1(dfa', cs) handle LEX_ERR => + if is_identifier a andalso not(null cs) andalso + is_letdig(hd cs) + then raise LEX_ERR else (Token(a), cs)); + + fun token(cs) = tokenize1(dfa, cs) handle LEX_ERR => raise FAIL(lexerr, cs); + + fun id([]) = raise FAIL(eof_id, []) + | id(cs as c::cs') = + if is_letter(c) + then let val (id, cs'') = xscan_ident FAIL cs + in (IdentSy(id), cs'') end + else + if c = "?" + then case cs' of + "'"::xs => let val ((a, i), ys) = xscan_varname FAIL xs + in (TVarSy("'"^a, i), ys) end + | _ => let val ((a, i), ys) = xscan_varname FAIL cs' + in (VarSy(a, i), ys) end + else + if c = "'" + then let val (a, cs'') = xscan_ident FAIL cs' + in (TFreeSy("'" ^ a), cs'') end + else raise FAIL(lexerr, cs); + + fun tknize([], ts) = rev(ts) + | tknize(cs as c::cs', ts) = + if is_blank(c) then tknize(cs', ts) else + let val (t, cs'') = + if c="?" then id(cs) handle FAIL _ => token(cs) + else (token(cs) handle FAIL _ => id(cs)) + in tknize(cs'', t::ts) end + +in tknize(explode s, []) handle FAIL(s, cs) => error(s^(implode cs)) end; + +(*Variables have the form ?nameidx, or ?name.idx if "name" ends with a digit*) +fun string_of_vname (a, idx) = case rev(explode a) of + [] => "?NULL_VARIABLE_NAME" + | c::_ => "?" ^ + (if is_digit c then a ^ "." ^ string_of_int idx + else if idx = 0 then a + else a ^ string_of_int idx); + +fun token_to_string (Token(s)) = s + | token_to_string (IdentSy(s)) = s + | token_to_string (VarSy v) = string_of_vname v + | token_to_string (TFreeSy a) = a + | token_to_string (TVarSy v) = string_of_vname v + | token_to_string end_token = "\n"; + +(* MMW *) +fun name_of_token (Token s) = quote s + | name_of_token (IdentSy _) = id + | name_of_token (VarSy _) = var + | name_of_token (TFreeSy _) = tfree + | name_of_token (TVarSy _) = tvar; + +fun matching_tokens(Token(i), Token(j)) = (i=j) | + matching_tokens(IdentSy(_), IdentSy(_)) = true | + matching_tokens(VarSy _, VarSy _) = true | + matching_tokens(TFreeSy _, TFreeSy _) = true | + matching_tokens(TVarSy _, TVarSy _) = true | + matching_tokens(end_token, end_token) = true | + matching_tokens(_, _) = false; + +fun valued_token(end_token) = false + | valued_token(Token _) = false + | valued_token(IdentSy _) = true + | valued_token(VarSy _) = true + | valued_token(TFreeSy _) = true + | valued_token(TVarSy _) = true; + +(* MMW *) +fun predef_term name = + if name = id then IdentSy name + else if name = var then VarSy (name, 0) + else if name = tfree then TFreeSy name + else if name = tvar then TVarSy (name, 0) + else end_token; + +(* MMW *) +fun is_terminal s = s mem [id, var, tfree, tvar]; + + + +type 'b TokenMap = (Token * 'b list) list * 'b list; +val first_token = 0; + +fun int_of_token(Token(tk)) = first_token | + int_of_token(IdentSy _) = first_token - 1 | + int_of_token(VarSy _) = first_token - 2 | + int_of_token(TFreeSy _) = first_token - 3 | + int_of_token(TVarSy _) = first_token - 4 | + int_of_token(end_token) = first_token - 5; + +fun lesstk(s, t) = int_of_token(s) < int_of_token(t) orelse + (case (s, t) of (Token(a), Token(b)) => a false); + +fun mkTokenMap(atll) = + let val aill = atll; + val dom = sort lesstk (distinct(flat(map snd aill))); + val mt = map fst (filter (null o snd) atll); + fun mktm(i) = + let fun add(l, (a, il)) = if i mem il then a::l else l + in (i, foldl add ([], aill)) end; + in (map mktm dom, mt) end; + +fun find_al (i) = + let fun find((j, al)::l) = if lesstk(i, j) then [] else + if matching_tokens(i, j) then al else find l | + find [] = []; + in find end; +fun applyTokenMap((l, mt), tk:Token) = mt@(find_al tk l); + + +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/parse_tree.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/parse_tree.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,63 @@ +(* Title: Pure/Syntax/parse_tree + ID: $Id$ + Author: Tobias Nipkow and Markus Wenzel, TU Muenchen +*) + +signature PARSE_TREE = +sig + structure Lexicon: LEXICON + structure Ast: AST + local open Lexicon Ast in + datatype ParseTree = + Node of string * ParseTree list | + Tip of Token + val mk_pt: string * ParseTree list -> ParseTree + val pt_to_ast: (string -> (ast list -> ast) option) -> ParseTree -> ast + end +end; + +functor ParseTreeFun(structure Lexicon: LEXICON and Ast: AST): PARSE_TREE = +struct + +structure Lexicon = Lexicon; +structure Ast = Ast; +open Lexicon Ast; + + +(* datatype ParseTree *) (* FIXME rename *) + +datatype ParseTree = + Node of string * ParseTree list | + Tip of Token; + + +(* mk_pt *) + +fun mk_pt("", [pt]) = pt + | mk_pt("", _) = error "mk_pt: funny copy op in parse tree" + | mk_pt(s, ptl) = Node (s, ptl); + + +(* pt_to_ast *) + +fun pt_to_ast trf pt = + let + fun trans a args = + (case trf a of + None => mk_appl (Constant a) args + | Some f => + (f args) handle _ => error ("pt_to_ast: error in translation for " ^ a)); + + fun trav (Node (a, pts)) = trans a (map trav pts) + | trav (Tip (IdentSy x)) = Variable x + | trav (Tip (VarSy xi)) = Variable (string_of_vname xi) + | trav (Tip (TFreeSy x)) = Variable x + | trav (Tip (TVarSy xi)) = Variable (string_of_vname xi) + | trav (Tip _) = error "pt_to_ast: unexpected token in parse tree"; + in + trav pt + end; + + +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/pretty.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/pretty.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,191 @@ +(* Title: Pure/Syntax/pretty + ID: $Id$ + Author: Lawrence C Paulson + Copyright 1991 University of Cambridge + +Pretty printing module + +Loosely based on + D. C. Oppen, "Pretty Printing", + ACM Transactions on Programming Languages and Systems (1980), 465-483. + +The object to be printed is given as a tree with indentation and line +breaking information. A "break" inserts a newline if the text until +the next break is too long to fit on the current line. After the newline, +text is indented to the level of the enclosing block. Normally, if a block +is broken then all enclosing blocks will also be broken. Only "inconsistent +breaks" are provided. + +The stored length of a block is used in breakdist (to treat each inner block as +a unit for breaking). +*) + +(* FIXME improve? *) +type pprint_args = (string -> unit) * (int -> unit) * (int -> unit) * (unit -> unit); + +signature PRETTY = + sig + type T + val blk : int * T list -> T + val brk : int -> T + val fbrk : T + val str : string -> T + val lst : string * string -> T list -> T + val quote : T -> T + val string_of : T -> string + val str_of : T -> string + val pprint : T -> pprint_args -> unit + val setdepth: int -> unit + val setmargin: int -> unit + end; + +functor PrettyFun () : PRETTY = +struct + +(*Printing items: compound phrases, strings, and breaks*) +datatype T = Block of T list * int * int (*indentation, length*) + | String of string + | Break of bool*int (*mandatory flag; width if not taken*); + +(*Add the lengths of the expressions until the next Break; if no Break then + include "after", to account for text following this block. *) +fun breakdist (Block(_,_,len)::es, after) = len + breakdist(es, after) + | breakdist (String s :: es, after) = size s + breakdist (es, after) + | breakdist (Break _ :: es, after) = 0 + | breakdist ([], after) = after; + +fun repstring a 0 = "" + | repstring a 1 = a + | repstring a k = + if k mod 2 = 0 then repstring(a^a) (k div 2) + else repstring(a^a) (k div 2) ^ a; + +(*** Type for lines of text: string, # of lines, position on line ***) + +type text = {tx: string, nl: int, pos: int}; + +val emptytext = {tx="", nl=0, pos=0}; + +fun blanks wd {tx,nl,pos} = + {tx = tx ^ repstring" "wd, + nl = nl, + pos = pos+wd}; + +fun newline {tx,nl,pos} = + {tx = tx ^ "\n", + nl = nl+1, + pos = 0}; + +fun string s {tx,nl,pos} = + {tx = tx ^ s, + nl = nl, + pos = pos + size(s)}; + +(*** Formatting ***) + +val margin = ref 80 (*right margin, or page width*) +and breakgain = ref 4 (*minimum added space required of a break*) +and emergencypos = ref 40; (*position too far to right*) + +fun setmargin m = + (margin := m; + breakgain := !margin div 20; + emergencypos := !margin div 2); + +fun forcenext [] = [] + | forcenext (Break(_,wd) :: es) = Break(true,0) :: es + | forcenext (e :: es) = e :: forcenext es; + +fun format ([], _, _) text = text + | format (e::es, blockin, after) (text as {pos,nl,...}) = + (case e of + Block(bes,indent,wd) => + let val blockin' = (pos + indent) mod !emergencypos + val btext = format(bes, blockin', breakdist(es,after)) text + val es2 = if nl < #nl(btext) then forcenext es + else es + in format (es2,blockin,after) btext end + | String s => format (es,blockin,after) (string s text) + | Break(force,wd) => (* no break if text fits on this line + or if breaking would add only breakgain to space *) + format (es,blockin,after) + (if not force andalso + pos+wd <= max[!margin - breakdist(es,after), + blockin + !breakgain] + then blanks wd text + else blanks blockin (newline text))); + +(*** Exported functions to create formatting expressions ***) + +fun length (Block(_,_,len)) = len + | length (String s) = size s + | length (Break(_,wd)) = wd; + +val str = String +and fbrk = Break(true,0); + +fun brk wd = Break(false,wd); + +fun blk (indent,es) = + let fun sum([], k) = k + | sum(e::es, k) = sum(es, length e + k) + in Block(es,indent, sum(es,0)) end; + +fun lst(lp,rp) es = + let fun add(e,es) = str"," :: brk 1 :: e :: es; + fun list(e :: (es as _::_)) = str lp :: e :: foldr add (es,[str rp]) + | list(e::[]) = [str lp, e, str rp] + | list([]) = [] + in blk(size lp, list es) end; + +fun quote prt = blk (1, [str "\"", prt, str "\""]); + + +(*** Pretty printing with depth limitation ***) + +val depth = ref 0; (*maximum depth; 0 means no limit*) + +fun setdepth dp = (depth := dp); + + +fun pruning dp (Block(bes,indent,wd)) = + if dp>0 then blk(indent, map (pruning(dp-1)) bes) + else String"..." + | pruning dp e = e; + +fun prune dp e = if dp>0 then pruning dp e else e; + + +fun string_of e = #tx (format ([prune (!depth) e],0,0) emptytext); + + +fun brk_width (force, wd) = if force andalso wd = 0 then 1 else wd; + +fun str_of prt = + let + fun s_of (Block (prts, _, _)) = implode (map s_of prts) + | s_of (String s) = s + | s_of (Break f_w) = repstring " " (brk_width f_w); + in + s_of (prune (! depth) prt) + end; + + +fun pprint prt (put_str, begin_blk, put_brk, end_blk) = + let + fun pp (Block (prts, ind, _)) = (begin_blk ind; pp_lst prts; end_blk ()) + | pp (String s) = put_str s + | pp (Break f_w) = put_brk (brk_width f_w) + and pp_lst [] = () + | pp_lst (prt :: prts) = (pp prt; pp_lst prts); + in + pp (prune (! depth) prt) + end; + + +(*** Initialization ***) + +val () = setmargin 80; + +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/printer.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/printer.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,307 @@ +(* Title: Pure/Syntax/printer + ID: $Id$ + Author: Tobias Nipkow and Markus Wenzel, TU Muenchen +*) + +signature PRINTER0 = +sig + val show_types: bool ref + val show_sorts: bool ref +end; + +signature PRINTER = +sig + include PRINTER0 + structure Symtab: SYMTAB + structure XGram: XGRAM + structure Pretty: PRETTY + local open XGram XGram.Ast in + val pretty_ast: ast -> Pretty.T + val pretty_rule: (ast * ast) -> Pretty.T + val term_to_ast: (string -> (term list -> term) option) -> term -> ast + val typ_to_ast: (string -> (term list -> term) option) -> typ -> ast + type tab + val mk_print_tab: (string prod list) -> (ast list -> ast) Symtab.table -> tab + val pretty_term_ast: tab -> ast -> Pretty.T + val pretty_typ_ast: tab -> ast -> Pretty.T + end +end; + +functor PrinterFun(structure Symtab: SYMTAB and Lexicon: LEXICON + and TypeExt: TYPE_EXT and SExtension: SEXTENSION and Pretty: PRETTY + sharing TypeExt.Extension = SExtension.Extension) (*: PRINTER *) = (* FIXME *) +struct + +structure Symtab = Symtab; +structure Extension = TypeExt.Extension; +structure XGram = Extension.XGram; +structure Pretty = Pretty; +open XGram XGram.Ast Lexicon TypeExt Extension SExtension; + + +(** options for printing **) + +val show_types = ref false; +val show_sorts = ref false; + + + +(** simple prettying **) + +(* pretty_ast *) + +fun pretty_ast (Constant a) = Pretty.str (quote a) + | pretty_ast (Variable x) = Pretty.str x + | pretty_ast (Appl asts) = + Pretty.blk + (2, (Pretty.str "(") :: + (separate (Pretty.brk 1) (map pretty_ast asts)) @ [Pretty.str ")"]); + + +(* pretty_rule *) + +fun pretty_rule (lhs, rhs) = + Pretty.blk + (2, [pretty_ast lhs, Pretty.str " ->", Pretty.brk 2, pretty_ast rhs]); + + + +(** convert term/typ to ast **) (* FIXME check *) + +fun apply_trans a f args = + ((f args) handle + Match => raise Match + | exn => (writeln ("Error in translation for " ^ quote a); raise exn)); + + +fun ast_of_term trf show_types show_sorts tm = + let + val aprop_const = Const (apropC, dummyT); + + fun fix_aprop (tm as Const _) = tm + | fix_aprop (tm as Free (x, ty)) = + if ty = propT then aprop_const $ Free (x, dummyT) else tm + | fix_aprop (tm as Var (xi, ty)) = + if ty = propT then aprop_const $ Var (xi, dummyT) else tm + | fix_aprop (tm as Bound _) = tm + | fix_aprop (Abs (x, ty, t)) = Abs (x, ty, fix_aprop t) + | fix_aprop (t1 $ t2) = fix_aprop t1 $ fix_aprop t2; + + + fun prune_typs (t_seen as (Const _, _)) = t_seen + | prune_typs (t as Free (x, ty), seen) = + if ty = dummyT then (t, seen) + else if t mem seen then (Free (x, dummyT), seen) + else (t, t :: seen) + | prune_typs (t as Var (xi, ty), seen) = + if ty = dummyT then (t, seen) + else if t mem seen then (Var (xi, dummyT), seen) + else (t, t :: seen) + | prune_typs (t_seen as (Bound _, _)) = t_seen + | prune_typs (Abs (x, ty, t), seen) = + let val (t', seen') = prune_typs (t, seen); + in (Abs (x, ty, t'), seen') end + | prune_typs (t1 $ t2, seen) = + let + val (t1', seen') = prune_typs (t1, seen); + val (t2', seen'') = prune_typs (t2, seen'); + in + (t1' $ t2', seen'') + end; + + + fun ast_of (Const (a, _)) = trans a [] + | ast_of (Free (x, ty)) = constrain x (Free (x, dummyT)) ty + | ast_of (Var (xi, ty)) = constrain (string_of_vname xi) (Var (xi, dummyT)) ty + | ast_of (Bound i) = Variable ("B." ^ string_of_int i) + | ast_of (t as Abs _) = ast_of (abs_tr' t) + | ast_of (t as _ $ _) = + (case strip_comb t of + (Const (a, _), args) => trans a args + | (f, args) => Appl (map ast_of (f :: args))) + + and trans a args = + (case trf a of + Some f => ast_of (apply_trans a f args) + | None => raise Match) + handle Match => mk_appl (Constant a) (map ast_of args) + + and constrain x t ty = + if show_types andalso ty <> dummyT then + ast_of (Const (constrainC, dummyT) $ t $ term_of_typ show_sorts ty) + else Variable x; + in + if show_types then ast_of (fst (prune_typs (fix_aprop tm, []))) + else ast_of (fix_aprop tm) + end; + + +(* term_to_ast *) + +fun term_to_ast trf tm = + ast_of_term trf (! show_types) (! show_sorts) tm; + + +(* typ_to_ast *) + +fun typ_to_ast trf ty = + ast_of_term trf false false (term_of_typ (! show_sorts) ty); + + + +(** type tab **) + +datatype symb = + Arg of int | + TypArg of int | + String of string | + Break of int | + Block of int * symb list; + +datatype format = + Prnt of symb list * int * int | + Trns of ast list -> ast | + TorP of (ast list -> ast) * (symb list * int * int); + +type tab = format Symtab.table; + + + +(** mk_print_tab **) + +fun mk_print_tab prods ast_trans = + let + fun nargs (Arg _ :: symbs) = nargs symbs + 1 + | nargs (TypArg _ :: symbs) = nargs symbs + 1 + | nargs (String _ :: symbs) = nargs symbs + | nargs (Break _ :: symbs) = nargs symbs + | nargs (Block (_, bsymbs) :: symbs) = nargs symbs + nargs bsymbs + | nargs [] = 0; + + fun merge (s, String s' :: l) = String (s ^ s') :: l + | merge (s, l) = String s :: l; + + fun mk_prnp sy opn p = + let + val constr = (opn = constrainC orelse opn = constrainIdtC); + + fun syn2pr (Terminal s :: sy) = + let val (symbs, sy') = syn2pr sy + in (merge (s, symbs), sy') end + | syn2pr (Space s :: sy) = + let val (symbs, sy') = syn2pr sy + in (merge (s, symbs), sy') end + | syn2pr (Nonterminal (s, p) :: sy) = + let + val (symbs, sy') = syn2pr sy; + val symb = + (if constr andalso s = "type" then TypArg else Arg) + (if is_terminal s then 0 else p); + in (symb :: symbs, sy') end + | syn2pr (Bg i :: sy) = + let + val (bsymbs, sy') = syn2pr sy; + val (symbs, sy'') = syn2pr sy'; + in (Block (i, bsymbs) :: symbs, sy'') end + | syn2pr (Brk i :: sy) = + let val (symbs, sy') = syn2pr sy + in (Break i :: symbs, sy') end + | syn2pr (En :: sy) = ([], sy) + | syn2pr [] = ([], []); + + val (pr, _) = syn2pr sy; + in + (pr, nargs pr, p) + end; + + fun add_prod (Prod (_, _, "", _), tab) = tab + | add_prod (Prod (_, sy, opn, p), tab) = + (case Symtab.lookup (tab, opn) of + None => Symtab.update ((opn, Prnt (mk_prnp sy opn p)), tab) + | Some (Prnt _) => tab + | Some (Trns f) => Symtab.update ((opn, TorP (f, mk_prnp sy opn p)), tab) + | Some (TorP _) => tab); + + fun add_tr (tab, (opn, f)) = Symtab.update_new ((opn, Trns f), tab); + in + Symtab.balance + (foldr add_prod + (prods, foldl add_tr (Symtab.null, Symtab.alist_of ast_trans))) + end; + + + +(** pretty term/typ asts **) (* FIXME check *) + +(*assumes a syntax derived from Pure*) + +fun pretty tab appT ast0 p0 = + let + fun synT ([], args) = ([], args) + | synT (Arg p :: symbs, t :: args) = + let val (Ts, args') = synT (symbs, args) + in (astT (t, p) @ Ts, args') end + | synT (TypArg p :: symbs, t :: args) = + let val (Ts, args') = synT (symbs, args) + in (pretty tab tappl_ast_tr' t p @ Ts, args') end + | synT (String s :: symbs, args) = + let val (Ts, args') = synT (symbs, args) + in (Pretty.str s :: Ts, args') end + | synT (Block (i, bsymbs) :: symbs, args) = + let + val (bTs, args') = synT (bsymbs, args); + val (Ts, args'') = synT (symbs, args'); + in (Pretty.blk (i, bTs) :: Ts, args'') end + | synT (Break i :: symbs, args) = + let val (Ts, args') = synT (symbs, args) + in (Pretty.brk i :: Ts, args') end + | synT (_ :: _, []) = error "synT" + + and parT (pr, args, p, p': int) = + if p > p' then fst (synT ([Block(1, String"(" :: pr @ [String")"])], args)) + else fst (synT (pr, args)) + + and prefixT (_, a, [], _) = [Pretty.str a] + | prefixT (c, _, args, p) = astT (appT (c, args), p) + + and combT (tup as (c, a, args, p)) = + let + val nargs = length args; + fun prnt (pr, n, p') = + if n = nargs then parT (pr, args, p, p') + else if n > nargs then prefixT tup + else astT (appT (c, args), p); (* FIXME ??? *) + in + case Symtab.lookup (tab, a) of + None => prefixT tup + | Some (Prnt prnp) => prnt prnp + | Some (Trns f) => + (astT (apply_trans a f args, p) handle Match => prefixT tup) + | Some (TorP (f, prnp)) => + (astT (apply_trans a f args, p) handle Match => prnt prnp) + end + + and astT (c as Constant a, p) = combT (c, a, [], p) + | astT (Variable x, _) = [Pretty.str x] + | astT (Appl ((c as Constant a) :: args), p) = combT (c, a, args, p) + | astT (Appl (f :: (args as _ :: _)), p) = astT (appT (f, args), p) + | astT (ast as (Appl _), _) = raise_ast "pretty: malformed ast" [ast]; + in + astT (ast0, p0) + end; + + +(* pretty_term_ast *) + +fun pretty_term_ast tab ast = + Pretty.blk (0, pretty tab appl_ast_tr' ast 0); + + +(* pretty_typ_ast *) + +fun pretty_typ_ast tab ast = + Pretty.blk (0, pretty tab tappl_ast_tr' ast 0); + + +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/sextension.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/sextension.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,511 @@ +(* Title: Pure/Syntax/sextension + ID: $Id$ + Author: Tobias Nipkow and Markus Wenzel, TU Muenchen + +Syntax extensions: mixfix declarations, syntax rules, infixes, binders and +the Pure syntax. + +Changes: + SEXTENSION: added Ast, xrule + changed sext + added ast_to_term + ext_of_sext: added xconsts + SEXTENSION1: added empty_sext, appl_ast_tr' + SEXTENSION1: removed appl_tr' +TODO: +*) + + +infix |-> <-| <->; + +signature SEXTENSION0 = +sig + structure Ast: AST + local open Ast in + datatype mixfix = + Mixfix of string * string * string * int list * int | + Delimfix of string * string * string | + Infixl of string * string * int | + Infixr of string * string * int | + Binder of string * string * string * int * int | + TInfixl of string * string * int | + TInfixr of string * string * int + datatype xrule = + op |-> of (string * string) * (string * string) | + op <-| of (string * string) * (string * string) | + op <-> of (string * string) * (string * string) + datatype sext = + Sext of { + mixfix: mixfix list, + parse_translation: (string * (term list -> term)) list, + print_translation: (string * (term list -> term)) list} | + NewSext of { + mixfix: mixfix list, + xrules: xrule list, + parse_ast_translation: (string * (ast list -> ast)) list, + parse_preproc: (ast -> ast) option, + parse_postproc: (ast -> ast) option, + parse_translation: (string * (term list -> term)) list, + print_translation: (string * (term list -> term)) list, + print_preproc: (ast -> ast) option, + print_postproc: (ast -> ast) option, + print_ast_translation: (string * (ast list -> ast)) list} + val eta_contract: bool ref + val mk_binder_tr: string * string -> string * (term list -> term) + val mk_binder_tr': string * string -> string * (term list -> term) + val ndependent_tr: string -> term list -> term + val dependent_tr': string * string -> term list -> term + val max_pri: int + end +end; + +signature SEXTENSION1 = +sig + include SEXTENSION0 + val empty_sext: sext + val simple_sext: mixfix list -> sext + val constants: sext -> (string list * string) list + val pure_sext: sext + val syntax_types: string list + val constrainAbsC: string +end; + +signature SEXTENSION = +sig + include SEXTENSION1 + structure Extension: EXTENSION + sharing Extension.XGram.Ast = Ast + local open Extension Ast in + val xrules_of: sext -> xrule list + val abs_tr': term -> term + val appl_ast_tr': ast * ast list -> ast + val ext_of_sext: string list -> string list -> (string -> typ) -> sext -> ext + val ast_to_term: (string -> (term list -> term) option) -> ast -> term + val constrainIdtC: string + val apropC: string + end +end; + +functor SExtensionFun(structure TypeExt: TYPE_EXT and Lexicon: LEXICON)(*: SEXTENSION *) = (* FIXME *) +struct + +structure Extension = TypeExt.Extension; +structure Ast = Extension.XGram.Ast; +open Extension Ast; + + +(** datatype sext **) + +datatype mixfix = + Mixfix of string * string * string * int list * int | + Delimfix of string * string * string | + Infixl of string * string * int | + Infixr of string * string * int | + Binder of string * string * string * int * int | + TInfixl of string * string * int | + TInfixr of string * string * int; + +datatype xrule = + op |-> of (string * string) * (string * string) | + op <-| of (string * string) * (string * string) | + op <-> of (string * string) * (string * string); + +datatype sext = + Sext of { + mixfix: mixfix list, + parse_translation: (string * (term list -> term)) list, + print_translation: (string * (term list -> term)) list} | + NewSext of { + mixfix: mixfix list, + xrules: xrule list, + parse_ast_translation: (string * (ast list -> ast)) list, + parse_preproc: (ast -> ast) option, + parse_postproc: (ast -> ast) option, + parse_translation: (string * (term list -> term)) list, + print_translation: (string * (term list -> term)) list, + print_preproc: (ast -> ast) option, + print_postproc: (ast -> ast) option, + print_ast_translation: (string * (ast list -> ast)) list}; + + +(* simple_sext *) + +fun simple_sext mixfix = + Sext {mixfix = mixfix, parse_translation = [], print_translation = []}; + + +(* empty_sext *) + +val empty_sext = simple_sext []; + + +(* sext_components *) + +fun sext_components (Sext {mixfix, parse_translation, print_translation}) = + {mixfix = mixfix, + xrules = [], + parse_ast_translation = [], + parse_preproc = None, + parse_postproc = None, + parse_translation = parse_translation, + print_translation = print_translation, + print_preproc = None, + print_postproc = None, + print_ast_translation = []} + | sext_components (NewSext cmps) = cmps; + + +(* mixfix_of *) + +fun mixfix_of (Sext {mixfix, ...}) = mixfix + | mixfix_of (NewSext {mixfix, ...}) = mixfix; + + +(* xrules_of *) + +fun xrules_of (Sext _) = [] + | xrules_of (NewSext {xrules, ...}) = xrules; + + + +(** parse translations **) + +(* application *) + +fun appl_ast_tr (*"_appl"*) [f, args] = Appl (f :: unfold_ast "_args" args) + | appl_ast_tr (*"_appl"*) asts = raise_ast "appl_ast_tr" asts; + + +(* abstraction *) + +fun idtyp_ast_tr (*"_idtyp"*) [x, ty] = Appl [Constant constrainC, x, ty] + | idtyp_ast_tr (*"_idtyp"*) asts = raise_ast "idtyp_ast_tr" asts; + +fun lambda_ast_tr (*"_lambda"*) [idts, body] = + fold_ast_p "_%" (unfold_ast "_idts" idts, body) + | lambda_ast_tr (*"_lambda"*) asts = raise_ast "lambda_ast_tr" asts; + +fun abs_tr (*"_%"*) [Free (x, T), body] = absfree (x, T, body) + | abs_tr (*"_%"*) (ts as [Const (c, _) $ Free (x, T) $ tT, body]) = + if c = constrainC then + Const ("_constrainAbs", dummyT) $ absfree (x, T, body) $ tT + else raise (TERM ("abs_tr", ts)) + | abs_tr (*"_%"*) ts = raise (TERM ("abs_tr", ts)); + + +(* binder *) (* FIXME check *) (* FIXME check *) + +fun mk_binder_tr (sy, name) = + let + val const = Const (name, dummyT); + + fun tr (Free (x, T), t) = const $ absfree (x, T, t) + | tr (Const ("_idts", _) $ idt $ idts, t) = tr (idt, tr (idts, t)) + | tr (t1 as (Const (c, _) $ Free (x, T) $ tT), t) = (* FIXME *) + if c = constrainC then + const $ (Const ("_constrainAbs", dummyT) $ absfree (x, T, t) $ tT) + else raise (TERM ("binder_tr", [t1, t])) + | tr (t1, t2) = raise (TERM ("binder_tr", [t1, t2])); + + fun binder_tr (*sy*) [idts, body] = tr (idts, body) + | binder_tr (*sy*) ts = raise (TERM ("binder_tr", ts)); + in + (sy, binder_tr) + end; + + +(* atomic props *) + +fun aprop_ast_tr (*"_aprop"*) [ast] = ast + | aprop_ast_tr (*"_aprop"*) asts = raise_ast "aprop_ast_tr" asts; + + +(* meta implication *) + +fun bigimpl_ast_tr (*"_bigimpl"*) [asms, concl] = + fold_ast_p "==>" (unfold_ast "_asms" asms, concl) + | bigimpl_ast_tr (*"_bigimpl"*) asts = raise_ast "bigimpl_ast_tr" asts; + + +(* 'dependent' type operators *) + +fun ndependent_tr q [A, B] = + Const (q, dummyT) $ A $ Abs ("x", dummyT, incr_boundvars 1 B) + | ndependent_tr _ _ = raise Match; + + + +(** print translations **) + +(* application *) + +fun appl_ast_tr' (f, []) = raise_ast "appl_ast_tr'" [f] + | appl_ast_tr' (f, args) = Appl [Constant "_appl", f, fold_ast "_args" args]; + + +(* abstraction *) (* FIXME check *) + +fun strip_abss vars_of body_of tm = + let + val vars = vars_of tm; + val body = body_of tm; + val rev_new_vars = rename_wrt_term body vars; + in + (map Free (rev rev_new_vars), + subst_bounds (map (fn (x, _) => Free (x, dummyT)) rev_new_vars, body)) + end; + +(*do (partial) eta-contraction before printing*) + +val eta_contract = ref false; + +fun eta_contr tm = + let + fun eta_abs (Abs (a, T, t)) = + (case eta_abs t of + t' as (f $ u) => + (case eta_abs u of + Bound 0 => + if not (0 mem loose_bnos f) then incr_boundvars ~1 f + else Abs (a, T, t') + | _ => Abs (a, T, t')) + | t' => Abs (a, T, t')) + | eta_abs t = t; + in + if ! eta_contract then eta_abs tm else tm + end; + + +fun abs_tr' tm = + foldr (fn (x, t) => Const ("_%", dummyT) $ x $ t) + (strip_abss strip_abs_vars strip_abs_body (eta_contr tm)); + + +fun lambda_ast_tr' (*"_%"*) asts = + (case unfold_ast_p "_%" (Appl (Constant "_%" :: asts)) of + ([], _) => raise_ast "lambda_ast_tr'" asts + | (xs, body) => Appl [Constant "_lambda", fold_ast "_idts" xs, body]); + + +(* binder *) (* FIXME check *) + +fun mk_binder_tr' (name, sy) = + let + fun mk_idts [] = raise Match (*abort translation*) + | mk_idts [idt] = idt + | mk_idts (idt :: idts) = Const ("_idts", dummyT) $ idt $ mk_idts idts; + + fun tr' t = + let + val (xs, bd) = strip_abss (strip_qnt_vars name) (strip_qnt_body name) t; + in + Const (sy, dummyT) $ mk_idts xs $ bd + end; + + fun binder_tr' (*name*) (t :: ts) = + list_comb (tr' (Const (name, dummyT) $ t), ts) + | binder_tr' (*name*) [] = raise Match; + in + (name, binder_tr') + end; + + +(* idts *) + +fun idts_ast_tr' (*"_idts"*) [Appl [Constant c, x, ty], xs] = + if c = constrainC then + Appl [Constant "_idts", Appl [Constant "_idtyp", x, ty], xs] + else raise Match + | idts_ast_tr' (*"_idts"*) _ = raise Match; + + +(* meta implication *) + +fun impl_ast_tr' (*"==>"*) asts = + (case unfold_ast_p "==>" (Appl (Constant "==>" :: asts)) of + (asms as (_ :: _ :: _), concl) + => Appl [Constant "_bigimpl", fold_ast "_asms" asms, concl] + | _ => raise Match); + + +(* 'dependent' type operators *) + +fun dependent_tr' (q, r) [A, Abs (x, T, B)] = + if 0 mem (loose_bnos B) then + let val (x', B') = variant_abs (x, dummyT, B); + in Const (q, dummyT) $ Free (x', T) $ A $ B' end + else Const (r, dummyT) $ A $ B + | dependent_tr' _ _ = raise Match; + + + +(** constants **) + +(* FIXME opn, clean: move *) +val clean = + let + fun q ("'" :: c :: cs) = c ^ q cs + | q (c :: cs) = c ^ q cs + | q ([]) = "" + in q o explode end; + +val opn = "op "; + + +fun constants sext = + let + fun consts (Delimfix (_, ty, c)) = ([c], ty) + | consts (Mixfix (_, ty, c, _, _)) = ([c], ty) + | consts (Infixl (c, ty, _)) = ([opn ^ clean c], ty) + | consts (Infixr (c, ty, _)) = ([opn ^ clean c], ty) + | consts (Binder (_, ty, c, _, _)) = ([c], ty) + | consts _ = ([""], ""); (*is filtered out below*) + in + distinct (filter_out (fn (l, _) => l = [""]) (map consts (mixfix_of sext))) + end; + + + +(** ext_of_sext **) (* FIXME check, clean *) + +fun ext_of_sext roots xconsts read_typ sext = + let + val + {mixfix, parse_ast_translation, parse_preproc, parse_postproc, + parse_translation, print_translation, print_preproc, print_postproc, + print_ast_translation, ...} = sext_components sext; + + val infixT = [typeT, typeT] ---> typeT; + + fun binder (Binder (sy, _, name, _, _)) = Some (sy, name) + | binder _ = None; + + fun binder_typ ty = + (case read_typ ty of + Type ("fun", [Type ("fun", [_, T2]), T3]) => + [Type ("idts", []), T2] ---> T3 + | _ => error (quote ty ^ " is not a valid binder type.")); + + fun mfix_of (Mixfix (sy, ty, c, pl, p)) = [Mfix (sy, read_typ ty, c, pl, p)] + | mfix_of (Delimfix (sy, ty, c)) = [Mfix (sy, read_typ ty, c, [], max_pri)] + | mfix_of (Infixl (sy, ty, p)) = + let val T = read_typ ty and c = opn ^ sy and c' = opn ^ clean sy + in + [Mfix (c, T, c', [], max_pri), + Mfix("(_ " ^ sy ^ "/ _)", T, c', [p, p + 1], p)] + end + | mfix_of (Infixr (sy, ty, p)) = + let val T = read_typ ty and c = opn ^ sy and c' = opn ^ clean sy + in + [Mfix(c, T, c', [], max_pri), + Mfix("(_ " ^ sy ^ "/ _)", T, c', [p + 1, p], p)] + end + | mfix_of (Binder (sy, ty, _, p, q)) = + [Mfix ("(3" ^ sy ^ "_./ _)", binder_typ ty, sy, [0, p], q)] + | mfix_of (TInfixl (s, c, p)) = + [Mfix ("(_ " ^ s ^ "/ _)", infixT, c, [p, p + 1], p)] + | mfix_of (TInfixr (s, c, p)) = + [Mfix ("(_ " ^ s ^ "/ _)", infixT, c, [p + 1, p], p)]; + + val mfix = flat (map mfix_of mixfix); + val mfix_consts = map (fn (Mfix (_, _, c, _, _)) => c) mfix; + val bs = mapfilter binder mixfix; + val bparses = map mk_binder_tr bs; + val bprints = map (mk_binder_tr' o (fn (x, y) => (y, x))) bs; + in + Ext { + roots = roots, mfix = mfix, + extra_consts = distinct (filter Lexicon.is_identifier (xconsts @ mfix_consts)), + parse_ast_translation = parse_ast_translation, + parse_preproc = parse_preproc, + parse_postproc = parse_postproc, + parse_translation = bparses @ parse_translation, + print_translation = bprints @ print_translation, + print_preproc = print_preproc, + print_postproc = print_postproc, + print_ast_translation = print_ast_translation} + end; + + + +(** ast_to_term **) + +fun ast_to_term trf ast = + let + fun scan_vname prfx cs = + (case Lexicon.scan_varname cs of + ((x, i), []) => Var ((prfx ^ x, i), dummyT) + | _ => error ("ast_to_term: bad variable name " ^ quote (implode cs))); + + fun vname_to_var v = + (case explode v of + "?" :: "'" :: cs => scan_vname "'" cs + | "?" :: cs => scan_vname "" cs + | _ => Free (v, dummyT)); + + fun trans a args = + (case trf a of + None => list_comb (Const (a, dummyT), args) + | Some f => ((f args) + handle _ => error ("ast_to_term: error in translation for " ^ quote a))); + + fun trav (Constant a) = trans a [] + | trav (Appl (Constant a :: (asts as _ :: _))) = trans a (map trav asts) + | trav (Appl (ast :: (asts as _ :: _))) = + list_comb (trav ast, map trav asts) + | trav (ast as (Appl _)) = raise_ast "ast_to_term: malformed ast" [ast] + | trav (Variable x) = vname_to_var x; + in + trav ast + end; + + + +(** the Pure syntax **) + +val pure_sext = + NewSext { + mixfix = [ + Mixfix ("(3%_./ _)", "[idts, 'a] => ('b => 'a)", "_lambda", [0], 0), + Delimfix ("_", "'a => " ^ args, ""), + Delimfix ("_,/ _", "['a, " ^ args ^ "] => " ^ args, "_args"), + Delimfix ("_", "id => idt", ""), + Mixfix ("_::_", "[id, type] => idt", "_idtyp", [0, 0], 0), + Delimfix ("'(_')", "idt => idt", ""), + Delimfix ("_", "idt => idts", ""), + Mixfix ("_/ _", "[idt, idts] => idts", "_idts", [1, 0], 0), + Delimfix ("_", "id => aprop", ""), + Delimfix ("_", "var => aprop", ""), + Mixfix ("_'(_')", "[('b => 'a), " ^ args ^ "] => aprop", applC, [max_pri, 0], 0), + Delimfix ("PROP _", "aprop => prop", "_aprop"), + Delimfix ("_", "prop => asms", ""), + Delimfix ("_;/ _", "[prop, asms] => asms", "_asms"), + Mixfix ("((3[| _ |]) ==>/ _)", "[asms, prop] => prop", "_bigimpl", [0, 1], 1), + Mixfix ("(_ ==/ _)", "['a::{}, 'a] => prop", "==", [3, 2], 2), + Mixfix ("(_ =?=/ _)", "['a::{}, 'a] => prop", "=?=", [3, 2], 2), + Mixfix ("(_ ==>/ _)", "[prop, prop] => prop", "==>", [2, 1], 1), + Binder ("!!", "('a::logic => prop) => prop", "all", 0, 0)], + xrules = [], + parse_ast_translation = + [(applC, appl_ast_tr), ("_lambda", lambda_ast_tr), ("_idtyp", idtyp_ast_tr), + ("_aprop", aprop_ast_tr), ("_bigimpl", bigimpl_ast_tr)], + parse_preproc = None, + parse_postproc = None, + parse_translation = [("_%", abs_tr)], + print_translation = [], + print_preproc = None, + print_postproc = None, + print_ast_translation = [("_%", lambda_ast_tr'), ("_idts", idts_ast_tr'), + ("==>", impl_ast_tr')]}; + +val syntax_types = (* FIXME clean, check *) + [logic, "aprop", args, "asms", id, "idt", "idts", tfree, tvar, "type", "types", + var, "sort", "classes"] + +val constrainIdtC = "_idtyp"; +val constrainAbsC = "_constrainAbs"; +val apropC = "_aprop"; + + +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/syntax.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/syntax.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,430 @@ +(* Title: Pure/Syntax/syntax + ID: $Id$ + Author: Tobias Nipkow and Markus Wenzel, TU Muenchen +*) + +signature SYNTAX = +sig + (* FIXME include AST0 (?) *) + include LEXICON0 + include EXTENSION0 + include TYPE_EXT0 + include SEXTENSION1 + include PRINTER0 + structure Extension: EXTENSION + structure Pretty: PRETTY + local open Extension.XGram Extension.XGram.Ast in + type syntax + val print_gram: syntax -> unit + val print_trans: syntax -> unit + val print_syntax: syntax -> unit + val read_ast: syntax -> string * string -> ast + val read: syntax -> typ -> string -> term + val pretty_term: syntax -> term -> Pretty.T + val pretty_typ: syntax -> typ -> Pretty.T + val string_of_term: syntax -> term -> string + val string_of_typ: syntax -> typ -> string + val type_syn: syntax + val extend: syntax * (indexname -> sort) -> string list * string list * sext + -> syntax + val merge: syntax * syntax -> syntax + end +end; + +functor SyntaxFun(structure Symtab: SYMTAB and TypeExt: TYPE_EXT + and Parser: PARSER and SExtension: SEXTENSION and Printer: PRINTER + sharing TypeExt.Extension.XGram = Parser.XGram = Printer.XGram + and TypeExt.Extension = SExtension.Extension + and Parser.ParseTree.Ast = Parser.XGram.Ast)(*: SYNTAX *) = (* FIXME *) +struct + +structure Extension = TypeExt.Extension; +structure XGram = Extension.XGram; +structure Lexicon = Parser.ParseTree.Lexicon; +open Lexicon Extension TypeExt SExtension Printer XGram XGram.Ast; + + +fun lookup tab a = Symtab.lookup (tab, a); + + + +(** datatype syntax **) + +datatype tables = + Tab of { + gram: Parser.Gram, + lexicon: Lexicon, + const_tab: unit Symtab.table, + parse_ast_trtab: (ast list -> ast) Symtab.table, + parse_preproc: ast -> ast, + parse_ruletab: (ast * ast) list Symtab.table, + parse_postproc: ast -> ast, + parse_trtab: (term list -> term) Symtab.table, + print_trtab: (term list -> term) Symtab.table, + print_preproc: ast -> ast, + print_ruletab: (ast * ast) list Symtab.table, + print_postproc: ast -> ast, + print_tab: Printer.tab}; + +datatype gramgraph = + EmptyGG | + ExtGG of gramgraph ref * (ext * synrules) | + MergeGG of gramgraph ref * gramgraph ref; + +datatype syntax = Syntax of gramgraph ref * tables; + + + +(*** compile syntax ***) + +(* ggr_to_xgram *) + +fun ggr_to_xgram ggr = + let + fun flatGG ggr (xg, v) = + if ggr mem v then (xg, v) else flatGG' ggr (xg, ggr :: v) + and flatGG' (ref EmptyGG) xgv = xgv + | flatGG' (ref (ExtGG (ggr, ext))) xgv = + let + val (xg', v') = flatGG ggr xgv + in + (Extension.extend xg' ext, v') + end + | flatGG' (ref (MergeGG (ggr1, ggr2))) xgv = + flatGG ggr1 (flatGG ggr2 xgv); + in + fst (flatGG ggr (Extension.empty, [])) + end; + + +(* mk_ruletab *) + +fun mk_ruletab rules = + let + fun add_rule (r, tab) = + let + val a = head_of_rule r + in + case lookup tab a of + None => Symtab.update ((a, [r]), tab) + | Some rs => Symtab.update ((a, r :: rs), tab) + end; + in + Symtab.balance (foldr add_rule (rules, Symtab.null)) + end; + + +(* make_syntax *) + +fun make_syntax ggr = + let + fun mk_const_tab cs = + Symtab.balance + (Symtab.st_of_alist ((map (fn c => (c, ())) cs), Symtab.null)); + + fun mk_trtab alst name = + Symtab.balance (Symtab.st_of_alist (alst, Symtab.null)) + handle Symtab.DUPLICATE s => error ("More than one " ^ name ^ " for " ^ quote s); + + fun mk_proc (Some f) = f + | mk_proc None = I; + + fun all_strings (opl: string prod list): string list = + flat (map (fn Prod (_, syn, _, _) => terminals syn) opl); + + fun str_to_tok (opl: string prod list, lex: Lexicon): Token prod list = + map + (fn Prod (t, syn, s, pa) => + Prod (t, translate (hd o tokenize lex) syn, s, pa)) + opl; + + fun xgram_to_tables (XGram xgram) = + let + val {roots, prods, consts, parse_ast_translation, parse_preproc, parse_rules, + parse_postproc, parse_translation, print_translation, print_preproc, + print_rules, print_postproc, print_ast_translation} = xgram; + val lexicon = mk_lexicon (all_strings prods); + in + Tab { + gram = Parser.compile_xgram (roots, str_to_tok (prods, lexicon)), + lexicon = lexicon, + const_tab = mk_const_tab consts, + parse_ast_trtab = mk_trtab parse_ast_translation "parse ast translation", + parse_preproc = mk_proc parse_preproc, + parse_ruletab = mk_ruletab parse_rules, + parse_postproc = mk_proc parse_postproc, + parse_trtab = mk_trtab parse_translation "parse translation", + print_trtab = mk_trtab print_translation "print translation", + print_preproc = mk_proc print_preproc, + print_ruletab = mk_ruletab print_rules, + print_postproc = mk_proc print_postproc, + print_tab = mk_print_tab prods + (mk_trtab print_ast_translation "print ast translation")} + end; + in + Syntax (ggr, xgram_to_tables (ggr_to_xgram ggr)) + end; + + +(* add_synrules *) + +fun add_synrules (Tab tabs) (SynRules rules) = + let + val {gram, lexicon, const_tab, parse_ast_trtab, parse_preproc, parse_ruletab, + parse_postproc, parse_trtab, print_trtab, print_preproc, print_ruletab, + print_postproc, print_tab} = tabs; + val {parse_rules, print_rules} = rules; + val parse_rs = flat (map snd (Symtab.alist_of parse_ruletab)) @ parse_rules; + val print_rs = flat (map snd (Symtab.alist_of print_ruletab)) @ print_rules; + in + Tab { + gram = gram, lexicon = lexicon, const_tab = const_tab, + parse_ast_trtab = parse_ast_trtab, + parse_preproc = parse_preproc, + parse_ruletab = mk_ruletab parse_rs, + parse_postproc = parse_postproc, + parse_trtab = parse_trtab, + print_trtab = print_trtab, + print_preproc = print_preproc, + print_ruletab = mk_ruletab print_rs, + print_postproc = print_postproc, + print_tab = print_tab} + end; + + + +(*** inspect syntax ***) + +(* print_syntax_old *) (* FIXME remove *) + +fun print_syntax_old (Syntax (_, Tab {gram, lexicon, ...})) = + Parser.print_gram (gram, lexicon); + + + +fun xgram_of (Syntax (ggr, _)) = ggr_to_xgram ggr; + +fun string_of_big_list name prts = + Pretty.string_of (Pretty.blk (2, + separate Pretty.fbrk (Pretty.str name :: prts))); + + +(* print_gram *) (* FIXME check *) + +fun prt_gram (XGram {roots, prods, ...}) = + let + fun pretty_name name = [Pretty.str (name ^ " ="), Pretty.brk 1]; + + fun pretty_xsymbs (Terminal s :: xs) = + Pretty.str (quote s) :: Pretty.brk 1 :: pretty_xsymbs xs + | pretty_xsymbs (Nonterminal (s, p) :: xs) = + (if is_terminal s then Pretty.str s + else Pretty.str (s ^ "[" ^ string_of_int p ^ "]")) + :: Pretty.brk 1 :: pretty_xsymbs xs + | pretty_xsymbs (_ :: xs) = pretty_xsymbs xs + | pretty_xsymbs [] = []; + + fun pretty_const "" = [Pretty.brk 1] + | pretty_const c = [Pretty.str (" => " ^ quote c), Pretty.brk 1]; + + fun pretty_pri p = [Pretty.str ("(" ^ string_of_int p ^ ")")]; + + fun pretty_prod (Prod (name, xsymbs, const, pri)) = + Pretty.blk (2, pretty_name name @ pretty_xsymbs xsymbs @ + pretty_const const @ pretty_pri pri); + in + writeln (Pretty.string_of (Pretty.blk (2, + separate (Pretty.brk 1) (map Pretty.str ("roots:" :: roots))))); + writeln (string_of_big_list "prods:" (map pretty_prod prods)) + end; + + +val print_gram = prt_gram o xgram_of; + + +(* print_trans *) (* FIXME check *) + +fun prt_trans (XGram xgram) = + let + fun string_of_strings name strs = + Pretty.string_of (Pretty.blk (2, + separate (Pretty.brk 1) (map Pretty.str (name :: map quote strs)))); + + fun string_of_trs name trs = string_of_strings name (map fst trs); + + fun string_of_proc name proc = + Pretty.string_of (Pretty.blk (2, [Pretty.str name, Pretty.brk 1, + Pretty.str (if is_none proc then "None" else "Some fn")])); + + fun string_of_rules name rules = + string_of_big_list name (map pretty_rule rules); + + + val {consts, parse_ast_translation, parse_preproc, parse_rules, + parse_postproc, parse_translation, print_translation, print_preproc, + print_rules, print_postproc, print_ast_translation, ...} = xgram; + in + writeln (string_of_strings "consts:" consts); + writeln (string_of_trs "parse_ast_translation:" parse_ast_translation); + writeln (string_of_proc "parse_preproc:" parse_preproc); + writeln (string_of_rules "parse_rules:" parse_rules); + writeln (string_of_proc "parse_postproc:" parse_postproc); + writeln (string_of_trs "parse_translation:" parse_translation); + writeln (string_of_trs "print_translation:" print_translation); + writeln (string_of_proc "print_preproc:" print_preproc); + writeln (string_of_rules "print_rules:" print_rules); + writeln (string_of_proc "print_postproc:" print_postproc); + writeln (string_of_trs "print_ast_translation:" print_ast_translation) + end; + + +val print_trans = prt_trans o xgram_of; + + +(* print_syntax *) + +fun print_syntax syn = + let + val xgram = xgram_of syn; + in + prt_gram xgram; prt_trans xgram + end; + + + +(*** parsing and printing ***) + +(* read_ast *) + +fun read_ast syn (root, s) = + let + val Syntax (_, Tab {gram, lexicon, parse_ast_trtab, ...}) = syn; + val root = if Parser.parsable (gram, root) then root else error ("Unparsable root " ^ root) (* Extension.logic *); (* FIXME *) + fun syn_err toks = + error ("Syntax error at\n" ^ space_implode " " (map token_to_string toks)); + in + Parser.ParseTree.pt_to_ast (lookup parse_ast_trtab) + (Parser.parse (gram, root, tokenize lexicon s)) + handle Parser.SYNTAX_ERR toks => syn_err toks + end; + + +(* norm_ast *) + +fun norm_ast ruletab ast = + let + fun get_rules a = + (case lookup ruletab a of + Some rules => rules + | None => []); + in + normalize (if Symtab.is_null ruletab then None else Some get_rules) ast + end; + + + +(** read **) + +fun read (syn as Syntax (_, Tab tabs)) ty s = + let + val {parse_preproc, parse_ruletab, parse_postproc, parse_trtab, ...} = tabs; + val ast = read_ast syn (typ_to_nt ty, s); + in + ast_to_term (lookup parse_trtab) + (parse_postproc (norm_ast parse_ruletab (parse_preproc ast))) + end; + + + +(** pprint_ast **) + +val pprint_ast = Pretty.pprint o pretty_ast; + + + +(** pretty term, typ **) + +fun pretty_t t_to_ast pretty_t (syn as Syntax (_, Tab tabs)) t = + let + val {print_trtab, print_preproc, print_ruletab, print_postproc, print_tab, ...} = tabs; + val ast = t_to_ast (lookup print_trtab) t; + in + pretty_t print_tab + (print_postproc (norm_ast print_ruletab (print_preproc ast))) + end; + +val pretty_term = pretty_t term_to_ast pretty_term_ast; + +val pretty_typ = pretty_t typ_to_ast pretty_typ_ast; + +fun string_of_term syn t = Pretty.string_of (pretty_term syn t); + +fun string_of_typ syn ty = Pretty.string_of (pretty_typ syn ty); + + + +(*** build syntax ***) + +(* type_syn *) + +val type_syn = make_syntax (ref (ExtGG (ref EmptyGG, (type_ext, empty_synrules)))); + + +(* extend *) (* FIXME check *) + +fun extend (old_syn as Syntax (ggr, _), def_sort) (roots, xconsts, sext) = + let + fun read_typ s = typ_of_term def_sort (read old_syn typeT s); + val ext = ext_of_sext roots xconsts read_typ sext; + + fun right_rule (xpat1 |-> xpat2) = Some (xpat1, xpat2) + | right_rule (xpat1 <-| xpat2) = None + | right_rule (xpat1 <-> xpat2) = Some (xpat1, xpat2); + + fun left_rule (xpat1 |-> xpat2) = None + | left_rule (xpat1 <-| xpat2) = Some (xpat2, xpat1) + | left_rule (xpat1 <-> xpat2) = Some (xpat2, xpat1); + + val (tmp_syn as Syntax (_, tmp_tabs)) = + make_syntax (ref (ExtGG (ggr, (ext, empty_synrules)))); + val Syntax (_, Tab {const_tab, ...}) = tmp_syn; + + fun constantify (ast as (Constant _)) = ast + | constantify (ast as (Variable x)) = + if is_some (lookup const_tab x) then Constant x else ast + | constantify (Appl asts) = Appl (map constantify asts); + + fun read_pat (r_s as (root, s)) = + constantify ((read_ast tmp_syn r_s) + handle ERROR => error ("The error above occurred in " ^ quote s)); + + fun read_rule (xrule as ((_, lhs_src), (_, rhs_src))) = + let + val rule as (lhs, rhs) = (pairself read_pat) xrule; + in + case rule_error rule of + Some msg => + error ("Error in syntax translation rule: " ^ msg ^ + "\nexternal: " ^ quote lhs_src ^ " -> " ^ quote rhs_src ^ + "\ninternal: " ^ str_of_ast lhs ^ " -> " ^ str_of_ast rhs) + | None => rule + end; + + val xrules = xrules_of sext; + val new_rules = + SynRules { + parse_rules = map read_rule (mapfilter right_rule xrules), + print_rules = map read_rule (mapfilter left_rule xrules)}; + in + Syntax (ref (ExtGG (ggr, (ext, new_rules))), add_synrules tmp_tabs new_rules) + end; + + +(* merge *) + +fun merge (Syntax (ggr1, _), Syntax (ggr2, _)) = + make_syntax (ref (MergeGG (ggr1, ggr2))); + + +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/type_ext.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/type_ext.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,197 @@ +(* Title: Pure/Syntax/type_ext + ID: $Id$ + Author: Tobias Nipkow and Markus Wenzel, TU Muenchen + +The concrete syntax of types (used to bootstrap Pure) +*) + +signature TYPE_EXT0 = +sig + val typ_of_term: (indexname -> sort) -> term -> typ +end; + +signature TYPE_EXT = +sig + include TYPE_EXT0 + structure Extension: EXTENSION + local open Extension Extension.XGram.Ast in + val term_of_typ: bool -> typ -> term + val tappl_ast_tr': ast * ast list -> ast + val type_ext: ext + end +end; + +functor TypeExtFun(structure Extension: EXTENSION and Lexicon: LEXICON): TYPE_EXT = +struct + +structure Extension = Extension; +open Extension Extension.XGram.Ast; + + +(** typ_of_term **) (* FIXME check *) + +fun typ_of_term def_sort t = + let + fun sort_err (xi as (x, i)) = + error ("typ_of_term: inconsistent sort constraints for type variable " ^ + (if i < 0 then x else Lexicon.string_of_vname xi)); + + fun is_tfree name = + (case explode name of + "'" :: _ :: _ => true + | _ :: _ => false + | _ => error ("typ_of_term: bad var name " ^ quote name)); + + fun put_sort scs xi s = + (case assoc (scs, xi) of + None => (xi, s) :: scs + | Some s' => if s = s' then scs else sort_err xi); + + fun classes (Const (s, _)) = [s] + | classes (Free (s, _)) = [s] + | classes (Const ("_sortcons", _) $ Const (s, _) $ cls) = s :: classes cls + | classes (Const ("_sortcons", _) $ Free (s, _) $ cls) = s :: classes cls + | classes tm = raise (TERM ("typ_of_term: bad encoding of classes", [tm])); + + fun sort (Const ("_emptysort", _)) = [] + | sort (Const (s, _)) = [s] + | sort (Free (s, _)) = [s] + | sort (Const ("_classes", _) $ cls) = classes cls + | sort tm = raise (TERM ("typ_of_term: bad encoding of sort", [tm])); + + fun typ (Free (x, _), scs) = + (if is_tfree x then TFree (x, []) else Type (x, []), scs) + | typ (Var (xi, _), scs) = (TVar (xi, []), scs) + | typ (Const ("_ofsort", _) $ Free (x, _) $ st, scs) = + (TFree (x, []), put_sort scs (x, ~1) (sort st)) + | typ (Const ("_ofsort", _) $ Var (xi, _) $ st, scs) = + (TVar (xi, []), put_sort scs xi (sort st)) + | typ (Const (a, _), scs) = (Type (a, []), scs) + | typ (tm as (_ $ _), scs) = + let + val (t, ts) = strip_comb tm; + val a = + case t of + Const (x, _) => x + | Free (x, _) => x + | _ => raise (TERM ("typ_of_term: bad type application", [tm])); + val (tys, scs') = typs (ts, scs); + in + (Type (a, tys), scs') + end + | typ (tm, _) = raise (TERM ("typ_of_term: bad encoding of typ", [tm])) + and typs (t :: ts, scs) = + let + val (ty, scs') = typ (t, scs); + val (tys, scs'') = typs (ts, scs'); + in (ty :: tys, scs'') end + | typs ([], scs) = ([], scs); + + + val (ty, scs) = typ (t, []); + + fun get_sort xi = + (case assoc (scs, xi) of + None => def_sort xi + | Some s => s); + + fun add_sorts (Type (a, tys)) = Type (a, map add_sorts tys) + | add_sorts (TVar (xi, _)) = TVar (xi, get_sort xi) + | add_sorts (TFree (x, _)) = TFree (x, get_sort (x, ~1)); + in + add_sorts ty + end; + + + +(** term_of_typ **) + +fun term_of_typ show_sorts ty = + let + fun const x = Const (x, dummyT); + + fun classes [] = raise Match + | classes [s] = const s + | classes (s :: ss) = const "_sortcons" $ const s $ classes ss; + + fun sort [] = const "_emptysort" + | sort [s] = const s + | sort ss = const "_classes" $ classes ss; + + fun of_sort t ss = + if show_sorts then const "_ofsort" $ t $ sort ss else t; + + fun typ (Type (a, tys)) = list_comb (const a, map typ tys) + | typ (TFree (x, ss)) = of_sort (Free (x, dummyT)) ss + | typ (TVar (xi, ss)) = of_sort (Var (xi, dummyT)) ss; + in + typ ty + end; + + + +(** the type syntax **) + +(* parse translations *) + +fun tappl_ast_tr (*"_tappl"*) [types, f] = Appl (f :: unfold_ast "_types" types) + | tappl_ast_tr (*"_tappl"*) asts = raise_ast "tappl_ast_tr" asts; + +fun bracket_ast_tr (*"_bracket"*) [dom, cod] = + fold_ast_p "fun" (unfold_ast "_types" dom, cod) + | bracket_ast_tr (*"_bracket"*) asts = raise_ast "bracket_ast_tr" asts; + + +(* print translations *) + +fun tappl_ast_tr' (f, []) = raise_ast "tappl_ast_tr'" [f] + | tappl_ast_tr' (f, [ty]) = Appl [Constant "_tapp", ty, f] + | tappl_ast_tr' (f, tys) = Appl [Constant "_tappl", fold_ast "_types" tys, f]; + +fun fun_ast_tr' (*"fun"*) asts = + (case unfold_ast_p "fun" (Appl (Constant "fun" :: asts)) of + (dom as (_ :: _ :: _), cod) + => Appl [Constant "_bracket", fold_ast "_types" dom, cod] + | _ => raise Match); + + +(* type_ext *) + +val sortT = Type ("sort", []); +val classesT = Type ("classes", []); +val typesT = Type ("types", []); + +val type_ext = + Ext { + roots = [logic, "type"], + mfix = [ + Mfix ("_", idT --> sortT, "", [], max_pri), + Mfix ("{}", sortT, "_emptysort", [], max_pri), + Mfix ("{_}", classesT --> sortT, "_classes", [], max_pri), + Mfix ("_", idT --> classesT, "", [], max_pri), + Mfix ("_,_", [idT, classesT] ---> classesT, "_sortcons", [], max_pri), + Mfix ("_", tfreeT --> typeT, "", [], max_pri), + Mfix ("_", tvarT --> typeT, "", [], max_pri), + Mfix ("_", idT --> typeT, "", [], max_pri), + Mfix ("_::_", [tfreeT, sortT] ---> typeT, "_ofsort", [max_pri, 0], max_pri), + Mfix ("_::_", [tvarT, sortT] ---> typeT, "_ofsort", [max_pri, 0], max_pri), + Mfix ("_ _", [typeT, idT] ---> typeT, "_tapp", [max_pri, 0], max_pri), + Mfix ("((1'(_'))_)", [typesT, idT] ---> typeT, "_tappl", [], max_pri), + Mfix ("_", typeT --> typesT, "", [], max_pri), + Mfix ("_,/ _", [typeT, typesT] ---> typesT, "_types", [], max_pri), + Mfix ("(_/ => _)", [typeT, typeT] ---> typeT, "fun", [1, 0], 0), + Mfix ("([_]/ => _)", [typesT, typeT] ---> typeT, "_bracket", [0, 0], 0)], + extra_consts = ["fun"], + parse_ast_translation = [("_tapp", tappl_ast_tr), ("_tappl", tappl_ast_tr), + ("_bracket", bracket_ast_tr)], + parse_preproc = None, + parse_postproc = None, + parse_translation = [], + print_translation = [], + print_preproc = None, + print_postproc = None, + print_ast_translation = [("fun", fun_ast_tr')]}; + + +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/Syntax/xgram.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Syntax/xgram.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,117 @@ +(* Title: Pure/Syntax/xgram + ID: $Id$ + Author: Tobias Nipkow and Markus Wenzel, TU Muenchen + +External Grammar Representation + +Changes: + XGRAM: added Ast, changed XGram, + renamed (XSymb, Prod, XGram) to (xsymb, prod, xgram) + removed Symtab +TODO: + prod, xsymb, xgram: 'a -> string (?) +*) + +signature XGRAM = +sig + structure Ast: AST + datatype 'a xsymb = + Terminal of 'a | + Nonterminal of string * int | + Space of string | + Bg of int | Brk of int | En + datatype 'a prod = Prod of string * ('a xsymb list) * string * int + local open Ast in + datatype 'a xgram = + XGram of { + roots: string list, + prods: 'a prod list, + consts: string list, + parse_ast_translation: (string * (ast list -> ast)) list, + parse_preproc: (ast -> ast) option, + parse_rules: (ast * ast) list, + parse_postproc: (ast -> ast) option, + parse_translation: (string * (term list -> term)) list, + print_translation: (string * (term list -> term)) list, + print_preproc: (ast -> ast) option, + print_rules: (ast * ast) list, + print_postproc: (ast -> ast) option, + print_ast_translation: (string * (ast list -> ast)) list} + end + val copy_pri: int + val terminals: 'a xsymb list -> 'a list + val nonterminals: 'a xsymb list -> string list + val translate: ('a -> 'b) -> 'a xsymb list -> 'b xsymb list +end; + +functor XGramFun(Ast: AST)(*: XGRAM*) = (* FIXME *) +struct + +structure Ast = Ast; +open Ast; + +(** datatype 'a xgram **) + +(* Terminal a: terminal a + Nonterminal (s, p): nonterminal named s, require priority >= p + Space s: some white space for printing + Bg, Brk, En: see resp. routines in Pretty *) + +datatype 'a xsymb = + Terminal of 'a | + Nonterminal of string * int | + Space of string | + Bg of int | Brk of int | En; + + +(* Prod (lhs, rhs, opn, p): + lhs: name of nonterminal on the lhs of the production + rhs: list of symbols on the rhs of the production + opn: name of the corresponding Isabelle Const + p: priority of this production, 0 <= p <= max_pri *) + +datatype 'a prod = Prod of string * ('a xsymb list) * string * int; + + +datatype 'a xgram = + XGram of { + roots: string list, + prods: 'a prod list, + consts: string list, + parse_ast_translation: (string * (ast list -> ast)) list, + parse_preproc: (ast -> ast) option, + parse_rules: (ast * ast) list, + parse_postproc: (ast -> ast) option, + parse_translation: (string * (term list -> term)) list, + print_translation: (string * (term list -> term)) list, + print_preproc: (ast -> ast) option, + print_rules: (ast * ast) list, + print_postproc: (ast -> ast) option, + print_ast_translation: (string * (ast list -> ast)) list}; + + + +(** misc stuff **) + +val copy_pri = ~1; (* must be < all legal priorities, i.e. 0 *) + +fun terminals [] = [] + | terminals (Terminal s :: sl) = s :: terminals sl + | terminals (_ :: sl) = terminals sl; + +fun nonterminals [] = [] + | nonterminals (Nonterminal (s, _) :: sl) = s :: nonterminals sl + | nonterminals (_ :: sl) = nonterminals sl; + +fun translate trfn = + map + (fn Terminal t => Terminal (trfn t) + | Nonterminal s => Nonterminal s + | Space s => Space s + | Bg i => Bg i + | Brk i => Brk i + | En => En); + + +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/Thy/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Thy/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,31 @@ +(* Title: Pure/Thy/ROOT + ID: $Id$ + Author: Sonia Mahjoub / Tobias Nipkow + Copyright 1992 TU Muenchen + +This file builds the theory parser. +It assumes the standard Isabelle environment. +*) + +use "scan.ML"; +use "parse.ML"; +use "syntax.ML"; +use "read.ML"; + + +structure Keyword = + struct + val alphas = ["classes", "default", "arities", "types", + "consts", "rules", "end", "rules", "mixfix", + "infixr", "infixl", "binder", "translations"] + + val symbols = [",", "<", "{", "}", "(", ")", "(*", "*)", + "[", "]", "::", "=", "+", "==", "=>", "<="] + end; + +structure Lex = LexicalFUN (Keyword); +structure Parse = ParseFUN (Lex); +structure ThySyn = ThySynFUN (Parse); +structure Readthy = ReadthyFUN (ThySyn); + +open Readthy; diff -r 000000000000 -r a5a9c433f639 src/Pure/Thy/parse.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Thy/parse.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,134 @@ +(* Title: Pure/Thy/parse + ID: $Id$ + Author: Sonia Mahjoub / Tobias Nipkow + Copyright 1992 TU Muenchen + +The parser combinators. +Adapted from Larry Paulson's ML for the Working Programmer. +*) + +(*Global infix declarations for the parsing primitives*) +infix 5 -- --$$ $$--; +infix 3 >>; +infix 0 ||; + + +signature PARSE = +sig + +type token +val $$ : string -> + (token * int) list -> string * (token * int)list +val id : (token * int) list -> string * (token * int)list +val nat : (token * int) list -> string * (token * int)list +val stg : (token * int) list -> string * (token * int)list +val txt : (token * int) list -> string * (token * int)list +val >> : ('a -> 'b * 'c) * ('b -> 'd) -> 'a -> 'd * 'c +val || : ('a -> 'b) * ('a -> 'b) -> 'a -> 'b +val -- : ('a -> 'b * 'c) * ('c -> 'd * 'e) -> + 'a -> ('b * 'd) * 'e +val $$-- : string * ((token * int)list -> 'b * 'c) -> + (token * int) list -> 'b * 'c +val --$$ : ( 'a -> 'b * (token * int)list ) * string -> + 'a -> 'b * (token * int)list +val !! : ((token * int) list -> 'a * (token * int) list ) + -> (token * int) list -> 'a * (token * int) list +val empty : 'a -> 'b list * 'a +val repeat : ('a -> 'b * 'a) -> 'a -> 'b list * 'a +val repeat1: ('a -> 'b * 'a) -> 'a -> 'b list * 'a +val list_of: ((token * int)list -> 'b * (token * int)list ) -> + (token * int)list -> 'b list * (token * int)list +val list_of1: ((token * int)list -> 'b * (token * int)list ) -> + (token * int)list -> 'b list * (token * int)list +val reader : ((token * int) list -> 'a * ('b * int) list ) + -> string list -> 'a + +end; + + + +functor ParseFUN (Lex: LEXICAL): PARSE = +struct + +type token = Lex.token; + +datatype synerr = Line of string * string * int | EOF of string; + +exception SynError of synerr; + +fun synerr(Line(s1, s2, n)) = + error("Syntax error on line " ^ (string_of_int n) ^ ": " ^ s1 ^ + " expected and " ^ s2 ^ " was found") + | synerr(EOF(s)) = error("Syntax error on last line: " ^ s ^ + " expected and end-of-file was found"); + +fun string_of_token (Lex.Key b) = b + | string_of_token (Lex.Id b) = b + | string_of_token (Lex.Nat b) = b + | string_of_token (Lex.Txt b) = b + | string_of_token (Lex.Stg b) = b; + +fun line_err x = raise SynError(Line x); +fun eof_err s = raise SynError(EOF s); + +fun $$a ((Lex.Key b,n) :: toks) = + if a = b then (a,toks) else line_err(a,b,n) + | $$a ((t,n) :: toks) = line_err (a,string_of_token t, n) + | $$a _ = eof_err a; + + +fun id ((Lex.Id b,n) :: toks) = (b, toks) + | id ((t,n) :: toks) = line_err ("identifier", string_of_token t, n) + | id _ = eof_err "identifier"; + + +fun nat ((Lex.Nat b,n) :: toks) = (b, toks) + | nat ((t,n) :: toks) = + line_err ("natural number", string_of_token t, n) + | nat _ = eof_err "natural number"; + + +fun stg ((Lex.Stg b,n) :: toks) = (b, toks) + | stg ((t,n) :: toks) = line_err("string", string_of_token t, n) + | stg _ = eof_err"string"; + + +fun txt ((Lex.Txt b,n) :: toks) = (b, toks) + | txt ((t,n) :: toks) = line_err ("ML text", string_of_token t, n) + | txt _ = eof_err "ML text"; + + +fun ( ph >> f) toks = let val (x, toks2) = ph toks in (f x, toks2) end; + +fun (ph1 || ph2) toks = ph1 toks handle SynError _ => ph2 toks; + + +fun (ph1 -- ph2) toks = + let val (x, toks2) = ph1 toks + val (y, toks3) = ph2 toks2 + in ((x,y), toks3) end; + +fun (a $$-- ph) = $$a -- ph >> #2; + +fun (ph --$$ a) = ph -- $$a >> #1; + +fun !! ph toks = ph toks handle SynError s => synerr s; + +fun empty toks = ([], toks); + +fun repeat ph toks = ( ph -- repeat ph >> (op::) + || empty ) toks; + +fun repeat1 ph = ph -- repeat ph >> (op::); + +fun list_of1 ph = ph -- repeat("," $$-- !! ph) >> (op::); +fun list_of ph = list_of1 ph || empty; + +fun reader ph a = + ( case ph (Lex.scan a) of + (x, []) => x + | (_,(_, n)::_) => error + ("Syntax error on line " ^ (string_of_int n) ^ ": " ^ + "Extra characters in phrase") + ); +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/Thy/read.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Thy/read.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,61 @@ +(* Title: Pure/Thy/read + ID: $Id$ + Author: Sonia Mahjoub / Tobias Nipkow / L C Paulson + Copyright 1992 TU Muenchen + +Reading and writing the theory definition files. + +For theory XXX, the input file is called XXX.thy + the output file is called .XXX.thy.ML + and it then tries to read XXX.ML +*) + +signature READTHY = +sig + val split_filename : string -> string * string + val time_use_thy : string -> unit + val use_thy : string -> unit +end; + + +functor ReadthyFUN (ThySyn: THYSYN): READTHY = +struct + +(*Convert Unix filename of the form path/file to "path/" and "file" ; + if filename contains no slash, then it returns "" and "file" *) +fun split_filename name = + let val (file,path) = take_prefix (apr(op<>,"/")) (rev (explode name)) + in (implode(rev path), implode(rev file)) end; + +(*create name of the output ML file for the theory*) +fun out_name thy = "." ^ thy ^ ".thy.ML"; + +fun read_thy path thy = + let val instream = open_in (path ^ thy ^ ".thy") + val outstream = open_out (path ^ out_name thy) + in output (outstream, ThySyn.read (explode(input(instream, 999999)))); + close_out outstream; + close_in instream + end; + +fun file_exists file = + let val instream = open_in file in close_in instream; true end + handle Io _ => false; + +(*Use the file if it exists*) +fun try_use file = + if file_exists file then use file else (); + +(*Read and convert the theory to a .XXX.thy.ML file and also try to read XXX.ML*) +fun use_thy name = + let val (path,thy) = split_filename name + in read_thy path thy; + use (path ^ out_name thy); + try_use (path ^ thy ^ ".ML") + end; + +fun time_use_thy tname = timeit(fn()=> + (writeln("\n**** Starting Theory " ^ tname ^ " ****"); use_thy tname; + writeln("\n**** Finished Theory " ^ tname ^ " ****"))); + +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/Thy/scan.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Thy/scan.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,164 @@ +(* Title: Pure/Thy/scan + ID: $Id$ + Author: Sonia Mahjoub / Tobias Nipkow + Copyright 1992 TU Muenchen + +The scanner. Adapted from Larry Paulson's ML for the Working Programmer. +*) + +signature LEXICAL = +sig + + +datatype token = Id of string + | Key of string + | Nat of string + | Stg of string + | Txt of string + +val scan : string list -> (token * int) list +end; + +signature KEYWORD = +sig +val alphas : string list +val symbols : string list +end; + + +functor LexicalFUN (Keyword: KEYWORD): LEXICAL = + +struct + + + +datatype token = Id of string + | Key of string + | Nat of string + | Stg of string + | Txt of string; + + +fun lexerr(n,s) = + error("Lexical error on line " ^ (string_of_int n) ^ ": " ^ s); + +val specials = explode"!{}@#$%^&*()+-=[]:\";',./?`_~<>|\\"; + +fun is_symbol c = "_" = c orelse "'" = c; + +fun alphanum (id, c::cs) = + if is_letter c orelse is_digit c orelse is_symbol c + then alphanum (id ^ c , cs) + else (id , c :: cs) + | alphanum (id ,[]) = (id ,[]); + +fun numeric (nat, c::cs) = + if is_digit c + then numeric (nat^c, cs) + else (nat, c::cs) + | numeric (nat, []) = (nat,[]); + +fun tokenof (a, n) = + if a mem Keyword.alphas + then (Key a, n) + else (Id a, n); + +fun symbolic (sy, c::cs) = + if (sy mem Keyword.symbols) andalso + not((sy^c) mem Keyword.symbols) + orelse not (c mem specials) + then (sy, c::cs) + else symbolic(sy^c, cs) + | symbolic (sy, []) = (sy, []); + + +fun stringerr(n) = lexerr(n, "No matching quote found on this line"); + +fun is_control_chr ([],_,n) = stringerr(n) + | is_control_chr (c::cs,s,n) = + let val m = ord c + in if (m >= 64 andalso m <= 95) + then (cs, s^c, n) + else stringerr(n) + end; + +fun is_3_dgt (c1::c2::cs, c,n) = + let val s = c^c1^c2 + in if (s >= "000" andalso s <= "255") + then (cs, s) + else stringerr(n) + end + | is_3_dgt (_,_,n) = stringerr(n); + +fun is_imprt_seq ([],_,n) = stringerr(n) + | is_imprt_seq ((c::cs),s,n) = + if c = "\\" then (cs,s^"\\",n) + else if c = "\n" + then is_imprt_seq (cs,s^"\n",n+1) + else if (c = "\t") orelse (c = " ") + then is_imprt_seq (cs,s^c,n) + else stringerr(n); + +fun is_escape_seq ([],_,n) = stringerr(n) + | is_escape_seq ((c::cs),s,n) = + if c = "\\" then (cs,s^"\\",n) + else if c = "\n" + then is_imprt_seq (cs,s^"\n",n+1) + else if (c = "\t") orelse (c = " ") + then is_imprt_seq (cs,s^c,n) + else if c = "^" + then is_control_chr (cs,s^"^",n) + else if ("0" <= c andalso c <= "2") + then let val (cs',s') = + is_3_dgt(cs,c,n) + in (cs',s^s',n) + end + else stringerr(n); + + +fun string (_,[],_,n) = stringerr(n) + | string (toks, c::cs, s, n) = + if c = "\"" then ((Stg s, n)::toks , cs, n) + else if c = "\\" + then let val (cs',s',n') = + is_escape_seq (cs, s^"\\",n) + in string (toks,cs',s',n') end + else string (toks,cs,s^c,n); + + +fun comment ((c1::c2::cs), n) = + if c1 = "*" andalso c2 = ")" then (cs,n) else + if c1 = "\n" then comment((c2::cs), n+1) + else comment((c2::cs), n) + | comment (_, n) = lexerr(n, "Missing end of comment"); + +fun scanning (toks , [], n) = rev toks + | scanning (toks , c :: cs, n) = + if is_letter c + then let val (id , cs2) = alphanum (c , cs) + in if id = "ML" + then let val text = implode cs2 + in scanning ((Txt text,n) :: toks , [], n) + end + else scanning (tokenof(id,n) :: toks , cs2, n) + end + else if is_digit c + then let val (nat , cs2) = numeric(c , cs) + in scanning ((Nat nat,n) :: toks , cs2, n) end + else if c mem specials + then if c = "\"" + then let val (toks', cs', n') = string (toks, cs, "", n) + in scanning (toks', cs', n') end + else let val (sy , cs2) = symbolic (c , cs) + in if sy = "(*" + then let val (cs3,n2) = comment(cs2,n) + in scanning (toks , cs3, n2) end + else scanning ((Key sy,n) :: toks, cs2, n) + end + else if c = "\n" then scanning (toks, cs, n+1) + else if c = " " orelse c = "\t" then scanning (toks , cs, n) + else lexerr(n,"Illegal character " ^ c); + +fun scan a = scanning ([] , a, 1); + +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/Thy/syntax.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/Thy/syntax.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,359 @@ +(* Title: Pure/Thy/syntax + ID: $Id$ + Author: Sonia Mahjoub and Tobias Nipkow and Markus Wenzel + Copyright 1992 TU Muenchen + +Definition of theory syntax together with translation to ML code. +*) + +signature THYSYN = + sig + val read: string list -> string + end; + + + +functor ThySynFUN (Parse: PARSE): THYSYN = +struct + + +(*-------------- OBJECT TO STRING TRANSLATION ---------------*) + +fun string a = "\"" ^ a ^ "\""; + +fun parent s = "(" ^ s ^ ")"; + +fun pair(a,b) = parent(a ^ ", " ^ b); + +fun pair_string(a,b) = pair(string a,string b); + +fun pair_string2(a,b) = pair(a,string b); + +fun bracket s = "[" ^ s ^ "]"; + +val comma = space_implode ", "; + +val bracket_comma = bracket o comma; + +val big_bracket_comma = bracket o space_implode ",\n"; + +fun big_bracket_comma_ind ind strs = bracket (space_implode (",\n" ^ ind) strs); + +val bracket_comma_string = bracket_comma o (map string); + + +(*------------------- ABSTRACT SYNTAX FUNCTIONS ---------------------*) + +datatype mixfix = Mixfix of string * string * string + | Delimfix of string + | Infixl of string + | Infixr of string + | Binder of string * string + | TInfixl of string + | TInfixr of string; + + +datatype pfix_or_mfix = Pref of string | Mixf of string; + +fun pm_proj(Pref s) = s + | pm_proj(Mixf s) = s; + +fun split_decls l = + let val (p,m) = partition (fn Pref _ => true | _ => false) l; + in (big_bracket_comma_ind " " (map pm_proj p), map pm_proj m) end; + +fun delim_mix (s, None) = Delimfix(s) + | delim_mix (s, Some(l,n)) = Mixfix(s,l,n); + +fun mixfix (sy,c,ty,l,n) = "Mixfix(" ^ comma[string sy, c, ty, l, n] ^ ")"; + +fun infixrl(ty,c,n) = parent(comma[ty,c,n]); + +fun binder(sy, c, ty, n) = "Binder(" ^ comma[string sy, c, ty, "0", n] ^ ")"; + +fun delimfix (sy,c,ty) = "Delimfix(" ^ comma[string sy, c, ty] ^ ")"; + +fun tinfixrl (ty, n) = "(" ^ comma[ty, ty, n] ^ ")"; + +fun mk_mfix((c,ty),mfix) = + let val cs = string c and tys = string ty + in case mfix of + Mixfix(sy, l, n) => mixfix (sy, tys, cs, l, n) + | Infixr(n) => "Infixr" ^ infixrl(cs, tys, n) + | Infixl(n) => "Infixl" ^ infixrl(cs, tys, n) + | Binder(sy,n) => binder(sy,tys,cs,n) + | TInfixl(n) => "TInfixl" ^ tinfixrl(cs, n) + | TInfixr(n) => "TInfixr" ^ tinfixrl(cs, n) + | Delimfix(sy) => delimfix(sy, tys, cs) + end; + + +fun mk_mixfix((cs,ty), None) = + [Pref(pair(bracket_comma_string cs, string ty))] + | mk_mixfix((c::cs,ty), Some(mfix)) = + Mixf(mk_mfix((c,ty),mfix)) :: mk_mixfix((cs,ty), Some(mfix)) + | mk_mixfix(([],_),_) = []; + +fun mk_type_decl((ts, n), None) = [Pref(pair(bracket_comma_string ts, n))] + | mk_type_decl((t::ts, n), Some(tinfix)) = + [Pref(pair(bracket(string t), n)), Mixf(mk_mfix((t,n), tinfix))] @ + mk_type_decl((ts, n), Some(tinfix)) + | mk_type_decl(([], n), Some(tinfix)) = []; + + +fun mk_extension (((((((cl, def), (ty, tinfix)), ar), (co, mfix)), tr), ax), ml) = + ((cl, def, ty, ar, co, ax), big_bracket_comma_ind " " tinfix, + big_bracket_comma_ind " " mfix, big_bracket_comma_ind " " tr, ml); + +fun add_val((id,_),s) = "val " ^ id ^ " = ax " ^ string id ^ "\n" ^ s; + +fun mk_rules ps = + let + val axs = big_bracket_comma_ind " " (map pair_string ps); + val vals = foldr add_val (ps, "") + in + axs ^ "\n\nval ax = get_axiom thy\n\n" ^ vals + end; + +fun mk_struct (id, s) = "structure " ^ id ^ " =\nstruct\n" ^ s ^ "\nend\n"; + + +fun mk_sext mfix trans = + "Some (NewSext {\n\ +\ mixfix =\n " ^ mfix ^ ",\n\ +\ xrules =\n " ^ trans ^ ",\n\ +\ parse_ast_translation = parse_ast_translation,\n\ +\ parse_preproc = parse_preproc,\n\ +\ parse_postproc = parse_postproc,\n\ +\ parse_translation = parse_translation,\n\ +\ print_translation = print_translation,\n\ +\ print_preproc = print_preproc,\n\ +\ print_postproc = print_postproc,\n\ +\ print_ast_translation = print_ast_translation})"; + +fun mk_simple_sext mfix = + "Some (Syntax.simple_sext\n " ^ mfix ^ ")"; + +fun mk_ext ((cl, def, ty, ar, co, ax), sext) = + " (" ^ space_implode ",\n " [cl, def, ty, ar, co, sext] ^ ")\n " ^ ax ^ "\n"; + +fun mk_ext_thy (base, name, ext, sext) = + "extend_theory (" ^ base ^ ")\n " ^ string name ^ "\n" ^ mk_ext (ext, sext); + +val preamble = + "\nlocal\n\ + \ val parse_ast_translation = []\n\ + \ val parse_preproc = None\n\ + \ val parse_postproc = None\n\ + \ val parse_translation = []\n\ + \ val print_translation = []\n\ + \ val print_preproc = None\n\ + \ val print_postproc = None\n\ + \ val print_ast_translation = []\n\ + \in\n\n\ + \(**** begin of user section ****)\n"; + +val postamble = "\n(**** end of user section ****)\n"; + +fun mk_structure ((name, base), Some (ext, tinfix, mfix, trans, ml)) = + let + val noext = ("[]", "[]", "[]", "[]", "[]", "[]"); + val basethy = + if tinfix = "[]" then base + else mk_ext_thy (base, name ^ "(type infix)", noext, mk_simple_sext tinfix); + val sext = + if mfix = "[]" andalso trans = "[]" andalso ml = "" then "None" + else mk_sext mfix trans; + val thy = "\nval thy = " ^ mk_ext_thy (basethy, name, ext, sext); + in + mk_struct (name, preamble ^ ml ^ postamble ^ thy ^ "\nend") + end + | mk_structure ((name, base), None) = mk_struct (name, "\nval thy = " ^ base); + + +fun merge (t :: ts) = + foldl (fn (base, thy) => "(merge_theories (" ^ base ^ ", " ^ thy ^ ".thy))") + (t ^ ".thy", ts) + | merge [] = raise Match; + + + +(*------------------------ PARSERS -------------------------*) + + +open Parse + +(*------------------- VARIOUS PARSERS ----------------------*) + +val emptyl = empty >> K"[]"; + +val ids = list_of1 id >> bracket_comma_string; +(* -> "[id1, id2, ..., idn]" *) + +val stgorids = list_of1 (stg || id); + +val sort = id >> (bracket o string) + || "{" $$-- (ids || emptyl) --$$ "}"; +(* -> "[id]" + -> "[id1, ...,idn]" *) + +val infxl = "infixl" $$-- !! nat +and infxr = "infixr" $$-- !! nat + + +(*------------------- CLASSES PARSER ----------------------*) + + + + +val class = (id >> string) -- ( "<" $$-- (!! ids) || emptyl) >> pair; + +(* -> "(id, [id1, ..., idn])" + || + -> "(id, [])" *) + + +val classes = "classes" $$-- !!(repeat1 class) >> bracket_comma + || emptyl; + + +(* "[(id, [..]), ...,(id, [...])]" *) + + + +(*------------------- DEFAULT PARSER ---------------------*) + + +val default = "default" $$-- !!sort + || emptyl; + +(* -> "[]" + -> "[id]" + -> "[id1, ...,idn]" *) + + +(*-------------------- TYPES PARSER ----------------------*) + + +val type_decl = stgorids -- nat; + +val tyinfix = infxl >> (Some o TInfixl) + || infxr >> (Some o TInfixr); + +val type_infix = "(" $$-- !! (tyinfix --$$ ")") + || empty >> K None; + +val types = "types" $$-- + !! (repeat1 (type_decl -- type_infix >> mk_type_decl)) + >> (split_decls o flat) + || empty >> (K ("[]", [])); + + (* ==> ("[(id, nat), ... ]", [strg, ...]) *) + + + +(*-------------------- ARITIES PARSER ----------------------*) + + +val sorts = list_of sort >> bracket_comma; + +(* -> "[[id1, ...], ..., [id, ...]]" *) + + +val arity = id >> (fn s => pair("[]",string s)) + || "(" $$-- sorts --$$")" -- id >> (fn (l,s) => pair(l,string s)); + +(* -> "([],id)" + -> "([[id,..], ...,[id,..]], id)" *) + +val tys = stgorids >> bracket_comma_string; + +val arities = "arities" $$-- !! (repeat1 (tys --$$ "::" -- arity >> pair)) + >> bracket_comma + || emptyl; + +(* -> "[([id,..], ([[id,...],...], id))]" *) + + +(*--------------------- CONSTS PARSER ---------------------*) + +val natlist = "[" $$-- !!(list_of nat --$$ "]") >> bracket_comma + || empty >> K"[]"; + + + (* "[nat, ...]" || "[]" *) + + +val prio_opt = natlist -- nat >> Some + || empty >> K None; + +val mfix = stg -- !! prio_opt >> delim_mix + || infxl >> Infixl + || infxr >> Infixr + || "binder" $$-- !!(stg -- nat) >> Binder + +val const_decl = stgorids -- !! ("::" $$-- stg); + +(*("[exid, ...]", stg) *) + + +val mixfix = "(" $$-- !! (mfix --$$ ")") >> Some + || empty >> K None; + +(* (s, e, l, n) *) + + +val consts = "consts" $$-- + !! (repeat1 (const_decl -- mixfix >> mk_mixfix)) + >> (split_decls o flat) + || empty >> K ("[]",[]); + +(* ("[([exid, ...], stg), ....]", [strg, ..]) *) + + +(*---------------- TRANSLATIONS PARSER --------------------*) + +val xpat = "(" $$-- id --$$ ")" -- stg >> pair_string + || stg >> (fn s => pair_string ("logic", s)); + +val arrow = $$ "=>" >> K " |-> " + || $$ "<=" >> K " <-| " + || $$ "==" >> K " <-> "; + +val xrule = xpat -- !! (arrow -- xpat) >> (fn (xp1, (a, xp2)) => xp1 ^ a ^ xp2); + +val translations = "translations" $$-- !! (repeat1 xrule) + || empty; + + +(*------------------- RULES PARSER -----------------------*) + +val rules = "rules" $$-- !! (repeat1 (id -- !! stg) >> mk_rules) + || emptyl; + +(* "[(id, stg), ...]" *) + + + +(*----------------------- ML_TEXT -------------------------*) + +val mltxt = txt || empty >> K""; + + +(*---------------------------------------------------------*) + +val extension = "+" $$-- !! (classes -- default -- types -- arities + -- consts -- translations -- rules --$$ "end" -- mltxt) + >> (Some o mk_extension) + || empty >> K None; + + +val bases = id -- repeat("+" $$-- id) >> op:: ; + +val theoryDef = !!(id --$$ "=" -- (bases >> merge) -- extension) + >> mk_structure; + +val read = reader theoryDef + +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/drule.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/drule.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,465 @@ +(* Title: drule + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Derived rules and other operations on theorems and theories +*) + +infix 0 RS RSN RL RLN COMP; + +signature DRULE = + sig + structure Thm : THM + local open Thm in + val asm_rl: thm + val assume_ax: theory -> string -> thm + val COMP: thm * thm -> thm + val compose: thm * int * thm -> thm list + val cterm_instantiate: (Sign.cterm*Sign.cterm)list -> thm -> thm + val cut_rl: thm + val equal_abs_elim: Sign.cterm -> thm -> thm + val equal_abs_elim_list: Sign.cterm list -> thm -> thm + val eq_sg: Sign.sg * Sign.sg -> bool + val eq_thm: thm * thm -> bool + val eq_thm_sg: thm * thm -> bool + val flexpair_abs_elim_list: Sign.cterm list -> thm -> thm + val forall_intr_list: Sign.cterm list -> thm -> thm + val forall_intr_frees: thm -> thm + val forall_elim_list: Sign.cterm list -> thm -> thm + val forall_elim_var: int -> thm -> thm + val forall_elim_vars: int -> thm -> thm + val implies_elim_list: thm -> thm list -> thm + val implies_intr_list: Sign.cterm list -> thm -> thm + val print_cterm: Sign.cterm -> unit + val print_ctyp: Sign.ctyp -> unit + val print_goals: int -> thm -> unit + val print_sg: Sign.sg -> unit + val print_theory: theory -> unit + val pprint_sg: Sign.sg -> pprint_args -> unit + val pprint_theory: theory -> pprint_args -> unit + val print_thm: thm -> unit + val prth: thm -> thm + val prthq: thm Sequence.seq -> thm Sequence.seq + val prths: thm list -> thm list + val read_instantiate: (string*string)list -> thm -> thm + val read_instantiate_sg: Sign.sg -> (string*string)list -> thm -> thm + val reflexive_thm: thm + val revcut_rl: thm + val rewrite_goal_rule: (meta_simpset -> thm -> thm option) -> meta_simpset -> + int -> thm -> thm + val rewrite_goals_rule: thm list -> thm -> thm + val rewrite_rule: thm list -> thm -> thm + val RS: thm * thm -> thm + val RSN: thm * (int * thm) -> thm + val RL: thm list * thm list -> thm list + val RLN: thm list * (int * thm list) -> thm list + val show_hyps: bool ref + val size_of_thm: thm -> int + val standard: thm -> thm + val string_of_thm: thm -> string + val symmetric_thm: thm + val pprint_thm: thm -> pprint_args -> unit + val transitive_thm: thm + val triv_forall_equality: thm + val types_sorts: thm -> (indexname-> typ option) * (indexname-> sort option) + val zero_var_indexes: thm -> thm + end + end; + +functor DruleFun (structure Logic: LOGIC and Thm: THM) : DRULE = +struct +structure Thm = Thm; +structure Sign = Thm.Sign; +structure Type = Sign.Type; +structure Pretty = Sign.Syntax.Pretty +local open Thm +in + +(**** More derived rules and operations on theorems ****) + +(*** Find the type (sort) associated with a (T)Var or (T)Free in a term + Used for establishing default types (of variables) and sorts (of + type variables) when reading another term. + Index -1 indicates that a (T)Free rather than a (T)Var is wanted. +***) + +fun types_sorts thm = + let val {prop,hyps,...} = rep_thm thm; + val big = list_comb(prop,hyps); (* bogus term! *) + val vars = map dest_Var (term_vars big); + val frees = map dest_Free (term_frees big); + val tvars = term_tvars big; + val tfrees = term_tfrees big; + fun typ(a,i) = if i<0 then assoc(frees,a) else assoc(vars,(a,i)); + fun sort(a,i) = if i<0 then assoc(tfrees,a) else assoc(tvars,(a,i)); + in (typ,sort) end; + +(** Standardization of rules **) + +(*Generalization over a list of variables, IGNORING bad ones*) +fun forall_intr_list [] th = th + | forall_intr_list (y::ys) th = + let val gth = forall_intr_list ys th + in forall_intr y gth handle THM _ => gth end; + +(*Generalization over all suitable Free variables*) +fun forall_intr_frees th = + let val {prop,sign,...} = rep_thm th + in forall_intr_list + (map (Sign.cterm_of sign) (sort atless (term_frees prop))) + th + end; + +(*Replace outermost quantified variable by Var of given index. + Could clash with Vars already present.*) +fun forall_elim_var i th = + let val {prop,sign,...} = rep_thm th + in case prop of + Const("all",_) $ Abs(a,T,_) => + forall_elim (Sign.cterm_of sign (Var((a,i), T))) th + | _ => raise THM("forall_elim_var", i, [th]) + end; + +(*Repeat forall_elim_var until all outer quantifiers are removed*) +fun forall_elim_vars i th = + forall_elim_vars i (forall_elim_var i th) + handle THM _ => th; + +(*Specialization over a list of cterms*) +fun forall_elim_list cts th = foldr (uncurry forall_elim) (rev cts, th); + +(* maps [A1,...,An], B to [| A1;...;An |] ==> B *) +fun implies_intr_list cAs th = foldr (uncurry implies_intr) (cAs,th); + +(* maps [| A1;...;An |] ==> B and [A1,...,An] to B *) +fun implies_elim_list impth ths = foldl (uncurry implies_elim) (impth,ths); + +(*Reset Var indexes to zero, renaming to preserve distinctness*) +fun zero_var_indexes th = + let val {prop,sign,...} = rep_thm th; + val vars = term_vars prop + val bs = foldl add_new_id ([], map (fn Var((a,_),_)=>a) vars) + val inrs = add_term_tvars(prop,[]); + val nms' = rev(foldl add_new_id ([], map (#1 o #1) inrs)); + val tye = map (fn ((v,rs),a) => (v, TVar((a,0),rs))) (inrs ~~ nms') + val ctye = map (fn (v,T) => (v,Sign.ctyp_of sign T)) tye; + fun varpairs([],[]) = [] + | varpairs((var as Var(v,T)) :: vars, b::bs) = + let val T' = typ_subst_TVars tye T + in (Sign.cterm_of sign (Var(v,T')), + Sign.cterm_of sign (Var((b,0),T'))) :: varpairs(vars,bs) + end + | varpairs _ = raise TERM("varpairs", []); + in instantiate (ctye, varpairs(vars,rev bs)) th end; + + +(*Standard form of object-rule: no hypotheses, Frees, or outer quantifiers; + all generality expressed by Vars having index 0.*) +fun standard th = + let val {maxidx,...} = rep_thm th + in varifyT (zero_var_indexes (forall_elim_vars(maxidx+1) + (forall_intr_frees(implies_intr_hyps th)))) + end; + +(*Assume a new formula, read following the same conventions as axioms. + Generalizes over Free variables, + creates the assumption, and then strips quantifiers. + Example is [| ALL x:?A. ?P(x) |] ==> [| ?P(?a) |] + [ !(A,P,a)[| ALL x:A. P(x) |] ==> [| P(a) |] ] *) +fun assume_ax thy sP = + let val sign = sign_of thy + val prop = Logic.close_form (Sign.term_of (Sign.read_cterm sign + (sP, propT))) + in forall_elim_vars 0 (assume (Sign.cterm_of sign prop)) end; + +(*Resolution: exactly one resolvent must be produced.*) +fun tha RSN (i,thb) = + case Sequence.chop (2, biresolution false [(false,tha)] i thb) of + ([th],_) => th + | ([],_) => raise THM("RSN: no unifiers", i, [tha,thb]) + | _ => raise THM("RSN: multiple unifiers", i, [tha,thb]); + +(*resolution: P==>Q, Q==>R gives P==>R. *) +fun tha RS thb = tha RSN (1,thb); + +(*For joining lists of rules*) +fun thas RLN (i,thbs) = + let val resolve = biresolution false (map (pair false) thas) i + fun resb thb = Sequence.list_of_s (resolve thb) handle THM _ => [] + in flat (map resb thbs) end; + +fun thas RL thbs = thas RLN (1,thbs); + +(*compose Q and [...,Qi,Q(i+1),...]==>R to [...,Q(i+1),...]==>R + with no lifting or renaming! Q may contain ==> or meta-quants + ALWAYS deletes premise i *) +fun compose(tha,i,thb) = + Sequence.list_of_s (bicompose false (false,tha,0) i thb); + +(*compose Q and [Q1,Q2,...,Qk]==>R to [Q2,...,Qk]==>R getting unique result*) +fun tha COMP thb = + case compose(tha,1,thb) of + [th] => th + | _ => raise THM("COMP", 1, [tha,thb]); + +(*Instantiate theorem th, reading instantiations under signature sg*) +fun read_instantiate_sg sg sinsts th = + let val ts = types_sorts th; + val instpair = Sign.read_insts sg ts ts sinsts + in instantiate instpair th end; + +(*Instantiate theorem th, reading instantiations under theory of th*) +fun read_instantiate sinsts th = + read_instantiate_sg (#sign (rep_thm th)) sinsts th; + + +(*Left-to-right replacements: tpairs = [...,(vi,ti),...]. + Instantiates distinct Vars by terms, inferring type instantiations. *) +local + fun add_types ((ct,cu), (sign,tye)) = + let val {sign=signt, t=t, T= T, ...} = Sign.rep_cterm ct + and {sign=signu, t=u, T= U, ...} = Sign.rep_cterm cu + val sign' = Sign.merge(sign, Sign.merge(signt, signu)) + val tye' = Type.unify (#tsig(Sign.rep_sg sign')) ((T,U), tye) + handle Type.TUNIFY => raise TYPE("add_types", [T,U], [t,u]) + in (sign', tye') end; +in +fun cterm_instantiate ctpairs0 th = + let val (sign,tye) = foldr add_types (ctpairs0, (#sign(rep_thm th),[])) + val tsig = #tsig(Sign.rep_sg sign); + fun instT(ct,cu) = let val inst = subst_TVars tye + in (Sign.cfun inst ct, Sign.cfun inst cu) end + fun ctyp2 (ix,T) = (ix, Sign.ctyp_of sign T) + in instantiate (map ctyp2 tye, map instT ctpairs0) th end + handle TERM _ => + raise THM("cterm_instantiate: incompatible signatures",0,[th]) + | TYPE _ => raise THM("cterm_instantiate: types", 0, [th]) +end; + + +(*** Printing of theorems ***) + +(*If false, hypotheses are printed as dots*) +val show_hyps = ref true; + +fun pretty_thm th = +let val {sign, hyps, prop,...} = rep_thm th + val hsymbs = if null hyps then [] + else if !show_hyps then + [Pretty.brk 2, + Pretty.lst("[","]") (map (Sign.pretty_term sign) hyps)] + else Pretty.str" [" :: map (fn _ => Pretty.str".") hyps @ + [Pretty.str"]"]; +in Pretty.blk(0, Sign.pretty_term sign prop :: hsymbs) end; + +val string_of_thm = Pretty.string_of o pretty_thm; + +val pprint_thm = Pretty.pprint o Pretty.quote o pretty_thm; + + +(** Top-level commands for printing theorems **) +val print_thm = writeln o string_of_thm; + +fun prth th = (print_thm th; th); + +(*Print and return a sequence of theorems, separated by blank lines. *) +fun prthq thseq = + (Sequence.prints (fn _ => print_thm) 100000 thseq; + thseq); + +(*Print and return a list of theorems, separated by blank lines. *) +fun prths ths = (print_list_ln print_thm ths; ths); + +(*Other printing commands*) +val print_cterm = writeln o Sign.string_of_cterm; +val print_ctyp = writeln o Sign.string_of_ctyp; +fun pretty_sg sg = + Pretty.lst ("{", "}") (map (Pretty.str o !) (#stamps (Sign.rep_sg sg))); + +val pprint_sg = Pretty.pprint o pretty_sg; + +val pprint_theory = pprint_sg o sign_of; + +val print_sg = writeln o Pretty.string_of o pretty_sg; +val print_theory = print_sg o sign_of; + + +(** Print thm A1,...,An/B in "goal style" -- premises as numbered subgoals **) + +fun prettyprints es = writeln(Pretty.string_of(Pretty.blk(0,es))); + +fun print_goals maxgoals th : unit = +let val {sign, hyps, prop,...} = rep_thm th; + fun printgoals (_, []) = () + | printgoals (n, A::As) = + let val prettyn = Pretty.str(" " ^ string_of_int n ^ ". "); + val prettyA = Sign.pretty_term sign A + in prettyprints[prettyn,prettyA]; + printgoals (n+1,As) + end; + fun prettypair(t,u) = + Pretty.blk(0, [Sign.pretty_term sign t, Pretty.str" =?=", Pretty.brk 1, + Sign.pretty_term sign u]); + fun printff [] = () + | printff tpairs = + writeln("\nFlex-flex pairs:\n" ^ + Pretty.string_of(Pretty.lst("","") (map prettypair tpairs))) + val (tpairs,As,B) = Logic.strip_horn(prop); + val ngoals = length As +in + writeln (Sign.string_of_term sign B); + if ngoals=0 then writeln"No subgoals!" + else if ngoals>maxgoals + then (printgoals (1, take(maxgoals,As)); + writeln("A total of " ^ string_of_int ngoals ^ " subgoals...")) + else printgoals (1, As); + printff tpairs +end; + + +(** theorem equality test is exported and used by BEST_FIRST **) + +(*equality of signatures means exact identity -- by ref equality*) +fun eq_sg (sg1,sg2) = (#stamps(Sign.rep_sg sg1) = #stamps(Sign.rep_sg sg2)); + +(*equality of theorems uses equality of signatures and + the a-convertible test for terms*) +fun eq_thm (th1,th2) = + let val {sign=sg1, hyps=hyps1, prop=prop1, ...} = rep_thm th1 + and {sign=sg2, hyps=hyps2, prop=prop2, ...} = rep_thm th2 + in eq_sg (sg1,sg2) andalso + aconvs(hyps1,hyps2) andalso + prop1 aconv prop2 + end; + +(*Do the two theorems have the same signature?*) +fun eq_thm_sg (th1,th2) = eq_sg(#sign(rep_thm th1), #sign(rep_thm th2)); + +(*Useful "distance" function for BEST_FIRST*) +val size_of_thm = size_of_term o #prop o rep_thm; + + +(*** Meta-Rewriting Rules ***) + + +val reflexive_thm = + let val cx = Sign.cterm_of Sign.pure (Var(("x",0),TVar(("'a",0),["logic"]))) + in Thm.reflexive cx end; + +val symmetric_thm = + let val xy = Sign.read_cterm Sign.pure ("x::'a::logic == y",propT) + in standard(Thm.implies_intr_hyps(Thm.symmetric(Thm.assume xy))) end; + +val transitive_thm = + let val xy = Sign.read_cterm Sign.pure ("x::'a::logic == y",propT) + val yz = Sign.read_cterm Sign.pure ("y::'a::logic == z",propT) + val xythm = Thm.assume xy and yzthm = Thm.assume yz + in standard(Thm.implies_intr yz (Thm.transitive xythm yzthm)) end; + + +(** Below, a "conversion" has type sign->term->thm **) + +(*In [A1,...,An]==>B, rewrite the selected A's only -- for rewrite_goals_tac*) +fun goals_conv pred cv sign = + let val triv = reflexive o Sign.cterm_of sign + fun gconv i t = + let val (A,B) = Logic.dest_implies t + val thA = if (pred i) then (cv sign A) else (triv A) + in combination (combination (triv implies) thA) + (gconv (i+1) B) + end + handle TERM _ => triv t + in gconv 1 end; + +(*Use a conversion to transform a theorem*) +fun fconv_rule cv th = + let val {sign,prop,...} = rep_thm th + in equal_elim (cv sign prop) th end; + +(*rewriting conversion*) +fun rew_conv prover mss sign t = + rewrite_cterm mss prover (Sign.cterm_of sign t); + +(*Rewrite a theorem*) +fun rewrite_rule thms = fconv_rule (rew_conv (K(K None)) (Thm.mss_of thms)); + +(*Rewrite the subgoals of a proof state (represented by a theorem) *) +fun rewrite_goals_rule thms = + fconv_rule (goals_conv (K true) (rew_conv (K(K None)) (Thm.mss_of thms))); + +(*Rewrite the subgoal of a proof state (represented by a theorem) *) +fun rewrite_goal_rule prover mss i = + fconv_rule (goals_conv (fn j => j=i) (rew_conv prover mss)); + + +(** Derived rules mainly for METAHYPS **) + +(*Given the term "a", takes (%x.t)==(%x.u) to t[a/x]==u[a/x]*) +fun equal_abs_elim ca eqth = + let val {sign=signa, t=a, ...} = Sign.rep_cterm ca + and combth = combination eqth (reflexive ca) + val {sign,prop,...} = rep_thm eqth + val (abst,absu) = Logic.dest_equals prop + val cterm = Sign.cterm_of (Sign.merge (sign,signa)) + in transitive (symmetric (beta_conversion (cterm (abst$a)))) + (transitive combth (beta_conversion (cterm (absu$a)))) + end + handle THM _ => raise THM("equal_abs_elim", 0, [eqth]); + +(*Calling equal_abs_elim with multiple terms*) +fun equal_abs_elim_list cts th = foldr (uncurry equal_abs_elim) (rev cts, th); + +local + open Logic + val alpha = TVar(("'a",0), []) (* type ?'a::{} *) + fun err th = raise THM("flexpair_inst: ", 0, [th]) + fun flexpair_inst def th = + let val {prop = Const _ $ t $ u, sign,...} = rep_thm th + val cterm = Sign.cterm_of sign + fun cvar a = cterm(Var((a,0),alpha)) + val def' = cterm_instantiate [(cvar"t", cterm t), (cvar"u", cterm u)] + def + in equal_elim def' th + end + handle THM _ => err th | bind => err th +in +val flexpair_intr = flexpair_inst (symmetric flexpair_def) +and flexpair_elim = flexpair_inst flexpair_def +end; + +(*Version for flexflex pairs -- this supports lifting.*) +fun flexpair_abs_elim_list cts = + flexpair_intr o equal_abs_elim_list cts o flexpair_elim; + + +(*** Some useful meta-theorems ***) + +(*The rule V/V, obtains assumption solving for eresolve_tac*) +val asm_rl = trivial(Sign.read_cterm Sign.pure ("PROP ?psi",propT)); + +(*Meta-level cut rule: [| V==>W; V |] ==> W *) +val cut_rl = trivial(Sign.read_cterm Sign.pure + ("PROP ?psi ==> PROP ?theta", propT)); + +(*Generalized elim rule for one conclusion; cut_rl with reversed premises: + [| PROP V; PROP V ==> PROP W |] ==> PROP W *) +val revcut_rl = + let val V = Sign.read_cterm Sign.pure ("PROP V", propT) + and VW = Sign.read_cterm Sign.pure ("PROP V ==> PROP W", propT); + in standard (implies_intr V + (implies_intr VW + (implies_elim (assume VW) (assume V)))) + end; + +(* (!!x. PROP ?V) == PROP ?V Allows removal of redundant parameters*) +val triv_forall_equality = + let val V = Sign.read_cterm Sign.pure ("PROP V", propT) + and QV = Sign.read_cterm Sign.pure ("!!x::'a. PROP V", propT) + and x = Sign.read_cterm Sign.pure ("x", TFree("'a",["logic"])); + in standard (equal_intr (implies_intr QV (forall_elim x (assume QV))) + (implies_intr V (forall_intr x (assume V)))) + end; + +end +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/envir.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/envir.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,207 @@ +(* Title: envir + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1988 University of Cambridge +*) + +(*Environments + Should take account that typ is part of variable name, + otherwise two variables with same name and different types + will cause errors! +*) + + +signature ENVIR = +sig + type 'a xolist + exception ENVIR of string * indexname; + datatype env = Envir of {asol: term xolist, + iTs: (indexname * typ) list, + maxidx: int} + val alist_of : env -> (indexname * term) list + val alist_of_olist : 'a xolist -> (indexname * 'a) list + val empty : int->env + val genvar : string -> env * typ -> env * term + val genvars : string -> env * typ list -> env * term list + val lookup : env * indexname -> term option + val norm_term : env -> term -> term + val null_olist : 'a xolist + val olist_of_alist: (indexname * 'a) list -> 'a xolist + val update : (indexname * term) * env -> env + val vupdate : (indexname * term) * env -> env +end; + +functor EnvirFun () : ENVIR = +struct + + +(*association lists with keys in ascending order, no repeated keys*) +datatype 'a xolist = Olist of (indexname * 'a) list; + + +exception ENVIR of string * indexname; + +(*look up key in ordered list*) +fun xsearch (Olist[], key) = None + | xsearch (Olist ((keyi,xi)::pairs), key) = + if xless(keyi,key) then xsearch (Olist pairs, key) + else if key=keyi then Some xi + else None; + +(*insert pair in ordered list, reject if key already present*) +fun xinsert_new (newpr as (key, x), Olist al) = + let fun insert [] = [newpr] + | insert ((pair as (keyi,xi)) :: pairs) = + if xless(keyi,key) then pair :: insert pairs + else if key=keyi then + raise ENVIR("xinsert_new: already present",key) + else newpr::pair::pairs + in Olist (insert al) end; + +(*insert pair in ordered list, overwrite if key already present*) +fun xinsert (newpr as (key, x), Olist al) = + let fun insert [] = [newpr] + | insert ((pair as (keyi,xi)) :: pairs) = + if xless(keyi,key) then pair :: insert pairs + else if key=keyi then newpr::pairs + else newpr::pair::pairs + in Olist (insert al) end; + +(*apply function to the contents part of each pair*) +fun xmap f (Olist pairs) = + let fun xmp [] = [] + | xmp ((key,x)::pairs) = (key, f x) :: xmp pairs + in Olist (xmp pairs) end; + +(*allows redefinition of a key by "join"ing the contents parts*) +fun xmerge_olists join (Olist pairsa, Olist pairsb) = + let fun xmrg ([],pairsb) = pairsb + | xmrg (pairsa,[]) = pairsa + | xmrg ((keya,x)::pairsa, (keyb,y)::pairsb) = + if xless(keya,keyb) then + (keya,x) :: xmrg (pairsa, (keyb,y)::pairsb) + else if xless(keyb,keya) then + (keyb,y) :: xmrg ((keya,x)::pairsa, pairsb) + else (*equal*) (keya, join(x,y)) :: xmrg (pairsa, pairsb) + in Olist (xmrg (pairsa,pairsb)) end; + +val null_olist = Olist[]; + +fun alist_of_olist (Olist pairs) = pairs; + +fun olist_of_alist pairs = foldr xinsert (pairs, Olist[]); + + + +(*updating can destroy environment in 2 ways!! + (1) variables out of range (2) circular assignments +*) +datatype env = Envir of + {maxidx: int, (*maximum index of vars*) + asol: term xolist, (*table of assignments to Vars*) + iTs: (indexname*typ)list} (*table of assignments to TVars*) + + +(*Generate a list of distinct variables. + Increments index to make them distinct from ALL present variables. *) +fun genvars name (Envir{maxidx, asol, iTs}, Ts) : env * term list = + let fun genvs (_, [] : typ list) : term list = [] + | genvs (n, [T]) = [ Var((name, maxidx+1), T) ] + | genvs (n, T::Ts) = + Var((name ^ radixstring(26,"a",n), maxidx+1), T) + :: genvs(n+1,Ts) + in (Envir{maxidx=maxidx+1, asol=asol, iTs=iTs}, genvs (0,Ts)) end; + +(*Generate a variable.*) +fun genvar name (env,T) : env * term = + let val (env',[v]) = genvars name (env,[T]) + in (env',v) end; + +fun lookup (Envir{asol,...}, xname) : term option = xsearch (asol,xname); + +fun update ((xname, t), Envir{maxidx, asol, iTs}) = + Envir{maxidx=maxidx, asol=xinsert_new ((xname,t), asol), iTs=iTs}; + +(*The empty environment. New variables will start with the given index.*) +fun empty m = Envir{maxidx=m, asol=null_olist, iTs=[]}; + +(*Update, checking Var-Var assignments: try to suppress higher indexes*) +fun vupdate((a,t), env) = case t of + Var(name',T) => + if a=name' then env (*cycle!*) + else if xless(a, name') then + (case lookup(env,name') of (*if already assigned, chase*) + None => update((name', Var(a,T)), env) + | Some u => vupdate((a,u), env)) + else update((a,t), env) + | _ => update((a,t), env); + + +(*Convert environment to alist*) +fun alist_of (Envir{maxidx,asol,...}) = alist_of_olist asol; + + +(*Beta normal form for terms (not eta normal form). + Chases variables in env; Does not exploit sharing of variable bindings + Does not check types, so could loop. *) +local + (*raised when norm has no effect on a term, + to encourage sharing instead of copying*) + exception SAME; + + fun norm_term1 (asol,t) : term = + let fun norm (Var (w,T)) = + (case xsearch(asol,w) of + Some u => normh u + | None => raise SAME) + | norm (Abs(a,T,body)) = Abs(a, T, norm body) + | norm (Abs(_,_,body) $ t) = normh(subst_bounds([t], body)) + | norm (f $ t) = + ((case norm f of + Abs(_,_,body) => normh(subst_bounds([t], body)) + | nf => nf $ normh t) + handle SAME => f $ norm t) + | norm _ = raise SAME + and normh t = (norm t) handle SAME => t + in normh t end + + and norm_term2(asol,iTs,t) : term = + let fun normT(Type(a,Ts)) = Type(a, normTs Ts) + | normT(TFree _) = raise SAME + | normT(TVar(v,_)) = (case assoc(iTs,v) of + Some(U) => normTh U + | None => raise SAME) + and normTh T = ((normT T) handle SAME => T) + and normTs([]) = raise SAME + | normTs(T::Ts) = ((normT T :: normTsh Ts) + handle SAME => T :: normTs Ts) + and normTsh Ts = ((normTs Ts) handle SAME => Ts) + and norm(Const(a,T)) = Const(a, normT T) + | norm(Free(a,T)) = Free(a, normT T) + | norm(Var (w,T)) = + (case xsearch(asol,w) of + Some u => normh u + | None => Var(w,normT T)) + | norm(Abs(a,T,body)) = + (Abs(a,normT T,normh body) handle SAME => Abs(a, T, norm body)) + | norm(Abs(_,_,body) $ t) = normh(subst_bounds([t], body)) + | norm(f $ t) = + ((case norm f of + Abs(_,_,body) => normh(subst_bounds([t], body)) + | nf => nf $ normh t) + handle SAME => f $ norm t) + | norm _ = raise SAME + and normh t = (norm t) handle SAME => t + in normh t end; + +in + +(*curried version might be slower in recursive calls*) +fun norm_term (env as Envir{asol,iTs,...}) t : term = + if iTs=[] then norm_term1(asol, t) else norm_term2(asol,iTs, t) + +end; + +end; + + diff -r 000000000000 -r a5a9c433f639 src/Pure/goals.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/goals.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,438 @@ +(* Title: goals + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Goal stack package. The goal stack initially holds a dummy proof, and can +never become empty. Each goal stack consists of a list of levels. The +undo list is a list of goal stacks. Finally, there may be a stack of +pending proofs. +*) + + +signature GOALS = + sig + structure Tactical: TACTICAL + local open Tactical Tactical.Thm in + type proof + val ba: int -> unit + val back: unit -> unit + val bd: thm -> int -> unit + val bds: thm list -> int -> unit + val be: thm -> int -> unit + val bes: thm list -> int -> unit + val br: thm -> int -> unit + val brs: thm list -> int -> unit + val bw: thm -> unit + val bws: thm list -> unit + val by: tactic -> unit + val byev: tactic list -> unit + val chop: unit -> unit + val choplev: int -> unit + val fa: unit -> unit + val fd: thm -> unit + val fds: thm list -> unit + val fe: thm -> unit + val fes: thm list -> unit + val filter_goal: (term*term->bool) -> thm list -> int -> thm list + val fr: thm -> unit + val frs: thm list -> unit + val getgoal: int -> term + val gethyps: int -> thm list + val goal: theory -> string -> thm list + val goalw: theory -> thm list -> string -> thm list + val goalw_cterm: thm list -> Sign.cterm -> thm list + val pop_proof: unit -> thm list + val pr: unit -> unit + val premises: unit -> thm list + val prin: term -> unit + val printyp: typ -> unit + val pprint_term: term -> pprint_args -> unit + val pprint_typ: typ -> pprint_args -> unit + val print_exn: exn -> 'a + val prlev: int -> unit + val proof_timing: bool ref + val prove_goal: theory -> string -> (thm list -> tactic list) -> thm + val prove_goalw: theory->thm list->string->(thm list->tactic list)->thm + val prove_goalw_cterm: thm list->Sign.cterm->(thm list->tactic list)->thm + val push_proof: unit -> unit + val read: string -> term + val ren: string -> int -> unit + val restore_proof: proof -> thm list + val result: unit -> thm + val rotate_proof: unit -> thm list + val uresult: unit -> thm + val save_proof: unit -> proof + val topthm: unit -> thm + val undo: unit -> unit + end + end; + +functor GoalsFun (structure Logic: LOGIC and Drule: DRULE and Tactic: TACTIC + and Pattern:PATTERN + sharing type Pattern.type_sig = Drule.Thm.Sign.Type.type_sig + and Drule.Thm = Tactic.Tactical.Thm) : GOALS = +struct +structure Tactical = Tactic.Tactical +local open Tactic Tactic.Tactical Tactic.Tactical.Thm Drule +in + +(*Each level of goal stack includes a proof state and alternative states, + the output of the tactic applied to the preceeding level. *) +type gstack = (thm * thm Sequence.seq) list; + +datatype proof = Proof of gstack list * thm list * (bool*thm->thm); + +(*** References ***) + +(*Should process time be printed after proof steps?*) +val proof_timing = ref false; + +(*Current assumption list -- set by "goal".*) +val curr_prems = ref([] : thm list); + +(*Return assumption list -- useful if you didn't save "goal"'s result. *) +fun premises() = !curr_prems; + +(*Current result maker -- set by "goal", used by "result". *) +val curr_mkresult = + ref((fn _=> error"No goal has been supplied in subgoal module") + : bool*thm->thm); + +val dummy = trivial(Sign.read_cterm Sign.pure + ("PROP No_goal_has_been_supplied",propT)); + +(*List of previous goal stacks, for the undo operation. Set by setstate. + A list of lists!*) +val undo_list = ref([[(dummy, Sequence.null)]] : gstack list); + +(* Stack of proof attempts *) +val proofstack = ref([]: proof list); + + +(*** Setting up goal-directed proof ***) + +(*Generates the list of new theories when the proof state's signature changes*) +fun sign_error (sign,sign') = + let val stamps = #stamps(Sign.rep_sg sign') \\ + #stamps(Sign.rep_sg sign) + in case stamps of + [stamp] => "\nNew theory: " ^ !stamp + | _ => "\nNew theories: " ^ space_implode ", " (map ! stamps) + end; + +(*Common treatment of "goal" and "prove_goal": + Return assumptions, initial proof state, and function to make result. *) +fun prepare_proof rths chorn = + let val {sign, t=horn,...} = Sign.rep_cterm chorn; + val (_,As,B) = Logic.strip_horn(horn); + val cAs = map (Sign.cterm_of sign) As; + val p_rew = if null rths then I else rewrite_rule rths; + val c_rew = if null rths then I else rewrite_goals_rule rths; + val prems = map (p_rew o forall_elim_vars(0) o assume) cAs + and st0 = c_rew (trivial (Sign.cterm_of sign B)) + fun result_error state msg = + (writeln ("Bad final proof state:"); + print_goals (!goals_limit) state; + error msg) + (*discharges assumptions from state in the order they appear in goal; + checks (if requested) that resulting theorem is equivalent to goal. *) + fun mkresult (check,state) = + let val ngoals = nprems_of state + val th = implies_intr_list cAs state + val {hyps,prop,sign=sign',...} = rep_thm th + in if not check then standard th + else if not (eq_sg(sign,sign')) then result_error state + ("Signature of proof state has changed!" ^ + sign_error (sign,sign')) + else if ngoals>0 then result_error state + (string_of_int ngoals ^ " unsolved goals!") + else if not (null hyps) then result_error state + ("Additional hypotheses:\n" ^ + cat_lines (map (Sign.string_of_term sign) hyps)) + else if Pattern.eta_matches (#tsig(Sign.rep_sg sign)) + (Sign.term_of chorn, prop) + then standard th + else result_error state "proved a different theorem" + end + in + if eq_sg(sign, #sign(rep_thm st0)) + then (prems, st0, mkresult) + else error ("Definitions would change the proof state's signature" ^ + sign_error (sign, #sign(rep_thm st0))) + end + handle THM(s,_,_) => error("prepare_proof: exception THM was raised!\n" ^ s); + +(*Prints exceptions readably to users*) +fun print_sign_exn sign e = + case e of + THM (msg,i,thms) => + (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg); + seq print_thm thms) + | THEORY (msg,thys) => + (writeln ("Exception THEORY raised:\n" ^ msg); + seq print_theory thys) + | TERM (msg,ts) => + (writeln ("Exception TERM raised:\n" ^ msg); + seq (writeln o Sign.string_of_term sign) ts) + | TYPE (msg,Ts,ts) => + (writeln ("Exception TYPE raised:\n" ^ msg); + seq (writeln o Sign.string_of_typ sign) Ts; + seq (writeln o Sign.string_of_term sign) ts) + | e => raise e; + +(** the prove_goal.... commands + Prove theorem using the listed tactics; check it has the specified form. + Augment signature with all type assignments of goal. + Syntax is similar to "goal" command for easy keyboard use. **) + +(*Version taking the goal as a cterm*) +fun prove_goalw_cterm rths chorn tacsf = + let val (prems, st0, mkresult) = prepare_proof rths chorn + val tac = EVERY (tacsf prems) + fun statef() = + (case Sequence.pull (tapply(tac,st0)) of + Some(st,_) => st + | _ => error ("prove_goal: tactic failed")) + in mkresult (true, cond_timeit (!proof_timing) statef) end; + +(*Version taking the goal as a string*) +fun prove_goalw thy rths agoal tacsf = + let val sign = sign_of thy + val chorn = Sign.read_cterm sign (agoal,propT) + in prove_goalw_cterm rths chorn tacsf + handle ERROR => error (*from type_assign, etc via prepare_proof*) + ("The error above occurred for " ^ agoal) + | e => (print_sign_exn sign e; (*other exceptions*) + error ("The exception above was raised for " ^ agoal)) + end; + +(*String version with no meta-rewrite-rules*) +fun prove_goal thy = prove_goalw thy []; + + +(*** Commands etc ***) + +(*Return the current goal stack, if any, from undo_list*) +fun getstate() : gstack = case !undo_list of + [] => error"No current state in subgoal module" + | x::_ => x; + +(*Pops the given goal stack*) +fun pop [] = error"Cannot go back past the beginning of the proof!" + | pop (pair::pairs) = (pair,pairs); + + +(*Print a level of the goal stack.*) +fun print_top ((th,_), pairs) = + (prs("Level " ^ string_of_int(length pairs) ^ "\n"); + print_goals (!goals_limit) th); + +(*Printing can raise exceptions, so the assignment occurs last. + Can do setstate[(st,Sequence.null)] to set st as the state. *) +fun setstate newgoals = + (print_top (pop newgoals); undo_list := newgoals :: !undo_list); + +(*Given a proof state transformation, return a command that updates + the goal stack*) +fun make_command com = setstate (com (pop (getstate()))); + +(*Apply a function on proof states to the current goal stack*) +fun apply_fun f = f (pop(getstate())); + +(*Return the top theorem, representing the proof state*) +fun topthm () = apply_fun (fn ((th,_), _) => th); + +(*Return the final result. *) +fun result () = !curr_mkresult (true, topthm()); + +(*Return the result UNCHECKED that it equals the goal -- for synthesis, + answer extraction, or other instantiation of Vars *) +fun uresult () = !curr_mkresult (false, topthm()); + +(*Get subgoal i from goal stack*) +fun getgoal i = + (case drop (i-1, prems_of (topthm())) of + [] => error"getgoal: Goal number out of range" + | Q::_ => Q); + +(*Return subgoal i's hypotheses as meta-level assumptions. + For debugging uses of METAHYPS*) +local exception GETHYPS of thm list +in +fun gethyps i = + (tapply(METAHYPS (fn hyps => raise (GETHYPS hyps)) i, topthm()); []) + handle GETHYPS hyps => hyps +end; + +(*Which thms could apply to goal i? (debugs tactics involving filter_thms) *) +fun filter_goal could ths i = filter_thms could (999, getgoal i, ths); + +(*For inspecting earlier levels of the backward proof*) +fun chop_level n (pair,pairs) = + let val level = length pairs + in if 0<=n andalso n<= level + then drop (level - n, pair::pairs) + else error ("Level number must lie between 0 and " ^ + string_of_int level) + end; + +(*Print the given level of the proof*) +fun prlev n = apply_fun (print_top o pop o (chop_level n)); +fun pr () = apply_fun print_top; + +(** the goal.... commands + Read main goal. Set global variables curr_prems, curr_mkresult. + Initial subgoal and premises are rewritten using rths. **) + +(*Version taking the goal as a cterm; if you have a term t and theory thy, use + goalw_cterm rths (Sign.cterm_of (sign_of thy) t); *) +fun goalw_cterm rths chorn = + let val (prems, st0, mkresult) = prepare_proof rths chorn + in undo_list := []; + setstate [ (st0, Sequence.null) ]; + curr_prems := prems; + curr_mkresult := mkresult; + prems + end; + +(*Version taking the goal as a string*) +fun goalw thy rths agoal = + goalw_cterm rths (Sign.read_cterm(sign_of thy)(agoal,propT)) + handle ERROR => error (*from type_assign, etc via prepare_proof*) + ("The error above occurred for " ^ agoal); + +(*String version with no meta-rewrite-rules*) +fun goal thy = goalw thy []; + +(*Proof step "by" the given tactic -- apply tactic to the proof state*) +fun by_com tac ((th,ths), pairs) : gstack = + (case Sequence.pull(tapply(tac, th)) of + None => error"by: tactic failed" + | Some(th2,ths2) => + (if eq_thm(th,th2) + then writeln"Warning: same as previous level" + else if eq_thm_sg(th,th2) then () + else writeln("Warning: signature of proof state has changed" ^ + sign_error (#sign(rep_thm th), #sign(rep_thm th2))); + ((th2,ths2)::(th,ths)::pairs))); + +fun by tac = cond_timeit (!proof_timing) + (fn() => make_command (by_com tac)); + +(* byev[tac1,...,tacn] applies tac1 THEN ... THEN tacn. + Good for debugging proofs involving prove_goal.*) +val byev = by o EVERY; + + +(*Backtracking means find an alternative result from a tactic. + If none at this level, try earlier levels*) +fun backtrack [] = error"back: no alternatives" + | backtrack ((th,thstr) :: pairs) = + (case Sequence.pull thstr of + None => (writeln"Going back a level..."; backtrack pairs) + | Some(th2,thstr2) => + (if eq_thm(th,th2) + then writeln"Warning: same as previous choice at this level" + else if eq_thm_sg(th,th2) then () + else writeln"Warning: signature of proof state has changed"; + (th2,thstr2)::pairs)); + +fun back() = setstate (backtrack (getstate())); + +(*Chop back to previous level of the proof*) +fun choplev n = make_command (chop_level n); + +(*Chopping back the goal stack*) +fun chop () = make_command (fn (_,pairs) => pairs); + +(*Restore the previous proof state; discard current state. *) +fun undo() = case !undo_list of + [] => error"No proof state" + | [_] => error"Already at initial state" + | _::newundo => (undo_list := newundo; pr()) ; + + +(*** Managing the proof stack ***) + +fun save_proof() = Proof(!undo_list, !curr_prems, !curr_mkresult); + +fun restore_proof(Proof(ul,prems,mk)) = + (undo_list:= ul; curr_prems:= prems; curr_mkresult := mk; prems); + + +fun top_proof() = case !proofstack of + [] => error("Stack of proof attempts is empty!") + | p::ps => (p,ps); + +(* push a copy of the current proof state on to the stack *) +fun push_proof() = (proofstack := (save_proof() :: !proofstack)); + +(* discard the top proof state of the stack *) +fun pop_proof() = + let val (p,ps) = top_proof() + val prems = restore_proof p + in proofstack := ps; pr(); prems end; + +(* rotate the stack so that the top element goes to the bottom *) +fun rotate_proof() = let val (p,ps) = top_proof() + in proofstack := ps@[save_proof()]; + restore_proof p; + pr(); + !curr_prems + end; + + +(** Shortcuts for commonly-used tactics **) + +fun bws rls = by (rewrite_goals_tac rls); +fun bw rl = bws [rl]; + +fun brs rls i = by (resolve_tac rls i); +fun br rl = brs [rl]; + +fun bes rls i = by (eresolve_tac rls i); +fun be rl = bes [rl]; + +fun bds rls i = by (dresolve_tac rls i); +fun bd rl = bds [rl]; + +fun ba i = by (assume_tac i); + +fun ren str i = by (rename_tac str i); + +(** Shortcuts to work on the first applicable subgoal **) + +fun frs rls = by (FIRSTGOAL (trace_goalno_tac (resolve_tac rls))); +fun fr rl = frs [rl]; + +fun fes rls = by (FIRSTGOAL (trace_goalno_tac (eresolve_tac rls))); +fun fe rl = fes [rl]; + +fun fds rls = by (FIRSTGOAL (trace_goalno_tac (dresolve_tac rls))); +fun fd rl = fds [rl]; + +fun fa() = by (FIRSTGOAL (trace_goalno_tac assume_tac)); + +(** Reading and printing terms wrt the current theory **) + +fun top_sg() = #sign(rep_thm(topthm())); + +fun read s = Sign.term_of (Sign.read_cterm (top_sg()) + (s, (TVar(("DUMMY",0),[])))); + +(*Print a term under the current signature of the proof state*) +fun prin t = writeln (Sign.string_of_term (top_sg()) t); + +fun printyp T = writeln (Sign.string_of_typ (top_sg()) T); + +fun pprint_term t = Sign.pprint_term (top_sg()) t; + +fun pprint_typ T = Sign.pprint_typ (top_sg()) T; + +(*Prints exceptions nicely at top level; + raises the exception in order to have a polymorphic type!*) +fun print_exn e = (print_sign_exn (top_sg()) e; raise e); + +end; +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/install_pp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/install_pp.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,15 @@ +(* Title: Pure/install_pp + ID: $Id$ + +Set up automatic toplevel printing +*) + +install_pp (make_pp ["Thm", "thm"] pprint_thm); +install_pp (make_pp ["Thm", "theory"] pprint_theory); +install_pp (make_pp ["Sign", "sg"] pprint_sg); +install_pp (make_pp ["term"] pprint_term); +install_pp (make_pp ["Sign", "cterm"] Sign.pprint_cterm); +install_pp (make_pp ["typ"] pprint_typ); +install_pp (make_pp ["Sign", "ctyp"] Sign.pprint_ctyp); +install_pp (make_pp ["Ast", "ast"] Syntax.pprint_ast); + diff -r 000000000000 -r a5a9c433f639 src/Pure/library.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/library.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,585 @@ +(* Title: library + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Basic library: booleans, lists, pairs, input/output, etc. +*) + + +(**** Booleans: operators for combining predicates ****) + +infix orf; +fun p orf q = fn x => p x orelse q x ; + +infix andf; +fun p andf q = fn x => p x andalso q x ; + +fun notf p x = not (p x) ; + +fun orl [] = false + | orl (x::l) = x orelse orl l; + +fun andl [] = true + | andl (x::l) = x andalso andl l; + +(*exists pred [x1,...,xn] ======> pred(x1) orelse ... orelse pred(xn)*) +fun exists (pred: 'a -> bool) : 'a list -> bool = + let fun boolf [] = false + | boolf (x::l) = (pred x) orelse boolf l + in boolf end; + +(*forall pred [x1,...,xn] ======> pred(x1) andalso ... andalso pred(xn)*) +fun forall (pred: 'a -> bool) : 'a list -> bool = + let fun boolf [] = true + | boolf (x::l) = (pred x) andalso (boolf l) + in boolf end; + + +(*** Lists ***) + +exception LIST of string; + +(*discriminator and selectors for lists. *) +fun null [] = true + | null (_::_) = false; + +fun hd [] = raise LIST "hd" + | hd (a::_) = a; + +fun tl [] = raise LIST "tl" + | tl (_::l) = l; + + +(*curried functions for pairing and reversed pairing*) +fun pair x y = (x,y); +fun rpair x y = (y,x); + +fun fst(x,y) = x and snd(x,y) = y; + +(*Handy combinators*) +fun curry f x y = f(x,y); +fun uncurry f(x,y) = f x y; +fun I x = x and K x y = x; + +(*Combine two functions forming the union of their domains*) +infix orelf; +fun f orelf g = fn x => f x handle Match=> g x; + + +(*Application of (infix) operator to its left or right argument*) +fun apl (x,f) y = f(x,y); +fun apr (f,y) x = f(x,y); + + +(*functional for pairs*) +fun pairself f (x,y) = (f x, f y); + +(*Apply the function to a component of a pair*) +fun apfst f (x, y) = (f x, y); +fun apsnd f (x, y) = (x, f y); + +fun square (n: int) = n*n; + +fun fact 0 = 1 + | fact n = n * fact(n-1); + + +(*The following versions of fold are designed to fit nicely with infixes.*) + +(* (op @) (e, [x1,...,xn]) ======> ((e @ x1) @ x2) ... @ xn + for operators that associate to the left. TAIL RECURSIVE*) +fun foldl (f: 'a * 'b -> 'a) : 'a * 'b list -> 'a = + let fun itl (e, []) = e + | itl (e, a::l) = itl (f(e,a), l) + in itl end; + +(* (op @) ([x1,...,xn], e) ======> x1 @ (x2 ... @ (xn @ e)) + for operators that associate to the right. Not tail recursive.*) +fun foldr f (l,e) = + let fun itr [] = e + | itr (a::l) = f(a, itr l) + in itr l end; + +(* (op @) [x1,...,xn] ======> x1 @ (x2 ..(x[n-1]. @ xn)) + for n>0, operators that associate to the right. Not tail recursive.*) +fun foldr1 f l = + let fun itr [x] = x + | itr (x::l) = f(x, itr l) + in itr l end; + + +(*Length of a list. Should unquestionably be a standard function*) +local fun length1 (n, [ ]) = n (*TAIL RECURSIVE*) + | length1 (n, x::l) = length1 (n+1, l) +in fun length l = length1 (0,l) end; + + +(*Take the first n elements from l.*) +fun take (n, []) = [] + | take (n, x::xs) = if n>0 then x::take(n-1,xs) + else []; + +(*Drop the first n elements from l.*) +fun drop (_, []) = [] + | drop (n, x::xs) = if n>0 then drop (n-1, xs) + else x::xs; + +(*Return nth element of l, where 0 designates the first element; + raise EXCEPTION if list too short.*) +fun nth_elem NL = case (drop NL) of + [] => raise LIST "nth_elem" + | x::l => x; + + +(*make the list [from, from+1, ..., to]*) +infix upto; +fun from upto to = + if from>to then [] else from :: ((from+1) upto to); + +(*make the list [from, from-1, ..., to]*) +infix downto; +fun from downto to = + if from is = [n,n-1,...,0] *) +fun downto0(i::is,n) = i=n andalso downto0(is,n-1) + | downto0([],n) = n = ~1; + +(*Like Lisp's MAPC -- seq proc [x1,...,xn] evaluates + proc(x1); ... ; proc(xn) for side effects.*) +fun seq (proc: 'a -> unit) : 'a list -> unit = + let fun seqf [] = () + | seqf (x::l) = (proc x; seqf l) + in seqf end; + + +(*** Balanced folding; access to balanced trees ***) + +exception Balance; (*indicates non-positive argument to balancing fun*) + +(*Balanced folding; avoids deep nesting*) +fun fold_bal f [x] = x + | fold_bal f [] = raise Balance + | fold_bal f xs = + let val k = length xs div 2 + in f (fold_bal f (take(k,xs)), + fold_bal f (drop(k,xs))) + end; + +(*Construct something of the form f(...g(...(x)...)) for balanced access*) +fun access_bal (f,g,x) n i = + let fun acc n i = (* 1<=i<=n*) + if n=1 then x else + let val n2 = n div 2 + in if i<=n2 then f (acc n2 i) + else g (acc (n-n2) (i-n2)) + end + in if 1<=i andalso i<=n then acc n i else raise Balance end; + +(*Construct ALL such accesses; could try harder to share recursive calls!*) +fun accesses_bal (f,g,x) n = + let fun acc n = + if n=1 then [x] else + let val n2 = n div 2 + val acc2 = acc n2 + in if n-n2=n2 then map f acc2 @ map g acc2 + else map f acc2 @ map g (acc (n-n2)) end + in if 1<=n then acc n else raise Balance end; + + +(*** Input/Output ***) + +fun prs s = output(std_out,s); +fun writeln s = prs (s ^ "\n"); + +(*Print error message and abort to top level*) +exception ERROR; +fun error (msg) = (writeln msg; raise ERROR); + +fun assert p msg = if p then () else error msg; +fun deny p msg = if p then error msg else (); + +(*For the "test" target in Makefiles -- signifies successful termination*) +fun maketest msg = + (writeln msg; + output(open_out "test", "Test examples ran successfully\n")); + +(*print a list surrounded by the brackets lpar and rpar, with comma separator + print nothing for empty list*) +fun print_list (lpar, rpar, pre: 'a -> unit) (l : 'a list) = + let fun prec(x) = (prs","; pre(x)) + in case l of + [] => () + | x::l => (prs lpar; pre x; seq prec l; prs rpar) + end; + +(*print a list of items separated by newlines*) +fun print_list_ln (pre: 'a -> unit) : 'a list -> unit = + seq (fn x => (pre x; writeln"")); + +fun is_letter ch = + (ord"A" <= ord ch) andalso (ord ch <= ord"Z") orelse + (ord"a" <= ord ch) andalso (ord ch <= ord"z"); + +fun is_digit ch = + (ord"0" <= ord ch) andalso (ord ch <= ord"9"); + +(*letter or _ or prime (') *) +fun is_quasi_letter "_" = true + | is_quasi_letter "'" = true + | is_quasi_letter ch = is_letter ch; + +(*white space: blanks, tabs, newlines*) +val is_blank : string -> bool = fn + " " => true | "\t" => true | "\n" => true | _ => false; + +val is_letdig = is_quasi_letter orf is_digit; + +val to_lower = + let + fun lower ch = + if ch >= "A" andalso ch <= "Z" then + chr (ord ch - ord "A" + ord "a") + else ch; + in + implode o (map lower) o explode + end; + + +(*** Timing ***) + +(*Unconditional timing function*) +val timeit = cond_timeit true; + +(*Timed application function*) +fun timeap f x = timeit(fn()=> f x); + +(*Timed "use" function, printing filenames*) +fun time_use fname = timeit(fn()=> + (writeln("\n**** Starting " ^ fname ^ " ****"); use fname; + writeln("\n**** Finished " ^ fname ^ " ****"))); + + +(*** Misc functions ***) + +(*Function exponentiation: f(...(f x)...) with n applications of f *) +fun funpow n f x = + let fun rep (0,x) = x + | rep (n,x) = rep (n-1, f x) + in rep (n,x) end; + +(*Combine two lists forming a list of pairs: + [x1,...,xn] ~~ [y1,...,yn] ======> [(x1,y1), ..., (xn,yn)] *) +infix ~~; +fun [] ~~ [] = [] + | (x::xs) ~~ (y::ys) = (x,y) :: (xs ~~ ys) + | _ ~~ _ = raise LIST "~~"; + +(*Inverse of ~~; the old 'split'. + [(x1,y1), ..., (xn,yn)] ======> ( [x1,...,xn] , [y1,...,yn] ) *) +fun split_list (l: ('a*'b)list) = (map #1 l, map #2 l); + +(*make the list [x; x; ...; x] of length n*) +fun replicate n (x: 'a) : 'a list = + let fun rep (0,xs) = xs + | rep (n,xs) = rep(n-1, x::xs) + in if n<0 then raise LIST "replicate" + else rep (n,[]) + end; + +(*Flatten a list of lists to a list.*) +fun flat (ls: 'c list list) : 'c list = foldr (op @) (ls,[]); + + +(*** polymorphic set operations ***) + +(*membership in a list*) +infix mem; +fun x mem [] = false + | x mem (y::l) = (x=y) orelse (x mem l); + +(*insertion into list if not already there*) +infix ins; +fun x ins xs = if x mem xs then xs else x::xs; + +(*union of sets represented as lists: no repetitions*) +infix union; +fun xs union [] = xs + | [] union ys = ys + | (x::xs) union ys = xs union (x ins ys); + +infix inter; +fun [] inter ys = [] + | (x::xs) inter ys = if x mem ys then x::(xs inter ys) + else xs inter ys; + +infix subset; +fun [] subset ys = true + | (x::xs) subset ys = x mem ys andalso xs subset ys; + +(*removing an element from a list WITHOUT duplicates*) +infix \; +fun (y::ys) \ x = if x=y then ys else y::(ys \ x) + | [] \ x = []; + +infix \\; +val op \\ = foldl (op \); + +(*** option stuff ***) + +datatype 'a option = None | Some of 'a; + +exception OPTION of string; + +fun the (Some x) = x + | the None = raise OPTION "the"; + +fun is_some (Some _) = true + | is_some None = false; + +fun is_none (Some _) = false + | is_none None = true; + + +(*** Association lists ***) + +(*Association list lookup*) +fun assoc ([], key) = None + | assoc ((keyi,xi)::pairs, key) = + if key=keyi then Some xi else assoc (pairs,key); + +fun assocs ps x = case assoc(ps,x) of None => [] | Some(ys) => ys; + +(*Association list update*) +fun overwrite(al,p as (key,_)) = + let fun over((q as (keyi,_))::pairs) = + if keyi=key then p::pairs else q::(over pairs) + | over[] = [p] + in over al end; + +(*Copy the list preserving elements that satisfy the predicate*) +fun filter (pred: 'a->bool) : 'a list -> 'a list = + let fun filt [] = [] + | filt (x::xs) = if pred(x) then x :: filt xs else filt xs + in filt end; + +fun filter_out f = filter (not o f); + + +(*** List operations, generalized to an arbitrary equality function "eq" + -- so what good are equality types?? ***) + +(*removing an element from a list -- possibly WITH duplicates*) +fun gen_rem eq (xs,y) = filter_out (fn x => eq(x,y)) xs; + +(*generalized membership test*) +fun gen_mem eq (x, []) = false + | gen_mem eq (x, y::ys) = eq(x,y) orelse gen_mem eq (x,ys); + +(*generalized insertion*) +fun gen_ins eq (x,xs) = if gen_mem eq (x,xs) then xs else x::xs; + +(*generalized union*) +fun gen_union eq (xs,[]) = xs + | gen_union eq ([],ys) = ys + | gen_union eq (x::xs,ys) = gen_union eq (xs, gen_ins eq (x,ys)); + +(*Generalized association list lookup*) +fun gen_assoc eq ([], key) = None + | gen_assoc eq ((keyi,xi)::pairs, key) = + if eq(key,keyi) then Some xi else gen_assoc eq (pairs,key); + +(** Finding list elements and duplicates **) + +(* find the position of an element in a list *) +fun find(x,ys) = + let fun f(y::ys,i) = if x=y then i else f(ys,i+1) + | f(_,_) = raise LIST "find" + in f(ys,0) end; + +(*Returns the tail beginning with the first repeated element, or []. *) +fun findrep [] = [] + | findrep (x::xs) = if x mem xs then x::xs else findrep xs; + +fun distinct1 (seen, []) = rev seen + | distinct1 (seen, x::xs) = + if x mem seen then distinct1 (seen, xs) + else distinct1 (x::seen, xs); + +(*Makes a list of the distinct members of the input*) +fun distinct xs = distinct1([],xs); + + +(*Use the keyfun to make a list of (x,key) pairs.*) +fun make_keylist (keyfun: 'a->'b) : 'a list -> ('a * 'b) list = + let fun keypair x = (x, keyfun x) + in map keypair end; + +(*Given a list of (x,key) pairs and a searchkey + return the list of xs from each pair whose key equals searchkey*) +fun keyfilter [] searchkey = [] + | keyfilter ((x,key)::pairs) searchkey = + if key=searchkey then x :: keyfilter pairs searchkey + else keyfilter pairs searchkey; + +fun mapfilter (f: 'a -> 'b option) ([]: 'a list) = [] : 'b list + | mapfilter f (x::xs) = + case (f x) of + None => mapfilter f xs + | Some y => y :: mapfilter f xs; + + +(*Partition list into elements that satisfy predicate and those that don't. + Preserves order of elements in both lists. *) +fun partition (pred: 'a->bool) (ys: 'a list) : ('a list * 'a list) = + let fun part ([], answer) = answer + | part (x::xs, (ys, ns)) = if pred(x) + then part (xs, (x::ys, ns)) + else part (xs, (ys, x::ns)) + in part (rev ys, ([],[])) end; + + +fun partition_eq (eq:'a * 'a -> bool) = + let fun part [] = [] + | part (x::ys) = let val (xs,xs') = partition (apl(x,eq)) ys + in (x::xs)::(part xs') end + in part end; + + +(*Partition a list into buckets [ bi, b(i+1),...,bj ] + putting x in bk if p(k)(x) holds. Preserve order of elements if possible.*) +fun partition_list p i j = + let fun part k xs = + if k>j then + (case xs of [] => [] + | _ => raise LIST "partition_list") + else + let val (ns,rest) = partition (p k) xs; + in ns :: part(k+1)rest end + in part i end; + + +(*Insertion sort. Stable (does not reorder equal elements) + 'less' is less-than test on type 'a. *) +fun sort (less: 'a*'a -> bool) = + let fun insert (x, []) = [x] + | insert (x, y::ys) = + if less(y,x) then y :: insert (x,ys) else x::y::ys; + fun sort1 [] = [] + | sort1 (x::xs) = insert (x, sort1 xs) + in sort1 end; + +(*Transitive Closure. Not Warshall's algorithm*) +fun transitive_closure [] = [] + | transitive_closure ((x,ys)::ps) = + let val qs = transitive_closure ps + val zs = foldl (fn (zs,y) => assocs qs y union zs) (ys,ys) + fun step(u,us) = (u, if x mem us then zs union us else us) + in (x,zs) :: map step qs end; + +(*** Converting integers to strings, generating identifiers, etc. ***) + +(*Expand the number in the given base + example: radixpand(2, 8) gives [1, 0, 0, 0] *) +fun radixpand (base,num) : int list = + let fun radix (n,tail) = + if nn then max(m::ns) else max(n::ns) + | max [] = raise LIST "max"; + +fun min[m : int] = m + | min(m::n::ns) = if m ([x1,...,x(i-1)], [xi,..., xn]) + where xi is the first element that does not satisfy the predicate*) +fun take_prefix (pred : 'a -> bool) (xs: 'a list) : 'a list * 'a list = + let fun take (rxs, []) = (rev rxs, []) + | take (rxs, x::xs) = + if pred x then take(x::rxs, xs) else (rev rxs, x::xs) + in take([],xs) end; + +infix prefix; +fun [] prefix _ = true + | (x::xs) prefix (y::ys) = (x=y) andalso (xs prefix ys) + | _ prefix _ = false; + +(* [x1, x2, ..., xn] ---> [x1, s, x2, s, ..., s, xn] *) +fun separate s (x :: (xs as _ :: _)) = x :: s :: separate s xs + | separate _ xs = xs; + +(*space_implode "..." (explode "hello"); gives "h...e...l...l...o" *) +fun space_implode a bs = implode (separate a bs); + +fun quote s = "\"" ^ s ^ "\""; + +(*Concatenate messages, one per line, into a string*) +val cat_lines = implode o (map (apr(op^,"\n"))); + +(*Scan a list of characters into "words" composed of "letters" (recognized + by is_let) and separated by any number of non-"letters".*) +fun scanwords is_let cs = + let fun scan1 [] = [] + | scan1 cs = + let val (lets, rest) = take_prefix is_let cs + in implode lets :: scanwords is_let rest end; + in scan1 (#2 (take_prefix (not o is_let) cs)) end; diff -r 000000000000 -r a5a9c433f639 src/Pure/logic.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/logic.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,298 @@ +(* Title: logic + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright Cambridge University 1992 + +Supporting code for defining the abstract type "thm" +*) + +infix occs; + +signature LOGIC = + sig + val assum_pairs: term -> (term*term)list + val auto_rename: bool ref + val close_form: term -> term + val count_prems: term * int -> int + val dest_equals: term -> term * term + val dest_flexpair: term -> term * term + val dest_implies: term -> term * term + val flatten_params: int -> term -> term + val freeze_vars: term -> term + val incr_indexes: typ list * int -> term -> term + val lift_fns: term * int -> (term -> term) * (term -> term) + val list_flexpairs: (term*term)list * term -> term + val list_implies: term list * term -> term + val list_rename_params: string list * term -> term + val mk_equals: term * term -> term + val mk_flexpair: term * term -> term + val mk_implies: term * term -> term + val occs: term * term -> bool + val rule_of: (term*term)list * term list * term -> term + val set_rename_prefix: string -> unit + val skip_flexpairs: term -> term + val strip_assums_concl: term -> term + val strip_assums_hyp: term -> term list + val strip_flexpairs: term -> (term*term)list * term + val strip_horn: term -> (term*term)list * term list * term + val strip_imp_concl: term -> term + val strip_imp_prems: term -> term list + val strip_params: term -> (string * typ) list + val strip_prems: int * term list * term -> term list * term + val thaw_vars: term -> term + val varify: term -> term + end; + +functor LogicFun (structure Unify: UNIFY and Net:NET) : LOGIC = +struct +structure Type = Unify.Sign.Type; + +(*** Abstract syntax operations on the meta-connectives ***) + +(** equality **) + +(*Make an equality. DOES NOT CHECK TYPE OF u! *) +fun mk_equals(t,u) = equals(type_of t) $ t $ u; + +fun dest_equals (Const("==",_) $ t $ u) = (t,u) + | dest_equals t = raise TERM("dest_equals", [t]); + +(** implies **) + +fun mk_implies(A,B) = implies $ A $ B; + +fun dest_implies (Const("==>",_) $ A $ B) = (A,B) + | dest_implies A = raise TERM("dest_implies", [A]); + +(** nested implications **) + +(* [A1,...,An], B goes to A1==>...An==>B *) +fun list_implies ([], B) = B : term + | list_implies (A::AS, B) = implies $ A $ list_implies(AS,B); + +(* A1==>...An==>B goes to [A1,...,An], where B is not an implication *) +fun strip_imp_prems (Const("==>", _) $ A $ B) = A :: strip_imp_prems B + | strip_imp_prems _ = []; + +(* A1==>...An==>B goes to B, where B is not an implication *) +fun strip_imp_concl (Const("==>", _) $ A $ B) = strip_imp_concl B + | strip_imp_concl A = A : term; + +(*Strip and return premises: (i, [], A1==>...Ai==>B) + goes to ([Ai, A(i-1),...,A1] , B) (REVERSED) + if i<0 or else i too big then raises TERM*) +fun strip_prems (0, As, B) = (As, B) + | strip_prems (i, As, Const("==>", _) $ A $ B) = + strip_prems (i-1, A::As, B) + | strip_prems (_, As, A) = raise TERM("strip_prems", A::As); + +(*Count premises -- quicker than (length ostrip_prems) *) +fun count_prems (Const("==>", _) $ A $ B, n) = count_prems (B,n+1) + | count_prems (_,n) = n; + +(** flex-flex constraints **) + +(*Make a constraint. DOES NOT CHECK TYPE OF u! *) +fun mk_flexpair(t,u) = flexpair(type_of t) $ t $ u; + +fun dest_flexpair (Const("=?=",_) $ t $ u) = (t,u) + | dest_flexpair t = raise TERM("dest_flexpair", [t]); + +(*make flexflex antecedents: ( [(a1,b1),...,(an,bn)] , C ) + goes to (a1=?=b1) ==>...(an=?=bn)==>C *) +fun list_flexpairs ([], A) = A + | list_flexpairs ((t,u)::pairs, A) = + implies $ (mk_flexpair(t,u)) $ list_flexpairs(pairs,A); + +(*Make the object-rule tpairs==>As==>B *) +fun rule_of (tpairs, As, B) = list_flexpairs(tpairs, list_implies(As, B)); + +(*Remove and return flexflex pairs: + (a1=?=b1)==>...(an=?=bn)==>C to ( [(a1,b1),...,(an,bn)] , C ) + [Tail recursive in order to return a pair of results] *) +fun strip_flex_aux (pairs, Const("==>", _) $ (Const("=?=",_)$t$u) $ C) = + strip_flex_aux ((t,u)::pairs, C) + | strip_flex_aux (pairs,C) = (rev pairs, C); + +fun strip_flexpairs A = strip_flex_aux([], A); + +(*Discard flexflex pairs*) +fun skip_flexpairs (Const("==>", _) $ (Const("=?=",_)$_$_) $ C) = + skip_flexpairs C + | skip_flexpairs C = C; + +(*strip a proof state (Horn clause): + (a1==b1)==>...(am==bm)==>B1==>...Bn==>C + goes to ( [(a1,b1),...,(am,bm)] , [B1,...,Bn] , C) *) +fun strip_horn A = + let val (tpairs,horn) = strip_flexpairs A + in (tpairs, strip_imp_prems horn, strip_imp_concl horn) end; + + +(*** Low-level term operations ***) + +(*Does t occur in u? Or is alpha-convertible to u? + The term t must contain no loose bound variables*) +fun t occs u = (t aconv u) orelse + (case u of + Abs(_,_,body) => t occs body + | f$t' => t occs f orelse t occs t' + | _ => false); + +(*Close up a formula over all free variables by quantification*) +fun close_form A = + list_all_free (map dest_Free (sort atless (term_frees A)), + A); + + +(*Freeze all (T)Vars by turning them into (T)Frees*) +fun freeze_vars(Var(ixn,T)) = Free(Syntax.string_of_vname ixn, + Type.freeze_vars T) + | freeze_vars(Const(a,T)) = Const(a,Type.freeze_vars T) + | freeze_vars(Free(a,T)) = Free(a,Type.freeze_vars T) + | freeze_vars(s$t) = freeze_vars s $ freeze_vars t + | freeze_vars(Abs(a,T,t)) = Abs(a,Type.freeze_vars T,freeze_vars t) + | freeze_vars(b) = b; + +(*Reverse the effect of freeze_vars*) +fun thaw_vars(Const(a,T)) = Const(a,Type.thaw_vars T) + | thaw_vars(Free(a,T)) = + let val T' = Type.thaw_vars T + in case explode a of + "?"::vn => let val (ixn,_) = Syntax.scan_varname vn + in Var(ixn,T') end + | _ => Free(a,T') + end + | thaw_vars(Abs(a,T,t)) = Abs(a,Type.thaw_vars T, thaw_vars t) + | thaw_vars(s$t) = thaw_vars s $ thaw_vars t + | thaw_vars(b) = b; + + +(*** Specialized operations for resolution... ***) + +(*For all variables in the term, increment indexnames and lift over the Us + result is ?Gidx(B.(lev+n-1),...,B.lev) where lev is abstraction level *) +fun incr_indexes (Us: typ list, inc:int) t = + let fun incr (Var ((a,i), T), lev) = + Unify.combound (Var((a, i+inc), Us---> incr_tvar inc T), + lev, length Us) + | incr (Abs (a,T,body), lev) = + Abs (a, incr_tvar inc T, incr(body,lev+1)) + | incr (Const(a,T),_) = Const(a, incr_tvar inc T) + | incr (Free(a,T),_) = Free(a, incr_tvar inc T) + | incr (f$t, lev) = incr(f,lev) $ incr(t,lev) + | incr (t,lev) = t + in incr(t,0) end; + +(*Make lifting functions from subgoal and increment. + lift_abs operates on tpairs (unification constraints) + lift_all operates on propositions *) +fun lift_fns (B,inc) = + let fun lift_abs (Us, Const("==>", _) $ _ $ B) u = lift_abs (Us,B) u + | lift_abs (Us, Const("all",_)$Abs(a,T,t)) u = + Abs(a, T, lift_abs (T::Us, t) u) + | lift_abs (Us, _) u = incr_indexes(rev Us, inc) u + fun lift_all (Us, Const("==>", _) $ A $ B) u = + implies $ A $ lift_all (Us,B) u + | lift_all (Us, Const("all",_)$Abs(a,T,t)) u = + all T $ Abs(a, T, lift_all (T::Us,t) u) + | lift_all (Us, _) u = incr_indexes(rev Us, inc) u; + in (lift_abs([],B), lift_all([],B)) end; + +(*Strips assumptions in goal, yielding list of hypotheses. *) +fun strip_assums_hyp (Const("==>", _) $ H $ B) = H :: strip_assums_hyp B + | strip_assums_hyp (Const("all",_)$Abs(a,T,t)) = strip_assums_hyp t + | strip_assums_hyp B = []; + +(*Strips assumptions in goal, yielding conclusion. *) +fun strip_assums_concl (Const("==>", _) $ H $ B) = strip_assums_concl B + | strip_assums_concl (Const("all",_)$Abs(a,T,t)) = strip_assums_concl t + | strip_assums_concl B = B; + +(*Make a list of all the parameters in a subgoal, even if nested*) +fun strip_params (Const("==>", _) $ H $ B) = strip_params B + | strip_params (Const("all",_)$Abs(a,T,t)) = (a,T) :: strip_params t + | strip_params B = []; + +(*Removes the parameters from a subgoal and renumber bvars in hypotheses, + where j is the total number of parameters (precomputed) + If n>0 then deletes assumption n. *) +fun remove_params j n A = + if j=0 andalso n<=0 then A (*nothing left to do...*) + else case A of + Const("==>", _) $ H $ B => + if n=1 then (remove_params j (n-1) B) + else implies $ (incr_boundvars j H) $ (remove_params j (n-1) B) + | Const("all",_)$Abs(a,T,t) => remove_params (j-1) n t + | _ => if n>0 then raise TERM("remove_params", [A]) + else A; + +(** Auto-renaming of parameters in subgoals **) + +val auto_rename = ref false +and rename_prefix = ref "ka"; + +(*rename_prefix is not exported; it is set by this function.*) +fun set_rename_prefix a = + if a<>"" andalso forall is_letter (explode a) + then (rename_prefix := a; auto_rename := true) + else error"rename prefix must be nonempty and consist of letters"; + +(*Makes parameters in a goal have distinctive names (not guaranteed unique!) + A name clash could cause the printer to rename bound vars; + then res_inst_tac would not work properly.*) +fun rename_vars (a, []) = [] + | rename_vars (a, (_,T)::vars) = + (a,T) :: rename_vars (bump_string a, vars); + +(*Move all parameters to the front of the subgoal, renaming them apart; + if n>0 then deletes assumption n. *) +fun flatten_params n A = + let val params = strip_params A; + val vars = if !auto_rename + then rename_vars (!rename_prefix, params) + else variantlist(map #1 params,[]) ~~ map #2 params + in list_all (vars, remove_params (length vars) n A) + end; + +(*Makes parameters in a goal have the names supplied by the list cs.*) +fun list_rename_params (cs, Const("==>", _) $ A $ B) = + implies $ A $ list_rename_params (cs, B) + | list_rename_params (c::cs, Const("all",_)$Abs(_,T,t)) = + all T $ Abs(c, T, list_rename_params (cs, t)) + | list_rename_params (cs, B) = B; + +(*Strips assumptions in goal yielding ( [Hn,...,H1], [xm,...,x1], B ) + where H1,...,Hn are the hypotheses and x1...xm are the parameters. *) +fun strip_assums_aux (Hs, params, Const("==>", _) $ H $ B) = + strip_assums_aux (H::Hs, params, B) + | strip_assums_aux (Hs, params, Const("all",_)$Abs(a,T,t)) = + strip_assums_aux (Hs, (a,T)::params, t) + | strip_assums_aux (Hs, params, B) = (Hs, params, B); + +fun strip_assums A = strip_assums_aux ([],[],A); + + +(*Produces disagreement pairs, one for each assumption proof, in order. + A is the first premise of the lifted rule, and thus has the form + H1 ==> ... Hk ==> B and the pairs are (H1,B),...,(Hk,B) *) +fun assum_pairs A = + let val (Hs, params, B) = strip_assums A + val D = Unify.rlist_abs(params, B) + fun pairrev ([],pairs) = pairs + | pairrev (H::Hs,pairs) = + pairrev(Hs, (Unify.rlist_abs(params,H), D) :: pairs) + in pairrev (Hs,[]) (*WAS: map pair (rev Hs) *) + end; + + +(*Converts Frees to Vars and TFrees to TVars so that axioms can be written + without (?) everywhere*) +fun varify (Const(a,T)) = Const(a, Type.varifyT T) + | varify (Free(a,T)) = Var((a,0), Type.varifyT T) + | varify (Var(ixn,T)) = Var(ixn, Type.varifyT T) + | varify (Abs (a,T,body)) = Abs (a, Type.varifyT T, varify body) + | varify (f$t) = varify f $ varify t + | varify t = t; + +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/net.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/net.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,207 @@ +(* Title: net + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Discrimination nets: a data structure for indexing items + +From the book + E. Charniak, C. K. Riesbeck, D. V. McDermott. + Artificial Intelligence Programming. + (Lawrence Erlbaum Associates, 1980). [Chapter 14] +*) + +signature NET = + sig + type key + type 'a net + exception DELETE and INSERT + val delete: (key list * 'a) * 'a net * ('a*'a -> bool) -> 'a net + val delete_term: (term * 'a) * 'a net * ('a*'a -> bool) -> 'a net + val empty: 'a net + val insert: (key list * 'a) * 'a net * ('a*'a -> bool) -> 'a net + val insert_term: (term * 'a) * 'a net * ('a*'a -> bool) -> 'a net + val lookup: 'a net * key list -> 'a list + val match_term: 'a net -> term -> 'a list + val key_of_term: term -> key list + val unify_term: 'a net -> term -> 'a list + end; + + +functor NetFun () : NET = +struct + +datatype key = CombK | VarK | AtomK of string; + +(*Only 'loose' bound variables could arise, since Abs nodes are skipped*) +fun string_of_bound i = "*B*" ^ chr i; + +(*Keys are preorder lists of symbols -- constants, Vars, bound vars, ... + Any term whose head is a Var is regarded entirely as a Var; + abstractions are also regarded as Vars (to cover eta-conversion) +*) +fun add_key_of_terms (t, cs) = + let fun rands (f$t, cs) = CombK :: rands (f, add_key_of_terms(t, cs)) + | rands (Const(c,_), cs) = AtomK c :: cs + | rands (Free(c,_), cs) = AtomK c :: cs + | rands (Bound i, cs) = AtomK (string_of_bound i) :: cs + in case (head_of t) of + Var _ => VarK :: cs + | Abs (_,_,t) => VarK :: cs + | _ => rands(t,cs) + end; + +(*convert a term to a key -- a list of keys*) +fun key_of_term t = add_key_of_terms (t, []); + + +(*Trees indexed by key lists: each arc is labelled by a key. + Each node contains a list of items, and arcs to children. + Keys in the association list (alist) are stored in ascending order. + The empty key addresses the entire net. + Lookup functions preserve order in items stored at same level. +*) +datatype 'a net = Leaf of 'a list + | Net of {comb: 'a net, + var: 'a net, + alist: (string * 'a net) list}; + +val empty = Leaf[]; +val emptynet = Net{comb=empty, var=empty, alist=[]}; + + +(*** Insertion into a discrimination net ***) + +exception INSERT; (*duplicate item in the net*) + + +(*Adds item x to the list at the node addressed by the keys. + Creates node if not already present. + eq is the equality test for items. + The empty list of keys generates a Leaf node, others a Net node. +*) +fun insert ((keys,x), net, eq) = + let fun ins1 ([], Leaf xs) = + if gen_mem eq (x,xs) then raise INSERT else Leaf(x::xs) + | ins1 (keys, Leaf[]) = ins1 (keys, emptynet) (*expand empty...*) + | ins1 (CombK :: keys, Net{comb,var,alist}) = + Net{comb=ins1(keys,comb), var=var, alist=alist} + | ins1 (VarK :: keys, Net{comb,var,alist}) = + Net{comb=comb, var=ins1(keys,var), alist=alist} + | ins1 (AtomK a :: keys, Net{comb,var,alist}) = + let fun newpair net = (a, ins1(keys,net)) + fun inslist [] = [newpair empty] + | inslist((b: string, netb) :: alist) = + if a=b then newpair netb :: alist + else if ab*) (b,netb) :: inslist alist + in Net{comb=comb, var=var, alist= inslist alist} end + in ins1 (keys,net) end; + +fun insert_term ((t,x), net, eq) = insert((key_of_term t, x), net, eq); + +(*** Deletion from a discrimination net ***) + +exception DELETE; (*missing item in the net*) + +(*Create a new Net node if it would be nonempty*) +fun newnet {comb=Leaf[], var=Leaf[], alist=[]} = empty + | newnet {comb,var,alist} = Net{comb=comb, var=var, alist=alist}; + +(*add new (b,net) pair to the alist provided net is nonempty*) +fun conspair((b, Leaf[]), alist) = alist + | conspair((b, net), alist) = (b, net) :: alist; + +(*Deletes item x from the list at the node addressed by the keys. + Raises DELETE if absent. Collapses the net if possible. + eq is the equality test for items. *) +fun delete ((keys, x), net, eq) = + let fun del1 ([], Leaf xs) = + if gen_mem eq (x,xs) then Leaf (gen_rem eq (xs,x)) + else raise DELETE + | del1 (keys, Leaf[]) = raise DELETE + | del1 (CombK :: keys, Net{comb,var,alist}) = + newnet{comb=del1(keys,comb), var=var, alist=alist} + | del1 (VarK :: keys, Net{comb,var,alist}) = + newnet{comb=comb, var=del1(keys,var), alist=alist} + | del1 (AtomK a :: keys, Net{comb,var,alist}) = + let fun newpair net = (a, del1(keys,net)) + fun dellist [] = raise DELETE + | dellist((b: string, netb) :: alist) = + if a=b then conspair(newpair netb, alist) + else if ab*) (b,netb) :: dellist alist + in newnet{comb=comb, var=var, alist= dellist alist} end + in del1 (keys,net) end; + +fun delete_term ((t,x), net, eq) = delete((key_of_term t, x), net, eq); + +(*** Retrieval functions for discrimination nets ***) + +exception OASSOC; + +(*Ordered association list lookup*) +fun oassoc ([], a: string) = raise OASSOC + | oassoc ((b,x)::pairs, a) = + if b []; + + +(*Skipping a term in a net. Recursively skip 2 levels if a combination*) +fun net_skip (Leaf _, nets) = nets + | net_skip (Net{comb,var,alist}, nets) = + foldr net_skip + (net_skip (comb,[]), + foldr (fn ((_,net), nets) => net::nets) (alist, var::nets)); + +(** Matching and Unification**) + +(*conses the linked net, if present, to nets*) +fun look1 (alist, a) nets = + oassoc(alist,a) :: nets handle OASSOC => nets; + +(*Return the nodes accessible from the term (cons them before nets) + "unif" signifies retrieval for unification rather than matching. + Var in net matches any term. + Abs in object (and Var if "unif") regarded as wildcard. + If not "unif", Var in object only matches a variable in net.*) +fun matching unif t (net,nets) = + let fun rands _ (Leaf _, nets) = nets + | rands t (Net{comb,alist,...}, nets) = + case t of + (f$t) => foldr (matching unif t) (rands f (comb,[]), nets) + | (Const(c,_)) => look1 (alist, c) nets + | (Free(c,_)) => look1 (alist, c) nets + | (Bound i) => look1 (alist, string_of_bound i) nets + in + case net of + Leaf _ => nets + | Net{var,...} => + case (head_of t) of + Var _ => if unif then net_skip (net,nets) + else var::nets (*only matches Var in net*) + | Abs(_,_,u) => net_skip (net,nets) (*could match anything*) + | _ => rands t (net, var::nets) (*var could match also*) + end; + +val extract_leaves = flat o map (fn Leaf(xs) => xs); + +(*return items whose key could match t*) +fun match_term net t = + extract_leaves (matching false t (net,[])); + +(*return items whose key could unify with t*) +fun unify_term net t = + extract_leaves (matching true t (net,[])); + +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/pattern.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/pattern.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,341 @@ +(* Title: pattern + ID: $Id$ + Author: Tobias Nipkow and Christine Heinzelmann, TU Muenchen + Copyright 1993 TU Muenchen + +Unification of Higher-Order Patterns. + +See also: +Tobias Nipkow. Functional Unification of Higher-Order Patterns. +In Proceedings of the 8th IEEE Symposium Logic in Computer Science, 1993. +*) + +signature PATTERN = +sig + type type_sig + type sg + type env + val eta_contract: term -> term + val match: type_sig -> term * term + -> (indexname*typ)list * (indexname*term)list + val eta_matches: type_sig -> term * term -> bool + val unify: sg * env * (term * term)list -> env + exception Unif + exception MATCH + exception Pattern +end; + +functor PatternFun(structure Sign:SIGN and Envir:ENVIR): PATTERN = +struct + +structure Type = Sign.Type; + +type type_sig = Type.type_sig +type sg = Sign.sg +type env = Envir.env + +exception Unif; +exception Pattern; + +fun occurs(F,t,env) = + let fun occ(Var(G,_)) = (case Envir.lookup(env,G) of + Some(t) => occ t + | None => F=G) + | occ(t1$t2) = occ t1 orelse occ t2 + | occ(Abs(_,_,t)) = occ t + | occ _ = false + in occ t end; + +(* Something's wrong *) +fun ill_formed s = error ("Ill-formed argument in "^s); + + +fun mapbnd f = + let fun mpb d (Bound(i)) = if i < d then Bound(i) else Bound(f(i-d)+d) + | mpb d (Free(c,T)) = Free(c,T) + | mpb d (Const(c,T)) = Const(c,T) + | mpb d (Var(iname,T)) = Var(iname,T) + | mpb d (Abs(s,T,t)) = Abs(s,T,mpb(d+1) t) + | mpb d ((u1 $ u2)) = mpb d (u1)$ mpb d (u2) + in mpb 0 end; + +fun idx [] j = ~10000 + | idx(i::is) j = if i=j then length is else idx is j; + +val nth_type = snd o nth_elem; + +fun at xs i = nth_elem (i,xs); + +fun mkabs (binders,is,t) = + let fun mk(i::is) = let val (x,T) = nth_elem(i,binders) + in Abs(x,T,mk is) end + | mk [] = t + in mk is end; + +val incr = mapbnd (fn i => i+1); + +(* termlist --> intlist *) +fun ints_of [] = [] + | ints_of (Bound i ::bs) = + let val is = ints_of bs + in if i mem is then raise Pattern else i::is end + | ints_of _ = raise Pattern; + + +fun app (s,(i::is)) = app (s$Bound(i),is) + | app (s,[]) = s; + +fun red (Abs(_,_,s)) (i::is) js = red s is (i::js) + | red s is jn = app (mapbnd (at jn) s,is); + +(* split_type ([T1,....,Tn]---> T,n,[]) = ([Tn,...,T1],T) *) +fun split_type (T,0,Ts) = (Ts,T) + | split_type (Type ("fun",[T1,T2]),n,Ts) = split_type (T2,n-1,T1::Ts) + | split_type _ = ill_formed("split_type"); + +fun type_of_G (T,n,is) = + let val (Ts,U) = split_type(T,n,[]) in map(at Ts)is ---> U end; + +fun mkhnf (binders,is,G,js) = mkabs (binders, is, app(G,js)); + +fun mknewhnf(env,binders,is,F as (a,_),T,js) = + let val (env',G) = Envir.genvar a (env,type_of_G(T,length is,js)) + in Envir.update((F,mkhnf(binders,is,G,js)),env') end; + + +fun devar env t = case strip_comb t of + (Var(F,_),ys) => + (case Envir.lookup(env,F) of + Some(t) => devar env (red t (ints_of ys) []) + | None => t) + | _ => t; + + +(* mk_proj_list(is) = [ |is| - k | 1 <= k <= |is| and is[k] >= 0 ] *) +fun mk_proj_list is = + let fun mk(i::is,j) = if i >= 0 then j :: mk(is,j-1) else mk(is,j-1) + | mk([],_) = [] + in mk(is,length is - 1) end; + +fun proj(s,env,binders,is) = + let fun trans d i = if i let val (t',env') = pr(t,env,d+1,((a,T)::binders)) + in (Abs(a,T,t'),env') end + | t => (case strip_comb t of + (c as Const _,ts) => + let val (ts',env') = prs(ts,env,d,binders) + in (list_comb(c,ts'),env') end + | (f as Free _,ts) => + let val (ts',env') = prs(ts,env,d,binders) + in (list_comb(f,ts'),env') end + | (Bound(i),ts) => + let val j = trans d i + in if j < 0 then raise Unif + else let val (ts',env') = prs(ts,env,d,binders) + in (list_comb(Bound j,ts'),env') end + end + | (Var(F as (a,_),Fty),ts) => + let val js = ints_of ts; + val js' = map (trans d) js; + val ks = mk_proj_list js'; + val ls = filter (fn i => i >= 0) js' + val Hty = type_of_G(Fty,length js,ks) + val (env',H) = Envir.genvar a (env,Hty) + val env'' = + Envir.update((F,mkhnf(binders,js,H,ks)),env') + in (app(H,ls),env'') end + | _ => raise Pattern)) + and prs(s::ss,env,d,binders) = + let val (s',env1) = pr(s,env,d,binders) + val (ss',env2) = prs(ss,env1,d,binders) + in (s'::ss',env2) end + | prs([],env,_,_) = ([],env) + in if downto0(is,length binders - 1) then (s,env) + else pr(s,env,0,binders) + end; + + +(* mk_ff_list(is,js) = [ length(is) - k | 1 <= k <= |is| and is[k] = js[k] ] *) +fun mk_ff_list(is,js) = + let fun mk([],[],_) = [] + | mk(i::is,j::js, k) = if i=j then k :: mk(is,js,k-1) + else mk(is,js,k-1) + | mk _ = ill_formed"mk_ff_list" + in mk(is,js,length is-1) end; + +fun flexflex1(env,binders,F,Fty,is,js) = + if is=js then env + else let val ks = mk_ff_list(is,js) + in mknewhnf(env,binders,is,F,Fty,ks) end; + +fun flexflex2(env,binders,F,Fty,is,G,Gty,js) = + let fun ff(F,Fty,is,G as (a,_),Gty,js) = + if js subset is + then let val t= mkabs(binders,is,app(Var(G,Gty),map (idx is) js)) + in Envir.update((F,t),env) end + else let val ks = is inter js + val Hty = type_of_G(Fty,length is,map (idx is) ks) + val (env',H) = Envir.genvar a (env,Hty) + fun lam(is) = mkabs(binders,is,app(H,map (idx is) ks)); + in Envir.update((G,lam js), Envir.update((F,lam is),env')) + end; + in if xless(G,F) then ff(F,Fty,is,G,Gty,js) else ff(G,Gty,js,F,Fty,is) end + +val tsgr = ref(Type.tsig0); + +fun unify_types(T,U, env as Envir.Envir{asol,iTs,maxidx}) = + if T=U then env + else let val iTs' = Type.unify (!tsgr) ((U,T),iTs) + in Envir.Envir{asol=asol,maxidx=maxidx,iTs=iTs'} end + handle Type.TUNIFY => raise Unif; + +fun unif binders (env,(s,t)) = case (devar env s,devar env t) of + (Abs(ns,Ts,ts),Abs(nt,Tt,tt)) => + let val name = if ns = "" then nt else ns + in unif ((name,Ts)::binders) (env,(ts,tt)) end + | (Abs(ns,Ts,ts),t) => unif ((ns,Ts)::binders) (env,(ts,(incr t)$Bound(0))) + | (t,Abs(nt,Tt,tt)) => unif ((nt,Tt)::binders) (env,((incr t)$Bound(0),tt)) + | p => cases(binders,env,p) + +and cases(binders,env,(s,t)) = case (strip_comb s,strip_comb t) of + ((Var(F,Fty),ss),(Var(G,Gty),ts)) => + if F = G then flexflex1(env,binders,F,Fty,ints_of ss,ints_of ts) + else flexflex2(env,binders,F,Fty,ints_of ss,G,Gty,ints_of ts) + | ((Var(F,_),ss),_) => flexrigid(env,binders,F,ints_of ss,t) + | (_,(Var(F,_),ts)) => flexrigid(env,binders,F,ints_of ts,s) + | ((Const c,ss),(Const d,ts)) => rigidrigid(env,binders,c,d,ss,ts) + | ((Free(f),ss),(Free(g),ts)) => rigidrigid(env,binders,f,g,ss,ts) + | ((Bound(i),ss),(Bound(j),ts)) => rigidrigidB (env,binders,i,j,ss,ts) + | ((Abs(_),_),_) => raise Pattern + | (_,(Abs(_),_)) => raise Pattern + | _ => raise Unif + +and rigidrigid (env,binders,(a,Ta),(b,Tb),ss,ts) = + if a<>b then raise Unif + else foldl (unif binders) (unify_types(Ta,Tb,env), ss~~ts) + +and rigidrigidB (env,binders,i,j,ss,ts) = + if i <> j then raise Unif else foldl (unif binders) (env ,ss~~ts) + +and flexrigid (env,binders,F,is,t) = + if occurs(F,t,env) then raise Unif + else let val (u,env') = proj(t,env,binders,is) + in Envir.update((F,mkabs(binders,is,u)),env') end; + +fun unify(sg,env,tus) = (tsgr := #tsig(Sign.rep_sg sg); + foldl (unif []) (env,tus)); + + +(*Perform eta-contractions upon a term*) +fun eta_contract (Abs(a,T,body)) = + (case eta_contract body of + body' as (f $ Bound i) => + if i=0 andalso not (0 mem loose_bnos f) then incr_boundvars ~1 f + else Abs(a,T,body') + | body' => Abs(a,T,body')) + | eta_contract(f$t) = eta_contract f $ eta_contract t + | eta_contract t = t; + + +(* Pattern matching. Raises MATCH if non-pattern *) +exception MATCH; +(* something wron with types, esp in abstractions +fun typ_match args = Type.typ_match (!tsgr) args + handle Type.TYPE_MATCH => raise MATCH; + +fun match_bind(itms,binders,ixn,is,t) = + let val js = loose_bnos t + in if null is + then if null js then (ixn,t)::itms else raise MATCH + else if js subset is + then let val t' = if downto0(is,length binders - 1) then t + else mapbnd (idx is) t + in (ixn, eta_contract(mkabs(binders,is,t'))) :: itms end + else raise MATCH + end; + +fun match_rr (iTs,(a,Ta),(b,Tb)) = + if a<>b then raise MATCH else typ_match (iTs,(Ta,Tb)) + +(* Pre: pat and obj have same type *) +fun mtch(binders,env as (iTs,itms),pat,obj) = case pat of + Var(ixn,_) => (case assoc(itms,ixn) of + None => (iTs,match_bind(itms,binders,ixn,[],obj)) + | Some u => if obj aconv u then env else raise MATCH) + | Abs(ns,Ts,ts) => + (case obj of + Abs(nt,Tt,tt) => mtch((nt,Tt)::binders,env,ts,tt) + | _ => let val Tt = typ_subst_TVars iTs Ts + in mtch((ns,Tt)::binders,env,ts,(incr obj)$Bound(0)) end) + | _ => (case obj of + Abs(nt,Tt,tt) => + mtch((nt,Tt)::binders,env,(incr pat)$Bound(0),tt) + | _ => cases(binders,env,pat,obj)) + +and cases(binders,env as (iTs,itms),pat,obj) = + let fun structural() = case (pat,obj) of + (Const c,Const d) => (match_rr(iTs,c,d),itms) + | (Free f,Free g) => (match_rr(iTs,f,g),itms) + | (Bound i,Bound j) => if i=j then env else raise MATCH + | (f$t,g$u) => mtch(binders,mtch(binders,env,t,u),f,g) + | _ => raise MATCH + in case strip_comb pat of + (Var(ixn,_),bs) => + (let val is = ints_of bs + in case assoc(itms,ixn) of + None => (iTs,match_bind(itms,binders,ixn,is,obj)) + | Some u => if obj aconv (red u is []) then env else raise MATCH + end (* if ints_of fails: *) handle Pattern => structural()) + | _ => structural() + end; + +fun match tsg = (tsgr := tsg; + fn (pat,obj) => + let val pT = fastype_of([],pat) + and oT = fastype_of([],obj) + val iTs = typ_match([],(pT,oT)) + in mtch([], (iTs,[]), pat, eta_contract obj) + handle Pattern => raise MATCH + end) + +(*Predicate: does the pattern match the object?*) +fun matches tsig args = (match tsig args; true) + handle MATCH => false; +*) + +(*First-order matching; term_match tsig (pattern, object) + returns a (tyvar,typ)list and (var,term)list. + The pattern and object may have variables in common. + Instantiation does not affect the object, so matching ?a with ?a+1 works. + A Const does not match a Free of the same name! + Does not notice eta-equality, thus f does not match %(x)f(x) *) +fun match tsig (pat,obj) = + let fun typ_match args = (Type.typ_match tsig args) + handle Type.TYPE_MATCH => raise MATCH; + fun mtch (tyinsts,insts) = fn + (Var(ixn,T), t) => + if null (loose_bnos t) + then case assoc(insts,ixn) of + None => (typ_match (tyinsts, (T,fastype_of([],t))), + (ixn,t)::insts) + | Some u => if t aconv u then (tyinsts,insts) else raise MATCH + else raise MATCH + | (Free (a,T), Free (b,U)) => + if a=b then (typ_match (tyinsts,(T,U)), insts) else raise MATCH + | (Const (a,T), Const (b,U)) => + if a=b then (typ_match (tyinsts,(T,U)), insts) else raise MATCH + | (Bound i, Bound j) => + if i=j then (tyinsts,insts) else raise MATCH + | (Abs(_,T,t), Abs(_,U,u)) => + mtch (typ_match (tyinsts,(T,U)),insts) (t,u) + | (f$t, g$u) => mtch (mtch (tyinsts,insts) (f,g)) (t, u) + | _ => raise MATCH + in mtch([],[]) (pat,obj) end; + +(*Predicate: does the pattern match the object?*) +fun eta_matches tsig (pat,obj) = + (match tsig (eta_contract pat,eta_contract obj); true) + handle MATCH => false; + +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/sequence.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/sequence.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,135 @@ +(* Title: sequence + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1988 University of Cambridge +*) + +(*Unbounded sequences implemented by closures. + + Could use 'a seq = Seq of ('a * (unit -> 'a seq)) option. + Recomputes if sequence is re-inspected; memoing would need polymorphic refs. +*) + + +signature SEQUENCE = + sig + type 'a seq + val append : 'a seq * 'a seq -> 'a seq + val chop: int * 'a seq -> 'a list * 'a seq + val cons: 'a * 'a seq -> 'a seq + val filters: ('a -> bool) -> 'a seq -> 'a seq + val flats: 'a seq seq -> 'a seq + val interleave : 'a seq * 'a seq -> 'a seq + val its_right: ('a * 'b seq -> 'b seq) -> 'a seq * 'b seq -> 'b seq + val list_of_s: 'a seq -> 'a list + val mapp: ('a -> 'b) -> 'a seq -> 'b seq -> 'b seq + val maps: ('a -> 'b) -> 'a seq -> 'b seq + val null: 'a seq + val prints: (int -> 'a -> unit) -> int -> 'a seq -> unit + val pull: 'a seq -> ('a * 'a seq) option + val s_of_list: 'a list -> 'a seq + val seqof: (unit -> ('a * 'a seq) option) -> 'a seq + val single: 'a -> 'a seq + end; + + +functor SequenceFun () : SEQUENCE = +struct + +datatype 'a seq = Seq of unit -> ('a * 'a seq)option; + +(*Return next sequence element as None or Some(x,str) *) +fun pull(Seq f) = f(); + +(*the abstraction for making a sequence*) +val seqof = Seq; + +(*prefix an element to the sequence + use cons(x,s) only if evaluation of s need not be delayed, + otherwise use seqof(fn()=> Some(x,s)) *) +fun cons all = Seq(fn()=>Some all); + +(*the empty sequence*) +val null = Seq(fn()=>None); + +fun single(x) = cons (x, null); + +(*The list of the first n elements, paired with rest of sequence. + If length of list is less than n, then sequence had less than n elements.*) +fun chop (n:int, s: 'a seq): 'a list * 'a seq = + if n<=0 then ([],s) + else case pull(s) of + None => ([],s) + | Some(x,s') => let val (xs,s'') = chop (n-1,s') + in (x::xs, s'') end; + +(*conversion from sequence to list*) +fun list_of_s (s: 'a seq) : 'a list = case pull(s) of + None => [] + | Some(x,s') => x :: list_of_s s'; + + +(*conversion from list to sequence*) +fun s_of_list [] = null + | s_of_list (x::l) = cons (x, s_of_list l); + + +(*functional to print a sequence, up to "count" elements + the function prelem should print the element number and also the element*) +fun prints (prelem: int -> 'a -> unit) count (s: 'a seq) : unit = + let fun pr (k,s) = + if k>count then () + else case pull(s) of + None => () + | Some(x,s') => (prelem k x; prs"\n"; pr (k+1, s')) + in pr(1,s) end; + +(*Map the function f over the sequence, making a new sequence*) +fun maps f xq = seqof (fn()=> case pull(xq) of + None => None + | Some(x,yq) => Some(f x, maps f yq)); + +(*Sequence append: put the elements of xq in front of those of yq*) +fun append (xq,yq) = + let fun copy xq = seqof (fn()=> + case pull xq of + None => pull yq + | Some(x,xq') => Some (x, copy xq')) + in copy xq end; + +(*Interleave elements of xq with those of yq -- fairer than append*) +fun interleave (xq,yq) = seqof (fn()=> + case pull(xq) of + None => pull yq + | Some(x,xq') => Some (x, interleave(yq, xq'))); + +(*map over a sequence xq, append the sequence yq*) +fun mapp f xq yq = + let fun copy s = seqof (fn()=> + case pull(s) of + None => pull(yq) + | Some(x,xq') => Some(f x, copy xq')) + in copy xq end; + +(*Flatten a sequence of sequences to a single sequence. *) +fun flats ss = seqof (fn()=> case pull(ss) of + None => None + | Some(s,ss') => pull(append(s, flats ss'))); + +(*accumulating a function over a sequence; this is lazy*) +fun its_right (f: 'a * 'b seq -> 'b seq) (s: 'a seq, bstr: 'b seq) : 'b seq = + let fun its s = seqof (fn()=> + case pull(s) of + None => pull bstr + | Some(a,s') => pull(f(a, its s'))) + in its s end; + +fun filters pred xq = + let fun copy s = seqof (fn()=> + case pull(s) of + None => None + | Some(x,xq') => if pred x then Some(x, copy xq') + else pull (copy xq') ) + in copy xq end + +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/sign.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/sign.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,332 @@ +(* Title: sign + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + + the abstract types "sg" (signatures) + and "cterm" (certified terms under a signature) +*) + +signature SIGN = +sig + structure Type: TYPE + structure Symtab: SYMTAB + structure Syntax: SYNTAX + sharing Symtab=Type.Symtab + type sg + type cterm + type ctyp + val cfun: (term -> term) -> (cterm -> cterm) + val cterm_of: sg -> term -> cterm + val ctyp_of: sg -> typ -> ctyp + val extend: sg -> string -> + (class * class list) list * class list * + (string list * int) list * + (string list * (sort list * class)) list * + (string list * string)list * Syntax.sext option -> sg + val merge: sg * sg -> sg + val pure: sg + val read_cterm: sg -> string * typ -> cterm + val read_ctyp: sg -> string -> ctyp + val read_insts: sg -> (indexname -> typ option) * (indexname -> sort option) + -> (indexname -> typ option) * (indexname -> sort option) + -> (string*string)list + -> (indexname*ctyp)list * (cterm*cterm)list + val read_typ: sg * (indexname -> sort option) -> string -> typ + val rep_cterm: cterm -> {T: typ, t: term, sign: sg, maxidx: int} + val rep_ctyp: ctyp -> {T: typ, sign: sg} + val rep_sg: sg -> {tsig: Type.type_sig, + const_tab: typ Symtab.table, + syn: Syntax.syntax, + stamps: string ref list} + val string_of_cterm: cterm -> string + val string_of_ctyp: ctyp -> string + val pprint_cterm: cterm -> pprint_args -> unit + val pprint_ctyp: ctyp -> pprint_args -> unit + val string_of_term: sg -> term -> string + val string_of_typ: sg -> typ -> string + val pprint_term: sg -> term -> pprint_args -> unit + val pprint_typ: sg -> typ -> pprint_args -> unit + val term_of: cterm -> term + val typ_of: ctyp -> typ + val pretty_term: sg -> term -> Syntax.Pretty.T +end; + + +functor SignFun (structure Type:TYPE and Syntax:SYNTAX) : SIGN = +struct +structure Type = Type; +structure Symtab = Type.Symtab; +structure Syntax = Syntax; +structure Pretty = Syntax.Pretty + +(*Signatures of theories. *) +datatype sg = + Sg of {tsig: Type.type_sig, (* order-sorted signature of types *) + const_tab: typ Symtab.table, (*types of constants*) + syn: Syntax.syntax, (*Parsing and printing operations*) + stamps: string ref list (*unique theory indentifier*) }; + + +fun rep_sg (Sg args) = args; + +fun string_of_typ(Sg{tsig,syn,...}) = Syntax.string_of_typ syn; + +fun pprint_typ(Sg{syn,...}) = Pretty.pprint o Pretty.quote o (Syntax.pretty_typ syn); + +(*Is constant present in table with more generic type?*) +fun valid_const tsig ctab (a,T) = case Symtab.lookup(ctab, a) of + Some U => Type.typ_instance(tsig,T,U) | _ => false; + + +(*Check a term for errors. Are all constants and types valid in signature? + Does not check that term is well-typed!*) +fun term_errors (sign as Sg{tsig,const_tab,...}) = +let val showtyp = string_of_typ sign; + fun terrs (Const (a,T), errs) = + if valid_const tsig const_tab (a,T) + then Type.type_errors (tsig,showtyp) (T,errs) + else ("Illegal type for constant: " ^ a ^ ": " ^ showtyp T) :: errs + | terrs (Free (_,T), errs) = Type.type_errors (tsig,showtyp) (T,errs) + | terrs (Var ((a,i),T), errs) = + if i>=0 then Type.type_errors (tsig,showtyp) (T,errs) + else ("Negative index for Var: " ^ a) :: errs + | terrs (Bound _, errs) = errs (*loose bvars detected by type_of*) + | terrs (Abs (_,T,t), errs) = + Type.type_errors(tsig,showtyp)(T,terrs(t,errs)) + | terrs (f$t, errs) = terrs(f, terrs (t,errs)) +in terrs end; + + +(** The Extend operation **) + + +(*Reset TVar indices to zero, renaming to preserve distinctness*) +fun zero_tvar_indices tsig T = + let val inxSs = typ_tvars T; + val nms' = variantlist(map (#1 o #1) inxSs,[]); + val tye = map (fn ((v,S),a) => (v, TVar((a,0),S))) (inxSs ~~ nms') + in typ_subst_TVars tye T end + +(*Check that all types mentioned in the list of declarations are valid. + If errors found then raise exception. + Zero type var indices because type inference requires it. +*) +fun read_consts(tsig,syn) = +let val showtyp = Syntax.string_of_typ syn; + fun read [] = [] + | read((cs,s)::pairs) = + let val t = Syntax.read syn Syntax.typeT s handle ERROR => + error("The error above occurred in type " ^ s); + val S = Type.defaultS tsig; + val T = Type.varifyT(Syntax.typ_of_term (K S) t) + val T0 = zero_tvar_indices tsig T; + in (case Type.type_errors (tsig,showtyp) (T0,[]) of + [] => (cs,T0) :: read pairs + | errs => error (cat_lines + (("Error in type of constants " ^ space_implode " " cs) :: errs))) + end +in read end; + + +(*Extend a signature: may add classes, types and constants. + Replaces syntax with "syn". + The "ref" in stamps ensures that no two signatures are identical -- + it is impossible to forge a signature. *) +fun extend (Sg{tsig, const_tab, syn, stamps, ...}) signame + (newclasses, newdefault, otypes, newtypes, const_decs, osext) : sg = +let val tsig' = Type.extend(tsig,newclasses,newdefault,otypes,newtypes); + val S = Type.defaultS tsig'; + val roots = filter (Type.logical_type tsig') + (distinct(flat(map #1 newtypes))) + val xconsts = map #1 newclasses @ flat (map #1 otypes) @ flat (map #1 const_decs); + val syn' = + case osext of + Some sext => Syntax.extend (syn, K S) (roots, xconsts, sext) + | None => if null roots andalso null xconsts then syn + else Syntax.extend (syn, K S) (roots, xconsts, Syntax.empty_sext); + val sconsts = case osext of + Some(sext) => Syntax.constants sext + | None => []; + val const_decs' = read_consts(tsig',syn') (sconsts @ const_decs) +in Sg{tsig= tsig', + const_tab= Symtab.st_of_declist (const_decs', const_tab) + handle Symtab.DUPLICATE(a) => + error("Constant '" ^ a ^ "' declared twice"), + syn=syn', stamps= ref signame :: stamps} +end; + + +(* The empty signature *) +val sg0 = Sg{tsig= Type.tsig0, const_tab= Symtab.null, + syn=Syntax.type_syn, stamps= []}; + +(*The pure signature*) +val pure : sg = extend sg0 "Pure" +([("logic", [])], + ["logic"], + [(["fun"],2), + (["prop"],0), + (Syntax.syntax_types,0)], + [(["fun"], ([["logic"], ["logic"]], "logic")), + (["prop"], ([], "logic"))], + [(["*NORMALIZED*"], "'a::{} => 'a"), + ([Syntax.constrainC], "'a::logic => 'a")], + Some(Syntax.pure_sext) +); + + +(** The Merge operation **) + +(*Update table with (a,x) providing any existing asgt to "a" equals x. *) +fun update_eq ((a,x),tab) = + case Symtab.lookup(tab,a) of + None => Symtab.update((a,x), tab) + | Some y => if x=y then tab + else raise TERM ("Incompatible types for constant: "^a, []); + +(*Combine tables, updating tab2 by tab1 and checking.*) +fun merge_tabs (tab1,tab2) = + Symtab.balance (foldr update_eq (Symtab.alist_of tab1, tab2)); + +(*Combine tables, overwriting tab2 with tab1.*) +fun smash_tabs (tab1,tab2) = + Symtab.balance (foldr Symtab.update (Symtab.alist_of tab1, tab2)); + +(*Combine stamps, checking that theory names are disjoint. *) +fun merge_stamps (stamps1,stamps2) = + let val stamps = stamps1 union stamps2 in + case findrep (map ! stamps) of + a::_ => error ("Attempt to merge different versions of theory: " ^ a) + | [] => stamps + end; + +(*Merge two signatures. Forms unions of tables. Prefers sign1. *) +fun merge (sign1 as Sg{tsig=tsig1,const_tab=ctab1,stamps=stamps1,syn=syn1}, + sign2 as Sg{tsig=tsig2,const_tab=ctab2,stamps=stamps2,syn=syn2}) = + if stamps2 subset stamps1 then sign1 + else if stamps1 subset stamps2 then sign2 + else (*neither is union already; must form union*) + Sg{tsig= Type.merge(tsig1,tsig2), + const_tab= merge_tabs (ctab1, ctab2), + stamps= merge_stamps (stamps1,stamps2), + syn = Syntax.merge(syn1,syn2)}; + + +(**** CERTIFIED TYPES ****) + + +(*Certified typs under a signature*) +datatype ctyp = Ctyp of {sign: sg, T: typ}; + +fun rep_ctyp(Ctyp ctyp) = ctyp; + +fun typ_of (Ctyp{sign,T}) = T; + +fun ctyp_of (sign as Sg{tsig,...}) T = + case Type.type_errors (tsig,string_of_typ sign) (T,[]) of + [] => Ctyp{sign= sign,T= T} + | errs => error (cat_lines ("Error in type:" :: errs)); + +(*The only use is a horrible hack in the simplifier!*) +fun read_typ(Sg{tsig,syn,...}, defS) s = + let val term = Syntax.read syn Syntax.typeT s; + val S0 = Type.defaultS tsig; + fun defS0 s = case defS s of Some S => S | None => S0; + in Syntax.typ_of_term defS0 term end; + +fun read_ctyp sign = ctyp_of sign o read_typ(sign, K None); + +fun string_of_ctyp (Ctyp{sign,T}) = string_of_typ sign T; + +fun pprint_ctyp (Ctyp{sign,T}) = pprint_typ sign T; + + +(**** CERTIFIED TERMS ****) + +(*Certified terms under a signature, with checked typ and maxidx of Vars*) +datatype cterm = Cterm of {sign: sg, t: term, T: typ, maxidx: int}; + +fun rep_cterm (Cterm args) = args; + +(*Return the underlying term*) +fun term_of (Cterm{sign,t,T,maxidx}) = t; + +(** pretty printing of terms **) + +fun pretty_term (Sg{tsig,syn,...}) = Syntax.pretty_term syn; + +fun string_of_term sign t = Pretty.string_of (pretty_term sign t); + +fun pprint_term sign = Pretty.pprint o Pretty.quote o (pretty_term sign); + +fun string_of_cterm (Cterm{sign,t,...}) = string_of_term sign t; + +fun pprint_cterm (Cterm{sign,t,...}) = pprint_term sign t; + +(*Create a cterm by checking a "raw" term with respect to a signature*) +fun cterm_of sign t = + case term_errors sign (t,[]) of + [] => Cterm{sign=sign, t=t, T= type_of t, maxidx= maxidx_of_term t} + | errs => raise TERM(cat_lines("Term not in signature"::errs), [t]); + +fun cfun f = fn Cterm{sign,t,...} => cterm_of sign (f t); + +(*Lexing, parsing, polymorphic typechecking of a term.*) +fun read_def_cterm (sign as Sg{tsig, const_tab, syn,...}, types, sorts) + (a,T) = + let val showtyp = string_of_typ sign + and showterm = string_of_term sign + fun termerr [] = "" + | termerr [t] = "\nInvolving this term:\n" ^ showterm t ^ "\n" + | termerr ts = "\nInvolving these terms:\n" ^ + cat_lines (map showterm ts) + val t = Syntax.read syn T a; + val (t',tye) = Type.infer_types (tsig, const_tab, types, + sorts, showtyp, T, t) + handle TYPE (msg, Ts, ts) => + error ("Type checking error: " ^ msg ^ "\n" ^ + cat_lines (map showtyp Ts) ^ termerr ts) + in (cterm_of sign t', tye) + end + handle TERM (msg, _) => error ("Error: " ^ msg); + + +fun read_cterm sign = #1 o (read_def_cterm (sign, K None, K None)); + +(** reading of instantiations **) + +fun indexname cs = case Syntax.scan_varname cs of (v,[]) => v + | _ => error("Lexical error in variable name: " ^ implode cs); + +fun absent ixn = + error("No such variable in term: " ^ Syntax.string_of_vname ixn); + +fun inst_failure ixn = + error("Instantiation of " ^ Syntax.string_of_vname ixn ^ " fails"); + +fun read_insts (sign as Sg{tsig,...}) (rtypes,rsorts) (types,sorts) insts = +let fun split([],tvs,vs) = (tvs,vs) + | split((sv,st)::l,tvs,vs) = (case explode sv of + "'"::cs => split(l,(indexname cs,st)::tvs,vs) + | cs => split(l,tvs,(indexname cs,st)::vs)); + val (tvs,vs) = split(insts,[],[]); + fun readT((a,i),st) = + let val ixn = ("'" ^ a,i); + val S = case rsorts ixn of Some S => S | None => absent ixn; + val T = read_typ (sign,sorts) st; + in if Type.typ_instance(tsig,T,TVar(ixn,S)) then (ixn,T) + else inst_failure ixn + end + val tye = map readT tvs; + fun add_cterm ((cts,tye), (ixn,st)) = + let val T = case rtypes ixn of + Some T => typ_subst_TVars tye T + | None => absent ixn; + val (ct,tye2) = read_def_cterm (sign,types,sorts) (st,T); + val cv = cterm_of sign (Var(ixn,typ_subst_TVars tye2 T)) + in ((cv,ct)::cts,tye2 @ tye) end + val (cterms,tye') = foldl add_cterm (([],tye), vs); +in (map (fn (ixn,T) => (ixn,ctyp_of sign T)) tye', cterms) end; + +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/symtab.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/symtab.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,114 @@ +(* Title: symtab + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1989 University of Cambridge +*) + +(*Unbalanced binary trees indexed by strings + No way to delete an entry + could generalize alist_of to a traversal functional +*) + +signature SYMTAB = +sig + type 'a table + val alist_of : 'a table -> (string*'a) list + val balance : 'a table -> 'a table + val lookup : 'a table * string -> 'a option + val null : 'a table + val is_null : 'a table -> bool + val st_of_alist : (string*'a)list * 'a table -> 'a table + val st_of_declist : (string list * 'a)list * 'a table -> 'a table + val update : (string*'a) * 'a table -> 'a table + val update_new : (string*'a) * 'a table -> 'a table + exception DUPLICATE of string +end; + + +functor SymtabFun () : SYMTAB = +struct + +(*symbol table errors, such as from update_new*) +exception DUPLICATE of string; + +datatype 'a table = Tip | Branch of (string * 'a * 'a table * 'a table); + + +val null = Tip; + +fun is_null Tip = true + | is_null _ = false; + + +fun lookup (symtab: 'a table, key: string) : 'a option = + let fun look Tip = None + | look (Branch (key',entry,left,right)) = + if key < key' then look left + else if key' < key then look right + else Some entry + in look symtab end; + +(*update, allows overwriting of an entry*) +fun update ((key: string, entry: 'a), symtab : 'a table) + : 'a table = + let fun upd Tip = Branch (key,entry,Tip,Tip) + | upd (Branch(key',entry',left,right)) = + if key < key' then Branch (key',entry', upd left, right) + else if key' < key then Branch (key',entry',left, upd right) + else Branch (key,entry,left,right) + in upd symtab end; + +(*Like update but fails if key is already defined in table. + Allows st_of_alist, etc. to detect multiple definitions*) +fun update_new ((key: string, entry: 'a), symtab : 'a table) + : 'a table = + let fun upd Tip = Branch (key,entry,Tip,Tip) + | upd (Branch(key',entry',left,right)) = + if key < key' then Branch (key',entry', upd left, right) + else if key' < key then Branch (key',entry',left, upd right) + else raise DUPLICATE(key) + in upd symtab end; + +(*conversion of symbol table to sorted association list*) +fun alist_of (symtab : 'a table) : (string * 'a) list = + let fun ali (symtab,cont) = case symtab of + Tip => cont + | Branch (key,entry,left,right) => + ali(left, (key,entry) :: ali(right,cont)) + in ali (symtab,[]) end; + + +(*Make a balanced tree of the first n members of the sorted alist (sal). + Utility for the function balance.*) +fun bal_of (sal, 0) = Tip + | bal_of (sal, n) = + let val mid = n div 2 + in case drop (mid,sal) of + [] => bal_of (sal, mid) (*should not occur*) + | ((key,entry):: pairs) => + Branch(key,entry, bal_of(sal,mid), bal_of(pairs, n-mid-1)) + end; + + +fun balance symtab = + let val sal = alist_of symtab + in bal_of (sal, length sal) end; + + +(*Addition of association list to a symbol table*) +fun st_of_alist (al, symtab) = + foldr update_new (al, symtab); + +(*A "declaration" associates the same entry with a list of keys; + does not allow overwriting of an entry*) +fun decl_update_new ((keys : string list, entry: 'a), symtab) + : 'a table = + let fun decl (key,symtab) = update_new((key,entry), symtab) + in foldr decl (keys, symtab) end; + +(*Addition of a list of declarations to a symbol table*) +fun st_of_declist (dl, symtab) = + balance (foldr decl_update_new (dl, symtab)) + +end; + diff -r 000000000000 -r a5a9c433f639 src/Pure/tactic.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/tactic.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,421 @@ +(* Title: tactic + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Tactics +*) + +signature TACTIC = +sig + structure Tactical: TACTICAL and Net: NET + local open Tactical Tactical.Thm Net + in + val ares_tac: thm list -> int -> tactic + val asm_rewrite_goal_tac: + (meta_simpset -> tactic) -> meta_simpset -> int -> tactic + val assume_tac: int -> tactic + val atac: int ->tactic + val bimatch_from_nets_tac: (int*(bool*thm)) net * (int*(bool*thm)) net -> int -> tactic + val bimatch_tac: (bool*thm)list -> int -> tactic + val biresolve_from_nets_tac: (int*(bool*thm)) net * (int*(bool*thm)) net -> int -> tactic + val biresolve_tac: (bool*thm)list -> int -> tactic + val build_net: thm list -> (int*thm) net + val build_netpair: (bool*thm)list -> (int*(bool*thm)) net * (int*(bool*thm)) net + val compose_inst_tac: (string*string)list -> (bool*thm*int) -> int -> tactic + val compose_tac: (bool * thm * int) -> int -> tactic + val cut_facts_tac: thm list -> int -> tactic + val dmatch_tac: thm list -> int -> tactic + val dresolve_tac: thm list -> int -> tactic + val dres_inst_tac: (string*string)list -> thm -> int -> tactic + val dtac: thm -> int ->tactic + val etac: thm -> int ->tactic + val eq_assume_tac: int -> tactic + val ematch_tac: thm list -> int -> tactic + val eresolve_tac: thm list -> int -> tactic + val eres_inst_tac: (string*string)list -> thm -> int -> tactic + val filter_thms: (term*term->bool) -> int*term*thm list -> thm list + val filt_resolve_tac: thm list -> int -> int -> tactic + val flexflex_tac: tactic + val fold_goals_tac: thm list -> tactic + val fold_tac: thm list -> tactic + val forward_tac: thm list -> int -> tactic + val forw_inst_tac: (string*string)list -> thm -> int -> tactic + val is_fact: thm -> bool + val lessb: (bool * thm) * (bool * thm) -> bool + val lift_inst_rule: thm * int * (string*string)list * thm -> thm + val make_elim: thm -> thm + val match_from_net_tac: (int*thm) net -> int -> tactic + val match_tac: thm list -> int -> tactic + val metacut_tac: thm -> int -> tactic + val net_bimatch_tac: (bool*thm) list -> int -> tactic + val net_biresolve_tac: (bool*thm) list -> int -> tactic + val net_match_tac: thm list -> int -> tactic + val net_resolve_tac: thm list -> int -> tactic + val PRIMITIVE: (thm -> thm) -> tactic + val PRIMSEQ: (thm -> thm Sequence.seq) -> tactic + val prune_params_tac: tactic + val rename_tac: string -> int -> tactic + val rename_last_tac: string -> string list -> int -> tactic + val resolve_from_net_tac: (int*thm) net -> int -> tactic + val resolve_tac: thm list -> int -> tactic + val res_inst_tac: (string*string)list -> thm -> int -> tactic + val rewrite_goals_tac: thm list -> tactic + val rewrite_tac: thm list -> tactic + val rewtac: thm -> tactic + val rtac: thm -> int -> tactic + val rule_by_tactic: tactic -> thm -> thm + val subgoals_of_brl: bool * thm -> int + val subgoal_tac: string -> int -> tactic + val trace_goalno_tac: (int -> tactic) -> int -> tactic + end +end; + + +functor TacticFun (structure Logic: LOGIC and Drule: DRULE and + Tactical: TACTICAL and Net: NET + sharing Drule.Thm = Tactical.Thm) : TACTIC = +struct +structure Tactical = Tactical; +structure Thm = Tactical.Thm; +structure Net = Net; +structure Sequence = Thm.Sequence; +structure Sign = Thm.Sign; +local open Tactical Tactical.Thm Drule +in + +(*Discover what goal is chosen: SOMEGOAL(trace_goalno_tac tac) *) +fun trace_goalno_tac tf i = Tactic (fn state => + case Sequence.pull(tapply(tf i, state)) of + None => Sequence.null + | seqcell => (prs("Subgoal " ^ string_of_int i ^ " selected\n"); + Sequence.seqof(fn()=> seqcell))); + +fun string_of (a,0) = a + | string_of (a,i) = a ^ "_" ^ string_of_int i; + +(*convert all Vars in a theorem to Frees -- export??*) +fun freeze th = + let val fth = freezeT th + val {prop,sign,...} = rep_thm fth + fun mk_inst (Var(v,T)) = + (Sign.cterm_of sign (Var(v,T)), + Sign.cterm_of sign (Free(string_of v, T))) + val insts = map mk_inst (term_vars prop) + in instantiate ([],insts) fth end; + +(*Makes a rule by applying a tactic to an existing rule*) +fun rule_by_tactic (Tactic tf) rl = + case Sequence.pull(tf (freeze (standard rl))) of + None => raise THM("rule_by_tactic", 0, [rl]) + | Some(rl',_) => standard rl'; + +(*** Basic tactics ***) + +(*Makes a tactic whose effect on a state is given by thmfun: thm->thm seq.*) +fun PRIMSEQ thmfun = Tactic (fn state => thmfun state + handle THM _ => Sequence.null); + +(*Makes a tactic whose effect on a state is given by thmfun: thm->thm.*) +fun PRIMITIVE thmfun = PRIMSEQ (Sequence.single o thmfun); + +(*** The following fail if the goal number is out of range: + thus (REPEAT (resolve_tac rules i)) stops once subgoal i disappears. *) + +(*Solve subgoal i by assumption*) +fun assume_tac i = PRIMSEQ (assumption i); + +(*Solve subgoal i by assumption, using no unification*) +fun eq_assume_tac i = PRIMITIVE (eq_assumption i); + +(** Resolution/matching tactics **) + +(*The composition rule/state: no lifting or var renaming. + The arg = (bires_flg, orule, m) ; see bicompose for explanation.*) +fun compose_tac arg i = PRIMSEQ (bicompose false arg i); + +(*Converts a "destruct" rule like P&Q==>P to an "elimination" rule + like [| P&Q; P==>R |] ==> R *) +fun make_elim rl = zero_var_indexes (rl RS revcut_rl); + +(*Attack subgoal i by resolution, using flags to indicate elimination rules*) +fun biresolve_tac brules i = PRIMSEQ (biresolution false brules i); + +(*Resolution: the simple case, works for introduction rules*) +fun resolve_tac rules = biresolve_tac (map (pair false) rules); + +(*Resolution with elimination rules only*) +fun eresolve_tac rules = biresolve_tac (map (pair true) rules); + +(*Forward reasoning using destruction rules.*) +fun forward_tac rls = resolve_tac (map make_elim rls) THEN' assume_tac; + +(*Like forward_tac, but deletes the assumption after use.*) +fun dresolve_tac rls = eresolve_tac (map make_elim rls); + +(*Shorthand versions: for resolution with a single theorem*) +fun rtac rl = resolve_tac [rl]; +fun etac rl = eresolve_tac [rl]; +fun dtac rl = dresolve_tac [rl]; +val atac = assume_tac; + +(*Use an assumption or some rules ... A popular combination!*) +fun ares_tac rules = assume_tac ORELSE' resolve_tac rules; + +(*Matching tactics -- as above, but forbid updating of state*) +fun bimatch_tac brules i = PRIMSEQ (biresolution true brules i); +fun match_tac rules = bimatch_tac (map (pair false) rules); +fun ematch_tac rules = bimatch_tac (map (pair true) rules); +fun dmatch_tac rls = ematch_tac (map make_elim rls); + +(*Smash all flex-flex disagreement pairs in the proof state.*) +val flexflex_tac = PRIMSEQ flexflex_rule; + +(*Lift and instantiate a rule wrt the given state and subgoal number *) +fun lift_inst_rule (state, i, sinsts, rule) = +let val {maxidx,sign,...} = rep_thm state + val (_, _, Bi, _) = dest_state(state,i) + val params = Logic.strip_params Bi (*params of subgoal i*) + val params = rev(rename_wrt_term Bi params) (*as they are printed*) + val paramTs = map #2 params + and inc = maxidx+1 + fun liftvar (Var ((a,j), T)) = Var((a, j+inc), paramTs---> incr_tvar inc T) + | liftvar t = raise TERM("Variable expected", [t]); + fun liftterm t = list_abs_free (params, + Logic.incr_indexes(paramTs,inc) t) + (*Lifts instantiation pair over params*) + fun liftpair (cv,ct) = (Sign.cfun liftvar cv, Sign.cfun liftterm ct) + fun lifttvar((a,i),ctyp) = + let val {T,sign} = Sign.rep_ctyp ctyp + in ((a,i+inc), Sign.ctyp_of sign (incr_tvar inc T)) end + val rts = types_sorts rule and (types,sorts) = types_sorts state + fun types'(a,~1) = (case assoc(params,a) of None => types(a,~1) | sm => sm) + | types'(ixn) = types ixn; + val (Tinsts,insts) = Sign.read_insts sign rts (types',sorts) sinsts +in instantiate (map lifttvar Tinsts, map liftpair insts) + (lift_rule (state,i) rule) +end; + + +(*** Resolve after lifting and instantation; may refer to parameters of the + subgoal. Fails if "i" is out of range. ***) + +(*compose version: arguments are as for bicompose.*) +fun compose_inst_tac sinsts (bires_flg, rule, nsubgoal) i = + STATE ( fn state => + compose_tac (bires_flg, lift_inst_rule (state, i, sinsts, rule), + nsubgoal) i + handle TERM (msg,_) => (writeln msg; no_tac) + | THM _ => no_tac ); + +(*Resolve version*) +fun res_inst_tac sinsts rule i = + compose_inst_tac sinsts (false, rule, nprems_of rule) i; + +(*eresolve (elimination) version*) +fun eres_inst_tac sinsts rule i = + compose_inst_tac sinsts (true, rule, nprems_of rule) i; + +(*For forw_inst_tac and dres_inst_tac: preserve Var indexes of rl. + Fails if rl's major premise contains !! or ==> ; it should not anyway!*) +fun make_elim_preserve rl = + let val revcut_rl' = lift_rule (rl,1) revcut_rl + val arg = (false, rl, nprems_of rl) + val [th] = Sequence.list_of_s (bicompose false arg 1 revcut_rl') + in th end + handle Bind => raise THM("make_elim_preserve", 1, [rl]); + +(*forward version*) +fun forw_inst_tac sinsts rule = + res_inst_tac sinsts (make_elim_preserve rule) THEN' assume_tac; + +(*dresolve version*) +fun dres_inst_tac sinsts rule = eres_inst_tac sinsts (make_elim_preserve rule); + +(*** Applications of cut_rl -- forward reasoning ***) + +(*Used by metacut_tac*) +fun bires_cut_tac arg i = + resolve_tac [cut_rl] i THEN biresolve_tac arg (i+1) ; + +(*The conclusion of the rule gets assumed in subgoal i, + while subgoal i+1,... are the premises of the rule.*) +fun metacut_tac rule = bires_cut_tac [(false,rule)]; + +(*Recognizes theorems that are not rules, but simple propositions*) +fun is_fact rl = + case prems_of rl of + [] => true | _::_ => false; + +(*"Cut" all facts from theorem list into the goal as assumptions. *) +fun cut_facts_tac ths i = + EVERY (map (fn th => metacut_tac th i) (filter is_fact ths)); + +(*Introduce the given proposition as a lemma and subgoal*) +fun subgoal_tac sprop = res_inst_tac [("psi", sprop)] cut_rl; + + +(**** Indexing and filtering of theorems ****) + +(*Returns the list of potentially resolvable theorems for the goal "prem", + using the predicate could(subgoal,concl). + Resulting list is no longer than "limit"*) +fun filter_thms could (limit, prem, ths) = + let val pb = Logic.strip_assums_concl prem; (*delete assumptions*) + fun filtr (limit, []) = [] + | filtr (limit, th::ths) = + if limit=0 then [] + else if could(pb, concl_of th) then th :: filtr(limit-1, ths) + else filtr(limit,ths) + in filtr(limit,ths) end; + + +(*** biresolution and resolution using nets ***) + +(** To preserve the order of the rules, tag them with increasing integers **) + +(*insert tags*) +fun taglist k [] = [] + | taglist k (x::xs) = (k,x) :: taglist (k+1) xs; + +(*remove tags and suppress duplicates -- list is assumed sorted!*) +fun untaglist [] = [] + | untaglist [(k:int,x)] = [x] + | untaglist ((k,x) :: (rest as (k',x')::_)) = + if k=k' then untaglist rest + else x :: untaglist rest; + +(*return list elements in original order*) +val orderlist = untaglist o sort (fn(x,y)=> #1 x < #1 y); + +(*insert one tagged brl into the pair of nets*) +fun insert_kbrl (kbrl as (k,(eres,th)), (inet,enet)) = + if eres then + case prems_of th of + prem::_ => (inet, Net.insert_term ((prem,kbrl), enet, K false)) + | [] => error"insert_kbrl: elimination rule with no premises" + else (Net.insert_term ((concl_of th, kbrl), inet, K false), enet); + +(*build a pair of nets for biresolution*) +fun build_netpair brls = + foldr insert_kbrl (taglist 1 brls, (Net.empty,Net.empty)); + +(*biresolution using a pair of nets rather than rules*) +fun biresolution_from_nets_tac match (inet,enet) = + SUBGOAL + (fn (prem,i) => + let val hyps = Logic.strip_assums_hyp prem + and concl = Logic.strip_assums_concl prem + val kbrls = Net.unify_term inet concl @ + flat (map (Net.unify_term enet) hyps) + in PRIMSEQ (biresolution match (orderlist kbrls) i) end); + +(*versions taking pre-built nets*) +val biresolve_from_nets_tac = biresolution_from_nets_tac false; +val bimatch_from_nets_tac = biresolution_from_nets_tac true; + +(*fast versions using nets internally*) +val net_biresolve_tac = biresolve_from_nets_tac o build_netpair; +val net_bimatch_tac = bimatch_from_nets_tac o build_netpair; + +(*** Simpler version for resolve_tac -- only one net, and no hyps ***) + +(*insert one tagged rl into the net*) +fun insert_krl (krl as (k,th), net) = + Net.insert_term ((concl_of th, krl), net, K false); + +(*build a net of rules for resolution*) +fun build_net rls = + foldr insert_krl (taglist 1 rls, Net.empty); + +(*resolution using a net rather than rules; pred supports filt_resolve_tac*) +fun filt_resolution_from_net_tac match pred net = + SUBGOAL + (fn (prem,i) => + let val krls = Net.unify_term net (Logic.strip_assums_concl prem) + in + if pred krls + then PRIMSEQ + (biresolution match (map (pair false) (orderlist krls)) i) + else no_tac + end); + +(*Resolve the subgoal using the rules (making a net) unless too flexible, + which means more than maxr rules are unifiable. *) +fun filt_resolve_tac rules maxr = + let fun pred krls = length krls <= maxr + in filt_resolution_from_net_tac false pred (build_net rules) end; + +(*versions taking pre-built nets*) +val resolve_from_net_tac = filt_resolution_from_net_tac false (K true); +val match_from_net_tac = filt_resolution_from_net_tac true (K true); + +(*fast versions using nets internally*) +val net_resolve_tac = resolve_from_net_tac o build_net; +val net_match_tac = match_from_net_tac o build_net; + + +(*** For Natural Deduction using (bires_flg, rule) pairs ***) + +(*The number of new subgoals produced by the brule*) +fun subgoals_of_brl (true,rule) = length (prems_of rule) - 1 + | subgoals_of_brl (false,rule) = length (prems_of rule); + +(*Less-than test: for sorting to minimize number of new subgoals*) +fun lessb (brl1,brl2) = subgoals_of_brl brl1 < subgoals_of_brl brl2; + + +(*** Meta-Rewriting Tactics ***) + +fun result1 tacf mss thm = + case Sequence.pull(tapply(tacf mss,thm)) of + None => None + | Some(thm,_) => Some(thm); + +(*Rewrite subgoal i only *) +fun asm_rewrite_goal_tac prover_tac mss i = + PRIMITIVE(rewrite_goal_rule (result1 prover_tac) mss i); + +(*Rewrite or fold throughout proof state. *) +fun rewrite_tac thms = PRIMITIVE(rewrite_rule thms); +fun fold_tac rths = rewrite_tac (map symmetric rths); + +(*Rewrite subgoals only, not main goal. *) +fun rewrite_goals_tac thms = PRIMITIVE (rewrite_goals_rule thms); +fun fold_goals_tac rths = rewrite_goals_tac (map symmetric rths); + +fun rewtac rth = rewrite_goals_tac [rth]; + + +(** Renaming of parameters in a subgoal + Names may contain letters, digits or primes and must be + separated by blanks **) + +(*Calling this will generate the warning "Same as previous level" since + it affects nothing but the names of bound variables!*) +fun rename_tac str i = + let val cs = explode str + in + if !Logic.auto_rename + then (writeln"Note: setting Logic.auto_rename := false"; + Logic.auto_rename := false) + else (); + case #2 (take_prefix (is_letdig orf is_blank) cs) of + [] => PRIMITIVE (rename_params_rule (scanwords is_letdig cs, i)) + | c::_ => error ("Illegal character: " ^ c) + end; + +(*Rename recent parameters using names generated from (a) and the suffixes, + provided the string (a), which represents a term, is an identifier. *) +fun rename_last_tac a sufs i = + let val names = map (curry op^ a) sufs + in if Syntax.is_identifier a + then PRIMITIVE (rename_params_rule (names,i)) + else all_tac + end; + +(*Prunes all redundant parameters from the proof state by rewriting*) +val prune_params_tac = rewrite_tac [triv_forall_equality]; + +end; +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/tctical.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/tctical.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,542 @@ +(* Title: tctical + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Tacticals +*) + +infix 1 THEN THEN' THEN_BEST_FIRST; +infix 0 ORELSE APPEND INTLEAVE ORELSE' APPEND' INTLEAVE'; + + +signature TACTICAL = + sig + structure Thm : THM + local open Thm in + datatype tactic = Tactic of thm -> thm Sequence.seq + val all_tac: tactic + val ALLGOALS: (int -> tactic) -> tactic + val APPEND: tactic * tactic -> tactic + val APPEND': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic + val BEST_FIRST: (thm -> bool) * (thm -> int) -> tactic -> tactic + val BREADTH_FIRST: (thm -> bool) -> tactic -> tactic + val CHANGED: tactic -> tactic + val COND: (thm -> bool) -> tactic -> tactic -> tactic + val DEPTH_FIRST: (thm -> bool) -> tactic -> tactic + val DEPTH_SOLVE: tactic -> tactic + val DEPTH_SOLVE_1: tactic -> tactic + val DETERM: tactic -> tactic + val EVERY: tactic list -> tactic + val EVERY': ('a -> tactic) list -> 'a -> tactic + val EVERY1: (int -> tactic) list -> tactic + val FILTER: (thm -> bool) -> tactic -> tactic + val FIRST: tactic list -> tactic + val FIRST': ('a -> tactic) list -> 'a -> tactic + val FIRST1: (int -> tactic) list -> tactic + val FIRSTGOAL: (int -> tactic) -> tactic + val goals_limit: int ref + val has_fewer_prems: int -> thm -> bool + val IF_UNSOLVED: tactic -> tactic + val INTLEAVE: tactic * tactic -> tactic + val INTLEAVE': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic + val METAHYPS: (thm list -> tactic) -> int -> tactic + val no_tac: tactic + val ORELSE: tactic * tactic -> tactic + val ORELSE': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic + val pause_tac: tactic + val print_tac: tactic + val REPEAT1: tactic -> tactic + val REPEAT: tactic -> tactic + val REPEAT_DETERM: tactic -> tactic + val REPEAT_FIRST: (int -> tactic) -> tactic + val REPEAT_SOME: (int -> tactic) -> tactic + val SELECT_GOAL: tactic -> int -> tactic + val SOMEGOAL: (int -> tactic) -> tactic + val STATE: (thm -> tactic) -> tactic + val strip_context: term -> (string * typ) list * term list * term + val SUBGOAL: ((term*int) -> tactic) -> int -> tactic + val tapply: tactic * thm -> thm Sequence.seq + val THEN: tactic * tactic -> tactic + val THEN': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic + val THEN_BEST_FIRST: tactic * ((thm->bool) * (thm->int) * tactic) -> tactic + val traced_tac: (thm -> (thm * thm Sequence.seq) option) -> tactic + val tracify: bool ref -> tactic -> thm -> thm Sequence.seq + val trace_BEST_FIRST: bool ref + val trace_DEPTH_FIRST: bool ref + val trace_REPEAT: bool ref + val TRY: tactic -> tactic + val TRYALL: (int -> tactic) -> tactic + end + end; + + +functor TacticalFun (structure Logic: LOGIC and Drule: DRULE) : TACTICAL = +struct +structure Thm = Drule.Thm; +structure Sequence = Thm.Sequence; +structure Sign = Thm.Sign; +local open Drule Thm +in + +(**** Tactics ****) + +(*A tactic maps a proof tree to a sequence of proof trees: + if length of sequence = 0 then the tactic does not apply; + if length > 1 then backtracking on the alternatives can occur.*) + +datatype tactic = Tactic of thm -> thm Sequence.seq; + +fun tapply(Tactic tf, state) = tf (state); + +(*Makes a tactic from one that uses the components of the state.*) +fun STATE tacfun = Tactic (fn state => tapply(tacfun state, state)); + + +(*** LCF-style tacticals ***) + +(*the tactical THEN performs one tactic followed by another*) +fun (Tactic tf1) THEN (Tactic tf2) = + Tactic (fn state => Sequence.flats (Sequence.maps tf2 (tf1 state))); + + +(*The tactical ORELSE uses the first tactic that returns a nonempty sequence. + Like in LCF, ORELSE commits to either tac1 or tac2 immediately. + Does not backtrack to tac2 if tac1 was initially chosen. *) +fun (Tactic tf1) ORELSE (Tactic tf2) = + Tactic (fn state => + case Sequence.pull(tf1 state) of + None => tf2 state + | sequencecell => Sequence.seqof(fn()=> sequencecell)); + + +(*The tactical APPEND combines the results of two tactics. + Like ORELSE, but allows backtracking on both tac1 and tac2. + The tactic tac2 is not applied until needed.*) +fun (Tactic tf1) APPEND (Tactic tf2) = + Tactic (fn state => Sequence.append(tf1 state, + Sequence.seqof(fn()=> Sequence.pull (tf2 state)))); + +(*Like APPEND, but interleaves results of tac1 and tac2.*) +fun (Tactic tf1) INTLEAVE (Tactic tf2) = + Tactic (fn state => Sequence.interleave(tf1 state, + Sequence.seqof(fn()=> Sequence.pull (tf2 state)))); + +(*Versions for combining tactic-valued functions, as in + SOMEGOAL (resolve_tac rls THEN' assume_tac) *) +fun tac1 THEN' tac2 = fn x => tac1 x THEN tac2 x; +fun tac1 ORELSE' tac2 = fn x => tac1 x ORELSE tac2 x; +fun tac1 APPEND' tac2 = fn x => tac1 x APPEND tac2 x; +fun tac1 INTLEAVE' tac2 = fn x => tac1 x INTLEAVE tac2 x; + +(*passes all proofs through unchanged; identity of THEN*) +val all_tac = Tactic (fn state => Sequence.single state); + +(*passes no proofs through; identity of ORELSE and APPEND*) +val no_tac = Tactic (fn state => Sequence.null); + + +(*Make a tactic deterministic by chopping the tail of the proof sequence*) +fun DETERM (Tactic tf) = Tactic (fn state => + case Sequence.pull (tf state) of + None => Sequence.null + | Some(x,_) => Sequence.cons(x, Sequence.null)); + + +(*Conditional tactical: testfun controls which tactic to use next. + Beware: due to eager evaluation, both thentac and elsetac are evaluated.*) +fun COND testfun (Tactic thenf) (Tactic elsef) = Tactic (fn prf => + if testfun prf then thenf prf else elsef prf); + +(*Do the tactic or else do nothing*) +fun TRY tac = tac ORELSE all_tac; + + +(*** List-oriented tactics ***) + +(* EVERY [tac1,...,tacn] equals tac1 THEN ... THEN tacn *) +fun EVERY tacs = foldr (op THEN) (tacs, all_tac); + +(* EVERY' [tf1,...,tfn] i equals tf1 i THEN ... THEN tfn i *) +fun EVERY' tfs = foldr (op THEN') (tfs, K all_tac); + +(*Apply every tactic to 1*) +fun EVERY1 tfs = EVERY' tfs 1; + +(* FIRST [tac1,...,tacn] equals tac1 ORELSE ... ORELSE tacn *) +fun FIRST tacs = foldr (op ORELSE) (tacs, no_tac); + +(* FIRST' [tf1,...,tfn] i equals tf1 i ORELSE ... ORELSE tfn i *) +fun FIRST' tfs = foldr (op ORELSE') (tfs, K no_tac); + +(*Apply first tactic to 1*) +fun FIRST1 tfs = FIRST' tfs 1; + + +(*** Tracing tactics ***) + +(*Max number of goals to print -- set by user*) +val goals_limit = ref 10; + +(*Print the current proof state and pass it on.*) +val print_tac = Tactic (fn state => + (print_goals (!goals_limit) state; Sequence.single state)); + +(*Pause until a line is typed -- if non-empty then fail. *) +val pause_tac = Tactic (fn state => + (prs"** Press RETURN to continue: "; + if input(std_in,1) = "\n" then Sequence.single state + else (prs"Goodbye\n"; Sequence.null))); + +exception TRACE_EXIT of thm +and TRACE_QUIT; + +(*Handle all tracing commands for current state and tactic *) +fun exec_trace_command flag (tf, state) = + case input_line(std_in) of + "\n" => tf state + | "f\n" => Sequence.null + | "o\n" => (flag:=false; tf state) + | "x\n" => (prs"Exiting now\n"; raise (TRACE_EXIT state)) + | "quit\n" => raise TRACE_QUIT + | _ => (prs +"Type RETURN to continue or...\n\ +\ f - to fail here\n\ +\ o - to switch tracing off\n\ +\ x - to exit at this point\n\ +\ quit - to abort this tracing run\n\ +\** Well? " ; exec_trace_command flag (tf, state)); + + +(*Extract from a tactic, a thm->thm seq function that handles tracing*) +fun tracify flag (Tactic tf) state = + if !flag then (print_goals (!goals_limit) state; + prs"** Press RETURN to continue: "; + exec_trace_command flag (tf,state)) + else tf state; + +(*Create a tactic whose outcome is given by seqf, handling TRACE_EXIT*) +fun traced_tac seqf = Tactic (fn st => + Sequence.seqof (fn()=> seqf st + handle TRACE_EXIT st' => Some(st', Sequence.null))); + + +(*Tracing flags*) +val trace_REPEAT= ref false +and trace_DEPTH_FIRST = ref false +and trace_BEST_FIRST = ref false; + +(*Deterministic REPEAT: only retains the first outcome; + uses less space than REPEAT; tail recursive*) +fun REPEAT_DETERM tac = + let val tf = tracify trace_REPEAT tac + fun drep st = + case Sequence.pull(tf st) of + None => Some(st, Sequence.null) + | Some(st',_) => drep st' + in traced_tac drep end; + +(*General REPEAT: maintains a stack of alternatives; tail recursive*) +fun REPEAT tac = + let val tf = tracify trace_REPEAT tac + fun rep qs st = + case Sequence.pull(tf st) of + None => Some(st, Sequence.seqof(fn()=> repq qs)) + | Some(st',q) => rep (q::qs) st' + and repq [] = None + | repq(q::qs) = case Sequence.pull q of + None => repq qs + | Some(st,q) => rep (q::qs) st + in traced_tac (rep []) end; + +(*Repeat 1 or more times*) +fun REPEAT1 tac = tac THEN REPEAT tac; + + +(** Search tacticals **) + +(*Seaarches "satp" reports proof tree as satisfied*) +fun DEPTH_FIRST satp tac = + let val tf = tracify trace_DEPTH_FIRST tac + fun depth [] = None + | depth(q::qs) = + case Sequence.pull q of + None => depth qs + | Some(st,stq) => + if satp st then Some(st, Sequence.seqof(fn()=> depth(stq::qs))) + else depth (tf st :: stq :: qs) + in traced_tac (fn st => depth([Sequence.single st])) end; + + +(*Predicate: Does the rule have fewer than n premises?*) +fun has_fewer_prems n rule = (nprems_of rule < n); + +(*Apply a tactic if subgoals remain, else do nothing.*) +val IF_UNSOLVED = COND (has_fewer_prems 1) all_tac; + +(*Tactical to reduce the number of premises by 1. + If no subgoals then it must fail! *) +fun DEPTH_SOLVE_1 tac = STATE + (fn state => + (case nprems_of state of + 0 => no_tac + | n => DEPTH_FIRST (has_fewer_prems n) tac)); + +(*Uses depth-first search to solve ALL subgoals*) +val DEPTH_SOLVE = DEPTH_FIRST (has_fewer_prems 1); + +(*** Best-first search ***) + +(*Insertion into priority queue of states *) +fun insert (nth: int*thm, []) = [nth] + | insert ((m,th), (n,th')::nths) = + if n some_of_list l)); + + +(* Best-first search for a state that satisfies satp (incl initial state) + Function sizef estimates size of problem remaining (smaller means better). + tactic tf0 sets up the initial priority queue, which is searched by tac. *) +fun (Tactic tf0) THEN_BEST_FIRST (satp, sizef, tac) = + let val tf = tracify trace_BEST_FIRST tac + fun pairsize th = (sizef th, th); + fun bfs (news,nprfs) = + (case partition satp news of + ([],nonsats) => next(foldr insert + (map pairsize nonsats, nprfs)) + | (sats,_) => some_of_list sats) + and next [] = None + | next ((n,prf)::nprfs) = + (if !trace_BEST_FIRST + then writeln("state size = " ^ string_of_int n ^ + " queue length =" ^ string_of_int (length nprfs)) + else (); + bfs (Sequence.list_of_s (tf prf), nprfs)) + fun tf st = bfs (Sequence.list_of_s (tf0 st), []) + in traced_tac tf end; + +(*Ordinary best-first search, with no initial tactic*) +fun BEST_FIRST (satp,sizef) tac = all_tac THEN_BEST_FIRST (satp,sizef,tac); + +(*Breadth-first search to satisfy satpred (including initial state) + SLOW -- SHOULD NOT USE APPEND!*) +fun BREADTH_FIRST satpred (Tactic tf) = + let val tacf = Sequence.list_of_s o tf; + fun bfs prfs = + (case partition satpred prfs of + ([],[]) => [] + | ([],nonsats) => + (prs("breadth=" ^ string_of_int(length nonsats) ^ "\n"); + bfs (flat (map tacf nonsats))) + | (sats,_) => sats) + in Tactic (fn state => Sequence.s_of_list (bfs [state])) end; + + +(** Filtering tacticals **) + +(*Returns all states satisfying the predicate*) +fun FILTER pred (Tactic tf) = Tactic + (fn state => Sequence.filters pred (tf state)); + +(*Returns all changed states*) +fun CHANGED (Tactic tf) = + Tactic (fn state => + let fun diff st = not (eq_thm(state,st)) + in Sequence.filters diff (tf state) + end ); + + +(*** Tacticals based on subgoal numbering ***) + +(*For n subgoals, performs tf(n) THEN ... THEN tf(1) + Essential to work backwards since tf(i) may add/delete subgoals at i. *) +fun ALLGOALS tf = + let fun tac 0 = all_tac + | tac n = tf(n) THEN tac(n-1) + in Tactic(fn state => tapply(tac(nprems_of state), state)) end; + +(*For n subgoals, performs tf(n) ORELSE ... ORELSE tf(1) *) +fun SOMEGOAL tf = + let fun tac 0 = no_tac + | tac n = tf(n) ORELSE tac(n-1) + in Tactic(fn state => tapply(tac(nprems_of state), state)) end; + +(*For n subgoals, performs tf(1) ORELSE ... ORELSE tf(n). + More appropriate than SOMEGOAL in some cases.*) +fun FIRSTGOAL tf = + let fun tac (i,n) = if i>n then no_tac else tf(i) ORELSE tac (i+1,n) + in Tactic(fn state => tapply(tac(1, nprems_of state), state)) end; + +(*Repeatedly solve some using tf. *) +fun REPEAT_SOME tf = REPEAT1 (SOMEGOAL (REPEAT1 o tf)); + +(*Repeatedly solve the first possible subgoal using tf. *) +fun REPEAT_FIRST tf = REPEAT1 (FIRSTGOAL (REPEAT1 o tf)); + +(*For n subgoals, tries to apply tf to n,...1 *) +fun TRYALL tf = ALLGOALS (TRY o tf); + + +(*Make a tactic for subgoal i, if there is one. *) +fun SUBGOAL goalfun i = Tactic(fn state => + case drop(i-1, prems_of state) of + [] => Sequence.null + | prem::_ => tapply(goalfun (prem,i), state)); + +(*Tactical for restricting the effect of a tactic to subgoal i. + Works by making a new state from subgoal i, applying tf to it, and + composing the resulting metathm with the original state. + The "main goal" of the new state will not be atomic, some tactics may fail! + DOES NOT work if tactic affects the main goal other than by instantiation.*) + +(* (!!x. ?V) ==> ?V ; used by protect_subgoal.*) +val dummy_quant_rl = + standard (forall_elim_var 0 (assume + (Sign.read_cterm Sign.pure ("!!x. PROP V",propT)))); + +(* Prevent the subgoal's assumptions from becoming additional subgoals in the + new proof state by enclosing them by a universal quantification *) +fun protect_subgoal state i = + case Sequence.chop (1, bicompose false (false,dummy_quant_rl,1) i state) + of + ([state'],_) => state' + | _ => error"SELECT_GOAL -- impossible error???"; + +(*Does the work of SELECT_GOAL. *) +fun select (Tactic tf) state i = + let val prem::_ = drop(i-1, prems_of state) + val st0 = trivial (Sign.cterm_of (#sign(rep_thm state)) prem); + fun next st = bicompose false (false, st, nprems_of st) i state + in Sequence.flats (Sequence.maps next (tf st0)) + end; + +(*If i=1 and there is only one subgoal then do nothing!*) +fun SELECT_GOAL tac i = Tactic (fn state => + case (i, drop(i-1, prems_of state)) of + (_,[]) => Sequence.null + | (1,[_]) => tapply(tac,state) + | (_, (Const("==>",_)$_$_) :: _) => select tac (protect_subgoal state i) i + | (_, _::_) => select tac state i); + + + +(*Strips assumptions in goal yielding ( [x1,...,xm], [H1,...,Hn], B ) + H1,...,Hn are the hypotheses; x1...xm are variants of the parameters. + Main difference from strip_assums concerns parameters: + it replaces the bound variables by free variables. *) +fun strip_context_aux (params, Hs, Const("==>", _) $ H $ B) = + strip_context_aux (params, H::Hs, B) + | strip_context_aux (params, Hs, Const("all",_)$Abs(a,T,t)) = + let val (b,u) = variant_abs(a,T,t) + in strip_context_aux ((b,T)::params, Hs, u) end + | strip_context_aux (params, Hs, B) = (rev params, rev Hs, B); + +fun strip_context A = strip_context_aux ([],[],A); + + +(**** METAHYPS -- tactical for using hypotheses as meta-level assumptions + METAHYPS (fn prems => tac (prems)) i + +converts subgoal i, of the form !!x1...xm. [| A1;...;An] ==> A into a new +proof state A==>A, supplying A1,...,An as meta-level assumptions (in +"prems"). The parameters x1,...,xm become free variables. If the +resulting proof state is [| B1;...;Bk] ==> C (possibly assuming A1,...,An) +then it is lifted back into the original context, yielding k subgoals. + +Replaces unknowns in the context by Frees having the prefix METAHYP_ +New unknowns in [| B1;...;Bk] ==> C are lifted over x1,...,xm. +DOES NOT HANDLE TYPE UNKNOWNS. +****) + +local + open Logic + + (*Left-to-right replacements: ctpairs = [...,(vi,ti),...]. + Instantiates distinct free variables by terms of same type.*) + fun free_instantiate ctpairs = + forall_elim_list (map snd ctpairs) o forall_intr_list (map fst ctpairs); + + fun free_of s ((a,i), T) = + Free(s ^ (case i of 0 => a | _ => a ^ "_" ^ string_of_int i), + T) + + fun mk_inst (var as Var(v,T)) = (var, free_of "METAHYP1_" (v,T)) +in + +fun metahyps_aux_tac tacf (prem,i) = Tactic (fn state => + let val {sign,maxidx,...} = rep_thm state + val cterm = Sign.cterm_of sign + (*find all vars in the hyps -- should find tvars also!*) + val hyps_vars = foldr add_term_vars (strip_assums_hyp prem, []) + val insts = map mk_inst hyps_vars + (*replace the hyps_vars by Frees*) + val prem' = subst_atomic insts prem + val (params,hyps,concl) = strip_context prem' + val fparams = map Free params + val cparams = map cterm fparams + and chyps = map cterm hyps + val hypths = map assume chyps + fun swap_ctpair (t,u) = (cterm u, cterm t) + (*Subgoal variables: make Free; lift type over params*) + fun mk_subgoal_inst concl_vars (var as Var(v,T)) = + if var mem concl_vars + then (var, true, free_of "METAHYP2_" (v,T)) + else (var, false, + free_of "METAHYP2_" (v, map #2 params --->T)) + (*Instantiate subgoal vars by Free applied to params*) + fun mk_ctpair (t,in_concl,u) = + if in_concl then (cterm t, cterm u) + else (cterm t, cterm (list_comb (u,fparams))) + (*Restore Vars with higher type and index*) + fun mk_subgoal_swap_ctpair + (t as Var((a,i),_), in_concl, u as Free(_,U)) = + if in_concl then (cterm u, cterm t) + else (cterm u, cterm(Var((a, i+maxidx), U))) + (*Embed B in the original context of params and hyps*) + fun embed B = list_all_free (params, list_implies (hyps, B)) + (*Strip the context using elimination rules*) + fun elim Bhyp = implies_elim_list (forall_elim_list cparams Bhyp) hypths + (*Embed an ff pair in the original params*) + fun embed_ff(t,u) = + mk_flexpair (list_abs_free (params, t), list_abs_free (params, u)) + (*Remove parameter abstractions from the ff pairs*) + fun elim_ff ff = flexpair_abs_elim_list cparams ff + (*A form of lifting that discharges assumptions.*) + fun relift st = + let val prop = #prop(rep_thm st) + val subgoal_vars = (*Vars introduced in the subgoals*) + foldr add_term_vars (strip_imp_prems prop, []) + and concl_vars = add_term_vars (strip_imp_concl prop, []) + val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars + val st' = instantiate ([], map mk_ctpair subgoal_insts) st + val emBs = map (cterm o embed) (prems_of st') + and ffs = map (cterm o embed_ff) (tpairs_of st') + val Cth = implies_elim_list st' + (map (elim_ff o assume) ffs @ + map (elim o assume) emBs) + in (*restore the unknowns to the hypotheses*) + free_instantiate (map swap_ctpair insts @ + map mk_subgoal_swap_ctpair subgoal_insts) + (*discharge assumptions from state in same order*) + (implies_intr_list (ffs@emBs) + (forall_intr_list cparams (implies_intr_list chyps Cth))) + end + val subprems = map (forall_elim_vars 0) hypths + and st0 = trivial (cterm concl) + (*function to replace the current subgoal*) + fun next st = bicompose false (false, relift st, nprems_of st) + i state + in Sequence.flats (Sequence.maps next (tapply(tacf subprems, st0))) + end); +end; + +fun METAHYPS tacf = SUBGOAL (metahyps_aux_tac tacf); + +end; +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/term.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/term.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,549 @@ +(* Title: term.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright Cambridge University 1992 +*) + + +(*Simply typed lambda-calculus: types, terms, and basic operations*) + + +(*Indexnames can be quickly renamed by adding an offset to the integer part, + for resolution.*) +type indexname = string*int; + +(* Types are classified by classes. *) +type class = string; +type sort = class list; + +(* The sorts attached to TFrees and TVars specify the sort of that variable *) +datatype typ = Type of string * typ list + | TFree of string * sort + | TVar of indexname * sort; + +infixr 5 -->; +fun S --> T = Type("fun",[S,T]); + +(*handy for multiple args: [T1,...,Tn]--->T gives T1-->(T2--> ... -->T)*) +infixr --->; +val op ---> = foldr (op -->); + + +(*terms. Bound variables are indicated by depth number. + Free variables, (scheme) variables and constants have names. + An term is "closed" if there every bound variable of level "lev" + is enclosed by at least "lev" abstractions. + + It is possible to create meaningless terms containing loose bound vars + or type mismatches. But such terms are not allowed in rules. *) + + + +infix 9 $; (*application binds tightly!*) +datatype term = + Const of string * typ + | Free of string * typ + | Var of indexname * typ + | Bound of int + | Abs of string*typ*term + | op $ of term*term; + + +(*For errors involving type mismatches*) +exception TYPE of string * typ list * term list; + +(*For system errors involving terms*) +exception TERM of string * term list; + + +(*Note variable naming conventions! + a,b,c: string + f,g,h: functions (including terms of function type) + i,j,m,n: int + t,u: term + v,w: indexnames + x,y: any + A,B,C: term (denoting formulae) + T,U: typ +*) + + +(** Discriminators **) + +fun is_Const (Const _) = true + | is_Const _ = false; + +fun is_Free (Free _) = true + | is_Free _ = false; + +fun is_Var (Var _) = true + | is_Var _ = false; + +fun is_TVar (TVar _) = true + | is_TVar _ = false; + +(** Destructors **) + +fun dest_Const (Const x) = x + | dest_Const t = raise TERM("dest_Const", [t]); + +fun dest_Free (Free x) = x + | dest_Free t = raise TERM("dest_Free", [t]); + +fun dest_Var (Var x) = x + | dest_Var t = raise TERM("dest_Var", [t]); + + +(* maps [T1,...,Tn]--->T to the list [T1,T2,...,Tn]*) +fun binder_types (Type("fun",[S,T])) = S :: binder_types T + | binder_types _ = []; + +(* maps [T1,...,Tn]--->T to T*) +fun body_type (Type("fun",[S,T])) = body_type T + | body_type T = T; + +(* maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T) *) +fun strip_type T : typ list * typ = + (binder_types T, body_type T); + + +(*Compute the type of the term, checking that combinations are well-typed + Ts = [T0,T1,...] holds types of bound variables 0, 1, ...*) +fun type_of1 (Ts, Const (_,T)) = T + | type_of1 (Ts, Free (_,T)) = T + | type_of1 (Ts, Bound i) = (nth_elem (i,Ts) + handle LIST _ => raise TYPE("type_of: bound variable", [], [Bound i])) + | type_of1 (Ts, Var (_,T)) = T + | type_of1 (Ts, Abs (_,T,body)) = T --> type_of1(T::Ts, body) + | type_of1 (Ts, f$u) = + let val U = type_of1(Ts,u) + and T = type_of1(Ts,f) + in case T of + Type("fun",[T1,T2]) => + if T1=U then T2 else raise TYPE + ("type_of: type mismatch in application", [T1,U], [f$u]) + | _ => raise TYPE ("type_of: Rator must have function type", + [T,U], [f$u]) + end; + + +fun type_of t : typ = type_of1 ([],t); + +(*Determines the type of a term, with minimal checking*) +fun fastype_of(Ts, f$u) = (case fastype_of(Ts,f) of + Type("fun",[_,T]) => T + | _ => raise TERM("fastype_of: expected function type", [f$u])) + | fastype_of(_, Const (_,T)) = T + | fastype_of(_, Free (_,T)) = T + | fastype_of(Ts, Bound i) = (nth_elem(i,Ts) + handle LIST _ => raise TERM("fastype_of: Bound", [Bound i])) + | fastype_of(_, Var (_,T)) = T + | fastype_of(Ts, Abs (_,T,u)) = T --> fastype_of(T::Ts, u); + + +(* maps (x1,...,xn)t to t *) +fun strip_abs_body (Abs(_,_,t)) = strip_abs_body t + | strip_abs_body u = u; + + +(* maps (x1,...,xn)t to [x1, ..., xn] *) +fun strip_abs_vars (Abs(a,T,t)) = (a,T) :: strip_abs_vars t + | strip_abs_vars u = [] : (string*typ) list; + + +fun strip_qnt_body qnt = +let fun strip(tm as Const(c,_)$Abs(_,_,t)) = if c=qnt then strip t else tm + | strip t = t +in strip end; + +fun strip_qnt_vars qnt = +let fun strip(Const(c,_)$Abs(a,T,t)) = if c=qnt then (a,T)::strip t else [] + | strip t = [] : (string*typ) list +in strip end; + + +(* maps (f, [t1,...,tn]) to f(t1,...,tn) *) +val list_comb : term * term list -> term = foldl (op $); + + +(* maps f(t1,...,tn) to (f, [t1,...,tn]) ; naturally tail-recursive*) +fun strip_comb u : term * term list = + let fun stripc (f$t, ts) = stripc (f, t::ts) + | stripc x = x + in stripc(u,[]) end; + + +(* maps f(t1,...,tn) to f , which is never a combination *) +fun head_of (f$t) = head_of f + | head_of u = u; + + +(*Number of atoms and abstractions in a term*) +fun size_of_term (Abs (_,_,body)) = 1 + size_of_term body + | size_of_term (f$t) = size_of_term f + size_of_term t + | size_of_term _ = 1; + + +(* apply a function to all types in a term *) +fun map_term_types f = +let fun map(Const(a,T)) = Const(a, f T) + | map(Free(a,T)) = Free(a, f T) + | map(Var(v,T)) = Var(v, f T) + | map(t as Bound _) = t + | map(Abs(a,T,t)) = Abs(a, f T, map t) + | map(f$t) = map f $ map t; +in map end; + +(* iterate a function over all types in a term *) +fun it_term_types f = +let fun iter(Const(_,T), a) = f(T,a) + | iter(Free(_,T), a) = f(T,a) + | iter(Var(_,T), a) = f(T,a) + | iter(Abs(_,T,t), a) = iter(t,f(T,a)) + | iter(f$u, a) = iter(f, iter(u, a)) + | iter(Bound _, a) = a +in iter end + + +(** Connectives of higher order logic **) + +val propT : typ = Type("prop",[]); + +val implies = Const("==>", propT-->propT-->propT); + +fun all T = Const("all", (T-->propT)-->propT); + +fun equals T = Const("==", T-->T-->propT); + +fun flexpair T = Const("=?=", T-->T-->propT); + +(* maps !!x1...xn. t to t *) +fun strip_all_body (Const("all",_)$Abs(_,_,t)) = strip_all_body t + | strip_all_body t = t; + +(* maps !!x1...xn. t to [x1, ..., xn] *) +fun strip_all_vars (Const("all",_)$Abs(a,T,t)) = + (a,T) :: strip_all_vars t + | strip_all_vars t = [] : (string*typ) list; + +(*increments a term's non-local bound variables + required when moving a term within abstractions + inc is increment for bound variables + lev is level at which a bound variable is considered 'loose'*) +fun incr_bv (inc, lev, u as Bound i) = if i>=lev then Bound(i+inc) else u + | incr_bv (inc, lev, Abs(a,T,body)) = + Abs(a, T, incr_bv(inc,lev+1,body)) + | incr_bv (inc, lev, f$t) = + incr_bv(inc,lev,f) $ incr_bv(inc,lev,t) + | incr_bv (inc, lev, u) = u; + +fun incr_boundvars 0 t = t + | incr_boundvars inc t = incr_bv(inc,0,t); + + +(*Accumulate all 'loose' bound vars referring to level 'lev' or beyond. + (Bound 0) is loose at level 0 *) +fun add_loose_bnos (Bound i, lev, js) = + if i= k + | loose_bvar(f$t, k) = loose_bvar(f,k) orelse loose_bvar(t,k) + | loose_bvar(Abs(_,_,t),k) = loose_bvar(t,k+1) + | loose_bvar _ = false; + + +(*Substitute arguments for loose bound variables. + Beta-reduction of arg(n-1)...arg0 into t replacing (Bound i) with (argi). + Note that for ((x,y)c)(a,b), the bound vars in c are x=1 and y=0 + and the appropriate call is subst_bounds([b,a], c) . + Loose bound variables >=n are reduced by "n" to + compensate for the disappearance of lambdas. +*) +fun subst_bounds (args: term list, t) : term = + let val n = length args; + fun subst (t as Bound i, lev) = + if i Bound(i-n) (*loose: change it*) + | arg::_ => incr_boundvars lev arg) + | subst (Abs(a,T,body), lev) = Abs(a, T, subst(body,lev+1)) + | subst (f$t, lev) = subst(f,lev) $ subst(t,lev) + | subst (t,lev) = t + in case args of [] => t | _ => subst (t,0) end; + +(*beta-reduce if possible, else form application*) +fun betapply (Abs(_,_,t), u) = subst_bounds([u],t) + | betapply (f,u) = f$u; + +(*Tests whether 2 terms are alpha-convertible and have same type. + Note that constants and Vars may have more than one type.*) +infix aconv; +fun (Const(a,T)) aconv (Const(b,U)) = a=b andalso T=U + | (Free(a,T)) aconv (Free(b,U)) = a=b andalso T=U + | (Var(v,T)) aconv (Var(w,U)) = v=w andalso T=U + | (Bound i) aconv (Bound j) = i=j + | (Abs(_,T,t)) aconv (Abs(_,U,u)) = t aconv u andalso T=U + | (f$t) aconv (g$u) = (f aconv g) andalso (t aconv u) + | _ aconv _ = false; + +(*are two term lists alpha-convertible in corresponding elements?*) +fun aconvs ([],[]) = true + | aconvs (t::ts, u::us) = t aconv u andalso aconvs(ts,us) + | aconvs _ = false; + +(*A fast unification filter: true unless the two terms cannot be unified. + Terms must be NORMAL. Treats all Vars as distinct. *) +fun could_unify (t,u) = + let fun matchrands (f$t, g$u) = could_unify(t,u) andalso matchrands(f,g) + | matchrands _ = true + in case (head_of t , head_of u) of + (_, Var _) => true + | (Var _, _) => true + | (Const(a,_), Const(b,_)) => a=b andalso matchrands(t,u) + | (Free(a,_), Free(b,_)) => a=b andalso matchrands(t,u) + | (Bound i, Bound j) => i=j andalso matchrands(t,u) + | (Abs _, _) => true (*because of possible eta equality*) + | (_, Abs _) => true + | _ => false + end; + +(*Substitute new for free occurrences of old in a term*) +fun subst_free [] = (fn t=>t) + | subst_free pairs = + let fun substf u = + case gen_assoc (op aconv) (pairs, u) of + Some u' => u' + | None => (case u of Abs(a,T,t) => Abs(a, T, substf t) + | t$u' => substf t $ substf u' + | _ => u) + in substf end; + +(*a total, irreflexive ordering on index names*) +fun xless ((a,i), (b,j): indexname) = i Abs(a, T, abst(lev+1, t)) + | f$rand => abst(lev,f) $ abst(lev,rand) + | _ => u) + in abst(0,body) end; + + +(*Form an abstraction over a free variable.*) +fun absfree (a,T,body) = Abs(a, T, abstract_over (Free(a,T), body)); + +(*Abstraction over a list of free variables*) +fun list_abs_free ([ ] , t) = t + | list_abs_free ((a,T)::vars, t) = + absfree(a, T, list_abs_free(vars,t)); + +(*Quantification over a list of free variables*) +fun list_all_free ([], t: term) = t + | list_all_free ((a,T)::vars, t) = + (all T) $ (absfree(a, T, list_all_free(vars,t))); + +(*Quantification over a list of variables (already bound in body) *) +fun list_all ([], t) = t + | list_all ((a,T)::vars, t) = + (all T) $ (Abs(a, T, list_all(vars,t))); + +(*Replace the ATOMIC term ti by ui; instl = [(t1,u1), ..., (tn,un)]. + A simultaneous substitution: [ (a,b), (b,a) ] swaps a and b. *) +fun subst_atomic [] t = t : term + | subst_atomic (instl: (term*term) list) t = + let fun subst (Abs(a,T,body)) = Abs(a, T, subst body) + | subst (f$t') = subst f $ subst t' + | subst t = (case assoc(instl,t) of + Some u => u | None => t) + in subst t end; + +fun typ_subst_TVars iTs T = if null iTs then T else + let fun subst(Type(a,Ts)) = Type(a, map subst Ts) + | subst(T as TFree _) = T + | subst(T as TVar(ixn,_)) = + (case assoc(iTs,ixn) of None => T | Some(U) => U) + in subst T end; + +val subst_TVars = map_term_types o typ_subst_TVars; + +fun subst_Vars itms t = if null itms then t else + let fun subst(v as Var(ixn,_)) = + (case assoc(itms,ixn) of None => v | Some t => t) + | subst(Abs(a,T,t)) = Abs(a,T,subst t) + | subst(f$t) = subst f $ subst t + | subst(t) = t + in subst t end; + +fun subst_vars(iTs,itms) = if null iTs then subst_Vars itms else + let fun subst(Const(a,T)) = Const(a,typ_subst_TVars iTs T) + | subst(Free(a,T)) = Free(a,typ_subst_TVars iTs T) + | subst(v as Var(ixn,T)) = (case assoc(itms,ixn) of + None => Var(ixn,typ_subst_TVars iTs T) + | Some t => t) + | subst(b as Bound _) = b + | subst(Abs(a,T,t)) = Abs(a,typ_subst_TVars iTs T,subst t) + | subst(f$t) = subst f $ subst t + in subst end; + + +(*Computing the maximum index of a typ*) +fun maxidx_of_typ(Type(_,Ts)) = + if Ts=[] then ~1 else max(map maxidx_of_typ Ts) + | maxidx_of_typ(TFree _) = ~1 + | maxidx_of_typ(TVar((_,i),_)) = i; + + +(*Computing the maximum index of a term*) +fun maxidx_of_term (Const(_,T)) = maxidx_of_typ T + | maxidx_of_term (Bound _) = ~1 + | maxidx_of_term (Free(_,T)) = maxidx_of_typ T + | maxidx_of_term (Var ((_,i), T)) = max[i, maxidx_of_typ T] + | maxidx_of_term (Abs (_,T,body)) = max[maxidx_of_term body, maxidx_of_typ T] + | maxidx_of_term (f$t) = max [maxidx_of_term f, maxidx_of_term t]; + + +(* Increment the index of all Poly's in T by k *) +fun incr_tvar k (Type(a,Ts)) = Type(a, map (incr_tvar k) Ts) + | incr_tvar k (T as TFree _) = T + | incr_tvar k (TVar((a,i),rs)) = TVar((a,i+k),rs); + + +(**** Syntax-related declarations ****) + + +(*Dummy type for parsing. Will be replaced during type inference. *) +val dummyT = Type("dummy",[]); + +(*scan a numeral of the given radix, normally 10*) +fun scan_radixint (radix: int, cs) : int * string list = + let val zero = ord"0" + val limit = zero+radix + fun scan (num,[]) = (num,[]) + | scan (num, c::cs) = + if zero <= ord c andalso ord c < limit + then scan(radix*num + ord c - zero, cs) + else (num, c::cs) + in scan(0,cs) end; + +fun scan_int cs = scan_radixint(10,cs); + + +(*** Printing ***) + + +(*Makes a variant of the name c distinct from the names in bs. + First attaches the suffix "a" and then increments this. *) +fun variant bs c : string = + let fun vary2 c = if (c mem bs) then vary2 (bump_string c) else c + fun vary1 c = if (c mem bs) then vary2 (c ^ "a") else c + in vary1 (if c="" then "u" else c) end; + +(*Create variants of the list of names, with priority to the first ones*) +fun variantlist ([], used) = [] + | variantlist(b::bs, used) = + let val b' = variant used b + in b' :: variantlist (bs, b'::used) end; + +(** TFrees and TVars **) + +(*maps (bs,v) to v'::bs this reverses the identifiers bs*) +fun add_new_id (bs, c) : string list = variant bs c :: bs; + +(*Accumulates the names in the term, suppressing duplicates. + Includes Frees and Consts. For choosing unambiguous bound var names.*) +fun add_term_names (Const(a,_), bs) = a ins bs + | add_term_names (Free(a,_), bs) = a ins bs + | add_term_names (f$u, bs) = add_term_names (f, add_term_names(u, bs)) + | add_term_names (Abs(_,_,t), bs) = add_term_names(t,bs) + | add_term_names (_, bs) = bs; + +(*Accumulates the TVars in a type, suppressing duplicates. *) +fun add_typ_tvars(Type(_,Ts),vs) = foldr add_typ_tvars (Ts,vs) + | add_typ_tvars(TFree(_),vs) = vs + | add_typ_tvars(TVar(v),vs) = v ins vs; + +(*Accumulates the TFrees in a type, suppressing duplicates. *) +fun add_typ_tfree_names(Type(_,Ts),fs) = foldr add_typ_tfree_names (Ts,fs) + | add_typ_tfree_names(TFree(f,_),fs) = f ins fs + | add_typ_tfree_names(TVar(_),fs) = fs; + +fun add_typ_tfrees(Type(_,Ts),fs) = foldr add_typ_tfrees (Ts,fs) + | add_typ_tfrees(TFree(f),fs) = f ins fs + | add_typ_tfrees(TVar(_),fs) = fs; + +(*Accumulates the TVars in a term, suppressing duplicates. *) +val add_term_tvars = it_term_types add_typ_tvars; +val add_term_tvar_ixns = (map #1) o (it_term_types add_typ_tvars); + +(*Accumulates the TFrees in a term, suppressing duplicates. *) +val add_term_tfrees = it_term_types add_typ_tfrees; +val add_term_tfree_names = it_term_types add_typ_tfree_names; + +(*Non-list versions*) +fun typ_tfrees T = add_typ_tfrees(T,[]); +fun typ_tvars T = add_typ_tvars(T,[]); +fun term_tfrees t = add_term_tfrees(t,[]); +fun term_tvars t = add_term_tvars(t,[]); + +(** Frees and Vars **) + +(*a partial ordering (not reflexive) for atomic terms*) +fun atless (Const (a,_), Const (b,_)) = a insert_aterm(t,vars) + | Abs (_,_,body) => add_term_vars(body,vars) + | f$t => add_term_vars (f, add_term_vars(t, vars)) + | _ => vars; + +fun term_vars t = add_term_vars(t,[]); + +(*Accumulates the Frees in the term, suppressing duplicates*) +fun add_term_frees (t, frees: term list) = case t of + Free _ => insert_aterm(t,frees) + | Abs (_,_,body) => add_term_frees(body,frees) + | f$t => add_term_frees (f, add_term_frees(t, frees)) + | _ => frees; + +fun term_frees t = add_term_frees(t,[]); + +(*Given an abstraction over P, replaces the bound variable by a Free variable + having a unique name. *) +fun variant_abs (a,T,P) = + let val b = variant (add_term_names(P,[])) a + in (b, subst_bounds ([Free(b,T)], P)) end; + +(* renames and reverses the strings in vars away from names *) +fun rename_aTs names vars : (string*typ)list = + let fun rename_aT (vars,(a,T)) = + (variant (map #1 vars @ names) a, T) :: vars + in foldl rename_aT ([],vars) end; + +fun rename_wrt_term t = rename_aTs (add_term_names(t,[])); diff -r 000000000000 -r a5a9c433f639 src/Pure/thm.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/thm.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,990 @@ +(* Title: thm + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +The abstract types "theory" and "thm" +*) + +signature THM = + sig + structure Envir : ENVIR + structure Sequence : SEQUENCE + structure Sign : SIGN + type meta_simpset + type theory + type thm + exception THM of string * int * thm list + exception THEORY of string * theory list + exception SIMPLIFIER of string * thm + val abstract_rule: string -> Sign.cterm -> thm -> thm + val add_congs: meta_simpset * thm list -> meta_simpset + val add_prems: meta_simpset * thm list -> meta_simpset + val add_simps: meta_simpset * thm list -> meta_simpset + val assume: Sign.cterm -> thm + val assumption: int -> thm -> thm Sequence.seq + val axioms_of: theory -> (string * thm) list + val beta_conversion: Sign.cterm -> thm + val bicompose: bool -> bool * thm * int -> int -> thm -> thm Sequence.seq + val biresolution: bool -> (bool*thm)list -> int -> thm -> thm Sequence.seq + val combination: thm -> thm -> thm + val concl_of: thm -> term + val dest_state: thm * int -> (term*term)list * term list * term * term + val empty_mss: meta_simpset + val eq_assumption: int -> thm -> thm + val equal_intr: thm -> thm -> thm + val equal_elim: thm -> thm -> thm + val extend_theory: theory -> string + -> (class * class list) list * sort + * (string list * int)list + * (string list * (sort list * class))list + * (string list * string)list * Sign.Syntax.sext option + -> (string*string)list -> theory + val extensional: thm -> thm + val flexflex_rule: thm -> thm Sequence.seq + val flexpair_def: thm + val forall_elim: Sign.cterm -> thm -> thm + val forall_intr: Sign.cterm -> thm -> thm + val freezeT: thm -> thm + val get_axiom: theory -> string -> thm + val implies_elim: thm -> thm -> thm + val implies_intr: Sign.cterm -> thm -> thm + val implies_intr_hyps: thm -> thm + val instantiate: (indexname*Sign.ctyp)list * (Sign.cterm*Sign.cterm)list + -> thm -> thm + val lift_rule: (thm * int) -> thm -> thm + val merge_theories: theory * theory -> theory + val mk_rews_of_mss: meta_simpset -> thm -> thm list + val mss_of: thm list -> meta_simpset + val nprems_of: thm -> int + val parents_of: theory -> theory list + val prems_of: thm -> term list + val prems_of_mss: meta_simpset -> thm list + val pure_thy: theory + val reflexive: Sign.cterm -> thm + val rename_params_rule: string list * int -> thm -> thm + val rep_thm: thm -> {prop: term, hyps: term list, maxidx: int, sign: Sign.sg} + val rewrite_cterm: meta_simpset -> (meta_simpset -> thm -> thm option) + -> Sign.cterm -> thm + val set_mk_rews: meta_simpset * (thm -> thm list) -> meta_simpset + val sign_of: theory -> Sign.sg + val syn_of: theory -> Sign.Syntax.syntax + val stamps_of_thm: thm -> string ref list + val stamps_of_thy: theory -> string ref list + val symmetric: thm -> thm + val tpairs_of: thm -> (term*term)list + val trace_simp: bool ref + val transitive: thm -> thm -> thm + val trivial: Sign.cterm -> thm + val varifyT: thm -> thm + end; + + + +functor ThmFun (structure Logic: LOGIC and Unify: UNIFY and Pattern:PATTERN + and Net:NET + sharing type Pattern.type_sig = Unify.Sign.Type.type_sig) + : THM = +struct +structure Sequence = Unify.Sequence; +structure Envir = Unify.Envir; +structure Sign = Unify.Sign; +structure Type = Sign.Type; +structure Syntax = Sign.Syntax; +structure Symtab = Sign.Symtab; + + +(*Meta-theorems*) +datatype thm = Thm of + {sign: Sign.sg, maxidx: int, hyps: term list, prop: term}; + +fun rep_thm (Thm x) = x; + +(*Errors involving theorems*) +exception THM of string * int * thm list; + +(*maps object-rule to tpairs *) +fun tpairs_of (Thm{prop,...}) = #1 (Logic.strip_flexpairs prop); + +(*maps object-rule to premises *) +fun prems_of (Thm{prop,...}) = + Logic.strip_imp_prems (Logic.skip_flexpairs prop); + +(*counts premises in a rule*) +fun nprems_of (Thm{prop,...}) = + Logic.count_prems (Logic.skip_flexpairs prop, 0); + +(*maps object-rule to conclusion *) +fun concl_of (Thm{prop,...}) = Logic.strip_imp_concl prop; + +(*Stamps associated with a signature*) +val stamps_of_thm = #stamps o Sign.rep_sg o #sign o rep_thm; + +(*Theories. There is one pure theory. + A theory can be extended. Two theories can be merged.*) +datatype theory = + Pure of {sign: Sign.sg} + | Extend of {sign: Sign.sg, axioms: thm Symtab.table, thy: theory} + | Merge of {sign: Sign.sg, thy1: theory, thy2: theory}; + +(*Errors involving theories*) +exception THEORY of string * theory list; + +fun sign_of (Pure {sign}) = sign + | sign_of (Extend {sign,...}) = sign + | sign_of (Merge {sign,...}) = sign; + +val syn_of = #syn o Sign.rep_sg o sign_of; + +(*return the axioms of a theory and its ancestors*) +fun axioms_of (Pure _) = [] + | axioms_of (Extend{axioms,thy,...}) = Symtab.alist_of axioms + | axioms_of (Merge{thy1,thy2,...}) = axioms_of thy1 @ axioms_of thy2; + +(*return the immediate ancestors -- also distinguishes the kinds of theories*) +fun parents_of (Pure _) = [] + | parents_of (Extend{thy,...}) = [thy] + | parents_of (Merge{thy1,thy2,...}) = [thy1,thy2]; + + +(*Merge theories of two theorems. Raise exception if incompatible. + Prefers (via Sign.merge) the signature of th1. *) +fun merge_theories(th1,th2) = + let val Thm{sign=sign1,...} = th1 and Thm{sign=sign2,...} = th2 + in Sign.merge (sign1,sign2) end + handle TERM _ => raise THM("incompatible signatures", 0, [th1,th2]); + +(*Stamps associated with a theory*) +val stamps_of_thy = #stamps o Sign.rep_sg o sign_of; + + +(**** Primitive rules ****) + +(* discharge all assumptions t from ts *) +val disch = gen_rem (op aconv); + +(*The assumption rule A|-A in a theory *) +fun assume ct : thm = + let val {sign, t=prop, T, maxidx} = Sign.rep_cterm ct + in if T<>propT then + raise THM("assume: assumptions must have type prop", 0, []) + else if maxidx <> ~1 then + raise THM("assume: assumptions may not contain scheme variables", + maxidx, []) + else Thm{sign = sign, maxidx = ~1, hyps = [prop], prop = prop} + end; + +(* Implication introduction + A |- B + ------- + A ==> B *) +fun implies_intr cA (thB as Thm{sign,maxidx,hyps,prop}) : thm = + let val {sign=signA, t=A, T, maxidx=maxidxA} = Sign.rep_cterm cA + in if T<>propT then + raise THM("implies_intr: assumptions must have type prop", 0, [thB]) + else Thm{sign= Sign.merge (sign,signA), maxidx= max[maxidxA, maxidx], + hyps= disch(hyps,A), prop= implies$A$prop} + handle TERM _ => + raise THM("implies_intr: incompatible signatures", 0, [thB]) + end; + +(* Implication elimination + A ==> B A + --------------- + B *) +fun implies_elim thAB thA : thm = + let val Thm{maxidx=maxA, hyps=hypsA, prop=propA,...} = thA + and Thm{sign, maxidx, hyps, prop,...} = thAB; + fun err(a) = raise THM("implies_elim: "^a, 0, [thAB,thA]) + in case prop of + imp$A$B => + if imp=implies andalso A aconv propA + then Thm{sign= merge_theories(thAB,thA), + maxidx= max[maxA,maxidx], + hyps= hypsA union hyps, (*dups suppressed*) + prop= B} + else err("major premise") + | _ => err("major premise") + end; + +(* Forall introduction. The Free or Var x must not be free in the hypotheses. + A + ------ + !!x.A *) +fun forall_intr cx (th as Thm{sign,maxidx,hyps,prop}) = + let val x = Sign.term_of cx; + fun result(a,T) = Thm{sign= sign, maxidx= maxidx, hyps= hyps, + prop= all(T) $ Abs(a, T, abstract_over (x,prop))} + in case x of + Free(a,T) => + if exists (apl(x, Logic.occs)) hyps + then raise THM("forall_intr: variable free in assumptions", 0, [th]) + else result(a,T) + | Var((a,_),T) => result(a,T) + | _ => raise THM("forall_intr: not a variable", 0, [th]) + end; + +(* Forall elimination + !!x.A + -------- + A[t/x] *) +fun forall_elim ct (th as Thm{sign,maxidx,hyps,prop}) : thm = + let val {sign=signt, t, T, maxidx=maxt} = Sign.rep_cterm ct + in case prop of + Const("all",Type("fun",[Type("fun",[qary,_]),_])) $ A => + if T<>qary then + raise THM("forall_elim: type mismatch", 0, [th]) + else Thm{sign= Sign.merge(sign,signt), + maxidx= max[maxidx, maxt], + hyps= hyps, prop= betapply(A,t)} + | _ => raise THM("forall_elim: not quantified", 0, [th]) + end + handle TERM _ => + raise THM("forall_elim: incompatible signatures", 0, [th]); + + +(*** Equality ***) + +(*Definition of the relation =?= *) +val flexpair_def = + Thm{sign= Sign.pure, hyps= [], maxidx= 0, + prop= Sign.term_of + (Sign.read_cterm Sign.pure + ("(?t =?= ?u) == (?t == ?u::?'a::{})", propT))}; + +(*The reflexivity rule: maps t to the theorem t==t *) +fun reflexive ct = + let val {sign, t, T, maxidx} = Sign.rep_cterm ct + in Thm{sign= sign, hyps= [], maxidx= maxidx, prop= Logic.mk_equals(t,t)} + end; + +(*The symmetry rule + t==u + ---- + u==t *) +fun symmetric (th as Thm{sign,hyps,prop,maxidx}) = + case prop of + (eq as Const("==",_)) $ t $ u => + Thm{sign=sign, hyps=hyps, maxidx=maxidx, prop= eq$u$t} + | _ => raise THM("symmetric", 0, [th]); + +(*The transitive rule + t1==u u==t2 + ------------ + t1==t2 *) +fun transitive th1 th2 = + let val Thm{maxidx=max1, hyps=hyps1, prop=prop1,...} = th1 + and Thm{maxidx=max2, hyps=hyps2, prop=prop2,...} = th2; + fun err(msg) = raise THM("transitive: "^msg, 0, [th1,th2]) + in case (prop1,prop2) of + ((eq as Const("==",_)) $ t1 $ u, Const("==",_) $ u' $ t2) => + if not (u aconv u') then err"middle term" else + Thm{sign= merge_theories(th1,th2), hyps= hyps1 union hyps2, + maxidx= max[max1,max2], prop= eq$t1$t2} + | _ => err"premises" + end; + +(*Beta-conversion: maps (%(x)t)(u) to the theorem (%(x)t)(u) == t[u/x] *) +fun beta_conversion ct = + let val {sign, t, T, maxidx} = Sign.rep_cterm ct + in case t of + Abs(_,_,bodt) $ u => + Thm{sign= sign, hyps= [], + maxidx= maxidx_of_term t, + prop= Logic.mk_equals(t, subst_bounds([u],bodt))} + | _ => raise THM("beta_conversion: not a redex", 0, []) + end; + +(*The extensionality rule (proviso: x not free in f, g, or hypotheses) + f(x) == g(x) + ------------ + f == g *) +fun extensional (th as Thm{sign,maxidx,hyps,prop}) = + case prop of + (Const("==",_)) $ (f$x) $ (g$y) => + let fun err(msg) = raise THM("extensional: "^msg, 0, [th]) + in (if x<>y then err"different variables" else + case y of + Free _ => + if exists (apl(y, Logic.occs)) (f::g::hyps) + then err"variable free in hyps or functions" else () + | Var _ => + if Logic.occs(y,f) orelse Logic.occs(y,g) + then err"variable free in functions" else () + | _ => err"not a variable"); + Thm{sign=sign, hyps=hyps, maxidx=maxidx, + prop= Logic.mk_equals(f,g)} + end + | _ => raise THM("extensional: premise", 0, [th]); + +(*The abstraction rule. The Free or Var x must not be free in the hypotheses. + The bound variable will be named "a" (since x will be something like x320) + t == u + ---------------- + %(x)t == %(x)u *) +fun abstract_rule a cx (th as Thm{sign,maxidx,hyps,prop}) = + let val x = Sign.term_of cx; + val (t,u) = Logic.dest_equals prop + handle TERM _ => + raise THM("abstract_rule: premise not an equality", 0, [th]) + fun result T = + Thm{sign= sign, maxidx= maxidx, hyps= hyps, + prop= Logic.mk_equals(Abs(a, T, abstract_over (x,t)), + Abs(a, T, abstract_over (x,u)))} + in case x of + Free(_,T) => + if exists (apl(x, Logic.occs)) hyps + then raise THM("abstract_rule: variable free in assumptions", 0, [th]) + else result T + | Var(_,T) => result T + | _ => raise THM("abstract_rule: not a variable", 0, [th]) + end; + +(*The combination rule + f==g t==u + ------------ + f(t)==g(u) *) +fun combination th1 th2 = + let val Thm{maxidx=max1, hyps=hyps1, prop=prop1,...} = th1 + and Thm{maxidx=max2, hyps=hyps2, prop=prop2,...} = th2 + in case (prop1,prop2) of + (Const("==",_) $ f $ g, Const("==",_) $ t $ u) => + Thm{sign= merge_theories(th1,th2), hyps= hyps1 union hyps2, + maxidx= max[max1,max2], prop= Logic.mk_equals(f$t, g$u)} + | _ => raise THM("combination: premises", 0, [th1,th2]) + end; + + +(*The equal propositions rule + A==B A + --------- + B *) +fun equal_elim th1 th2 = + let val Thm{maxidx=max1, hyps=hyps1, prop=prop1,...} = th1 + and Thm{maxidx=max2, hyps=hyps2, prop=prop2,...} = th2; + fun err(msg) = raise THM("equal_elim: "^msg, 0, [th1,th2]) + in case prop1 of + Const("==",_) $ A $ B => + if not (prop2 aconv A) then err"not equal" else + Thm{sign= merge_theories(th1,th2), hyps= hyps1 union hyps2, + maxidx= max[max1,max2], prop= B} + | _ => err"major premise" + end; + + +(* Equality introduction + A==>B B==>A + ------------- + A==B *) +fun equal_intr th1 th2 = +let val Thm{maxidx=max1, hyps=hyps1, prop=prop1,...} = th1 + and Thm{maxidx=max2, hyps=hyps2, prop=prop2,...} = th2; + fun err(msg) = raise THM("equal_intr: "^msg, 0, [th1,th2]) +in case (prop1,prop2) of + (Const("==>",_) $ A $ B, Const("==>",_) $ B' $ A') => + if A aconv A' andalso B aconv B' + then Thm{sign= merge_theories(th1,th2), hyps= hyps1 union hyps2, + maxidx= max[max1,max2], prop= Logic.mk_equals(A,B)} + else err"not equal" + | _ => err"premises" +end; + +(**** Derived rules ****) + +(*Discharge all hypotheses (need not verify cterms) + Repeated hypotheses are discharged only once; fold cannot do this*) +fun implies_intr_hyps (Thm{sign, maxidx, hyps=A::As, prop}) = + implies_intr_hyps + (Thm{sign=sign, maxidx=maxidx, + hyps= disch(As,A), prop= implies$A$prop}) + | implies_intr_hyps th = th; + +(*Smash" unifies the list of term pairs leaving no flex-flex pairs. + Instantiates the theorem and deletes trivial tpairs. + Resulting sequence may contain multiple elements if the tpairs are + not all flex-flex. *) +fun flexflex_rule (Thm{sign,maxidx,hyps,prop}) = + let fun newthm env = + let val (tpairs,horn) = + Logic.strip_flexpairs (Envir.norm_term env prop) + (*Remove trivial tpairs, of the form t=t*) + val distpairs = filter (not o op aconv) tpairs + val newprop = Logic.list_flexpairs(distpairs, horn) + in Thm{sign= sign, hyps= hyps, + maxidx= maxidx_of_term newprop, prop= newprop} + end; + val (tpairs,_) = Logic.strip_flexpairs prop + in Sequence.maps newthm + (Unify.smash_unifiers(sign, Envir.empty maxidx, tpairs)) + end; + + +(*Instantiation of Vars + A + -------------------- + A[t1/v1,....,tn/vn] *) + +(*Check that all the terms are Vars and are distinct*) +fun instl_ok ts = forall is_Var ts andalso null(findrep ts); + +(*For instantiate: process pair of cterms, merge theories*) +fun add_ctpair ((ct,cu), (sign,tpairs)) = + let val {sign=signt, t=t, T= T, ...} = Sign.rep_cterm ct + and {sign=signu, t=u, T= U, ...} = Sign.rep_cterm cu + in if T=U then (Sign.merge(sign, Sign.merge(signt, signu)), (t,u)::tpairs) + else raise TYPE("add_ctpair", [T,U], [t,u]) + end; + +fun add_ctyp ((v,ctyp), (sign',vTs)) = + let val {T,sign} = Sign.rep_ctyp ctyp + in (Sign.merge(sign,sign'), (v,T)::vTs) end; + +fun duplicates t = findrep (map (#1 o dest_Var) (term_vars t)); + +(*Left-to-right replacements: ctpairs = [...,(vi,ti),...]. + Instantiates distinct Vars by terms of same type. + Normalizes the new theorem! *) +fun instantiate (vcTs,ctpairs) (th as Thm{sign,maxidx,hyps,prop}) = + let val (newsign,tpairs) = foldr add_ctpair (ctpairs, (sign,[])); + val (newsign,vTs) = foldr add_ctyp (vcTs, (newsign,[])); + val prop = Type.inst_term_tvars(#tsig(Sign.rep_sg newsign),vTs) prop; + val newprop = Envir.norm_term (Envir.empty 0) (subst_atomic tpairs prop) + val newth = Thm{sign= newsign, hyps= hyps, + maxidx= maxidx_of_term newprop, prop= newprop} + in if not(instl_ok(map #1 tpairs)) orelse not(null(findrep(map #1 vTs))) + then raise THM("instantiate: not distinct Vars", 0, [th]) + else case duplicates newprop of + [] => newth + | ix::_ => raise THM("instantiate: conflicting types for " ^ + Syntax.string_of_vname ix ^ "\n", 0, [newth]) + end + handle TERM _ => + raise THM("instantiate: incompatible signatures",0,[th]) + | TYPE _ => raise THM("instantiate: types", 0, [th]); + + +(*The trivial implication A==>A, justified by assume and forall rules. + A can contain Vars, not so for assume! *) +fun trivial ct : thm = + let val {sign, t=A, T, maxidx} = Sign.rep_cterm ct + in if T<>propT then + raise THM("trivial: the term must have type prop", 0, []) + else Thm{sign= sign, maxidx= maxidx, hyps= [], prop= implies$A$A} + end; + +(* Replace all TFrees not in the hyps by new TVars *) +fun varifyT(Thm{sign,maxidx,hyps,prop}) = + let val tfrees = foldr add_term_tfree_names (hyps,[]) + in Thm{sign=sign, maxidx=max[0,maxidx], hyps=hyps, + prop= Type.varify(prop,tfrees)} + end; + +(* Replace all TVars by new TFrees *) +fun freezeT(Thm{sign,maxidx,hyps,prop}) = + let val prop' = Type.freeze (K true) prop + in Thm{sign=sign, maxidx=maxidx_of_term prop', hyps=hyps, prop=prop'} end; + + +(*** Inference rules for tactics ***) + +(*Destruct proof state into constraints, other goals, goal(i), rest *) +fun dest_state (state as Thm{prop,...}, i) = + let val (tpairs,horn) = Logic.strip_flexpairs prop + in case Logic.strip_prems(i, [], horn) of + (B::rBs, C) => (tpairs, rev rBs, B, C) + | _ => raise THM("dest_state", i, [state]) + end + handle TERM _ => raise THM("dest_state", i, [state]); + +(*Increment variables and parameters of rule as required for + resolution with goal i of state. *) +fun lift_rule (state, i) orule = + let val Thm{prop=sprop,maxidx=smax,...} = state; + val (Bi::_, _) = Logic.strip_prems(i, [], Logic.skip_flexpairs sprop) + handle TERM _ => raise THM("lift_rule", i, [orule,state]); + val (lift_abs,lift_all) = Logic.lift_fns(Bi,smax+1); + val (Thm{sign,maxidx,hyps,prop}) = orule + val (tpairs,As,B) = Logic.strip_horn prop + in Thm{hyps=hyps, sign= merge_theories(state,orule), + maxidx= maxidx+smax+1, + prop= Logic.rule_of(map (pairself lift_abs) tpairs, + map lift_all As, lift_all B)} + end; + +(*Solve subgoal Bi of proof state B1...Bn/C by assumption. *) +fun assumption i state = + let val Thm{sign,maxidx,hyps,prop} = state; + val (tpairs, Bs, Bi, C) = dest_state(state,i) + fun newth (env as Envir.Envir{maxidx,asol,iTs}, tpairs) = + Thm{sign=sign, hyps=hyps, maxidx=maxidx, prop= + case (Envir.alist_of_olist asol, iTs) of + (*avoid wasted normalizations*) + ([],[]) => Logic.rule_of(tpairs, Bs, C) + | _ => (*normalize the new rule fully*) + Envir.norm_term env (Logic.rule_of(tpairs, Bs, C))}; + fun addprfs [] = Sequence.null + | addprfs ((t,u)::apairs) = Sequence.seqof (fn()=> Sequence.pull + (Sequence.mapp newth + (Unify.unifiers(sign,Envir.empty maxidx, (t,u)::tpairs)) + (addprfs apairs))) + in addprfs (Logic.assum_pairs Bi) end; + +(*Solve subgoal Bi of proof state B1...Bn/C by assumption. + Checks if Bi's conclusion is alpha-convertible to one of its assumptions*) +fun eq_assumption i state = + let val Thm{sign,maxidx,hyps,prop} = state; + val (tpairs, Bs, Bi, C) = dest_state(state,i) + in if exists (op aconv) (Logic.assum_pairs Bi) + then Thm{sign=sign, hyps=hyps, maxidx=maxidx, + prop=Logic.rule_of(tpairs, Bs, C)} + else raise THM("eq_assumption", 0, [state]) + end; + + +(** User renaming of parameters in a subgoal **) + +(*Calls error rather than raising an exception because it is intended + for top-level use -- exception handling would not make sense here. + The names in cs, if distinct, are used for the innermost parameters; + preceding parameters may be renamed to make all params distinct.*) +fun rename_params_rule (cs, i) state = + let val Thm{sign,maxidx,hyps,prop} = state + val (tpairs, Bs, Bi, C) = dest_state(state,i) + val iparams = map #1 (Logic.strip_params Bi) + val short = length iparams - length cs + val newnames = + if short<0 then error"More names than abstractions!" + else variantlist(take (short,iparams), cs) @ cs + val freenames = map (#1 o dest_Free) (term_frees prop) + val newBi = Logic.list_rename_params (newnames, Bi) + in + case findrep cs of + c::_ => error ("Bound variables not distinct: " ^ c) + | [] => (case cs inter freenames of + a::_ => error ("Bound/Free variable clash: " ^ a) + | [] => Thm{sign=sign, hyps=hyps, maxidx=maxidx, prop= + Logic.rule_of(tpairs, Bs@[newBi], C)}) + end; + +(*** Preservation of bound variable names ***) + +(*Scan a pair of terms; while they are similar, + accumulate corresponding bound vars in "al"*) +fun match_bvs(Abs(x,_,s),Abs(y,_,t), al) = match_bvs(s,t,(x,y)::al) + | match_bvs(f$s, g$t, al) = match_bvs(f,g,match_bvs(s,t,al)) + | match_bvs(_,_,al) = al; + +(* strip abstractions created by parameters *) +fun match_bvars((s,t),al) = match_bvs(strip_abs_body s, strip_abs_body t, al); + + +(* strip_apply f A(,B) strips off all assumptions/parameters from A + introduced by lifting over B, and applies f to remaining part of A*) +fun strip_apply f = + let fun strip(Const("==>",_)$ A1 $ B1, + Const("==>",_)$ _ $ B2) = implies $ A1 $ strip(B1,B2) + | strip((c as Const("all",_)) $ Abs(a,T,t1), + Const("all",_) $ Abs(_,_,t2)) = c$Abs(a,T,strip(t1,t2)) + | strip(A,_) = f A + in strip end; + +(*Use the alist to rename all bound variables and some unknowns in a term + dpairs = current disagreement pairs; tpairs = permanent ones (flexflex); + Preserves unknowns in tpairs and on lhs of dpairs. *) +fun rename_bvs([],_,_,_) = I + | rename_bvs(al,dpairs,tpairs,B) = + let val vars = foldr add_term_vars + (map fst dpairs @ map fst tpairs @ map snd tpairs, []) + (*unknowns appearing elsewhere be preserved!*) + val vids = map (#1 o #1 o dest_Var) vars; + fun rename(t as Var((x,i),T)) = + (case assoc(al,x) of + Some(y) => if x mem vids orelse y mem vids then t + else Var((y,i),T) + | None=> t) + | rename(Abs(x,T,t)) = + Abs(case assoc(al,x) of Some(y) => y | None => x, + T, rename t) + | rename(f$t) = rename f $ rename t + | rename(t) = t; + fun strip_ren Ai = strip_apply rename (Ai,B) + in strip_ren end; + +(*Function to rename bounds/unknowns in the argument, lifted over B*) +fun rename_bvars(dpairs, tpairs, B) = + rename_bvs(foldr match_bvars (dpairs,[]), dpairs, tpairs, B); + + +(*** RESOLUTION ***) + +(*strip off pairs of assumptions/parameters in parallel -- they are + identical because of lifting*) +fun strip_assums2 (Const("==>", _) $ _ $ B1, + Const("==>", _) $ _ $ B2) = strip_assums2 (B1,B2) + | strip_assums2 (Const("all",_)$Abs(a,T,t1), + Const("all",_)$Abs(_,_,t2)) = + let val (B1,B2) = strip_assums2 (t1,t2) + in (Abs(a,T,B1), Abs(a,T,B2)) end + | strip_assums2 BB = BB; + + +(*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C) + Unifies B with Bi, replacing subgoal i (1 <= i <= n) + If match then forbid instantiations in proof state + If lifted then shorten the dpair using strip_assums2. + If eres_flg then simultaneously proves A1 by assumption. + nsubgoal is the number of new subgoals (written m above). + Curried so that resolution calls dest_state only once. +*) +local open Sequence; exception Bicompose +in +fun bicompose_aux match (state, (stpairs, Bs, Bi, C), lifted) + (eres_flg, orule, nsubgoal) = + let val Thm{maxidx=smax, hyps=shyps, ...} = state + and Thm{maxidx=rmax, hyps=rhyps, prop=rprop,...} = orule; + val sign = merge_theories(state,orule); + (** Add new theorem with prop = '[| Bs; As |] ==> C' to thq **) + fun addth As ((env as Envir.Envir{maxidx,asol,iTs}, tpairs), thq) = + let val minenv = case Envir.alist_of_olist asol of + [] => ~1 | ((_,i),_) :: _ => i; + val minx = min (minenv :: map (fn ((_,i),_) => i) iTs); + val normt = Envir.norm_term env; + (*Perform minimal copying here by examining env*) + val normp = if minx = ~1 then (tpairs, Bs@As, C) + else + let val ntps = map (pairself normt) tpairs + in if minx>smax then (*no assignments in state*) + (ntps, Bs @ map normt As, C) + else if match then raise Bicompose + else (*normalize the new rule fully*) + (ntps, map normt (Bs @ As), normt C) + end + val th = Thm{sign=sign, hyps=rhyps union shyps, maxidx=maxidx, + prop= Logic.rule_of normp} + in cons(th, thq) end handle Bicompose => thq + val (rtpairs,rhorn) = Logic.strip_flexpairs(rprop); + val (rAs,B) = Logic.strip_prems(nsubgoal, [], rhorn) + handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]); + (*Modify assumptions, deleting n-th if n>0 for e-resolution*) + fun newAs(As0, n, dpairs, tpairs) = + let val As1 = if !Logic.auto_rename orelse not lifted then As0 + else map (rename_bvars(dpairs,tpairs,B)) As0 + in (map (Logic.flatten_params n) As1) + handle TERM _ => + raise THM("bicompose: 1st premise", 0, [orule]) + end; + val env = Envir.empty(max[rmax,smax]); + val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi); + val dpairs = BBi :: (rtpairs@stpairs); + (*elim-resolution: try each assumption in turn. Initially n=1*) + fun tryasms (_, _, []) = null + | tryasms (As, n, (t,u)::apairs) = + (case pull(Unify.unifiers(sign, env, (t,u)::dpairs)) of + None => tryasms (As, n+1, apairs) + | cell as Some((_,tpairs),_) => + its_right (addth (newAs(As, n, [BBi,(u,t)], tpairs))) + (seqof (fn()=> cell), + seqof (fn()=> pull (tryasms (As, n+1, apairs))))); + fun eres [] = raise THM("bicompose: no premises", 0, [orule,state]) + | eres (A1::As) = tryasms (As, 1, Logic.assum_pairs A1); + (*ordinary resolution*) + fun res(None) = null + | res(cell as Some((_,tpairs),_)) = + its_right (addth(newAs(rev rAs, 0, [BBi], tpairs))) + (seqof (fn()=> cell), null) + in if eres_flg then eres(rev rAs) + else res(pull(Unify.unifiers(sign, env, dpairs))) + end; +end; (*open Sequence*) + + +fun bicompose match arg i state = + bicompose_aux match (state, dest_state(state,i), false) arg; + +(*Quick test whether rule is resolvable with the subgoal with hyps Hs + and conclusion B. If eres_flg then checks 1st premise of rule also*) +fun could_bires (Hs, B, eres_flg, rule) = + let fun could_reshyp (A1::_) = exists (apl(A1,could_unify)) Hs + | could_reshyp [] = false; (*no premise -- illegal*) + in could_unify(concl_of rule, B) andalso + (not eres_flg orelse could_reshyp (prems_of rule)) + end; + +(*Bi-resolution of a state with a list of (flag,rule) pairs. + Puts the rule above: rule/state. Renames vars in the rules. *) +fun biresolution match brules i state = + let val lift = lift_rule(state, i); + val (stpairs, Bs, Bi, C) = dest_state(state,i) + val B = Logic.strip_assums_concl Bi; + val Hs = Logic.strip_assums_hyp Bi; + val comp = bicompose_aux match (state, (stpairs, Bs, Bi, C), true); + fun res [] = Sequence.null + | res ((eres_flg, rule)::brules) = + if could_bires (Hs, B, eres_flg, rule) + then Sequence.seqof (*delay processing remainder til needed*) + (fn()=> Some(comp (eres_flg, lift rule, nprems_of rule), + res brules)) + else res brules + in Sequence.flats (res brules) end; + + +(**** THEORIES ****) + +val pure_thy = Pure{sign = Sign.pure}; + +(*Look up the named axiom in the theory*) +fun get_axiom thy axname = + let fun get (Pure _) = raise Match + | get (Extend{axioms,thy,...}) = + (case Symtab.lookup(axioms,axname) of + Some th => th + | None => get thy) + | get (Merge{thy1,thy2,...}) = + get thy1 handle Match => get thy2 + in get thy + handle Match => raise THEORY("get_axiom: No axiom "^axname, [thy]) + end; + +(*Converts Frees to Vars: axioms can be written without question marks*) +fun prepare_axiom sign sP = + Logic.varify (Sign.term_of (Sign.read_cterm sign (sP,propT))); + +(*Read an axiom for a new theory*) +fun read_ax sign (a, sP) : string*thm = + let val prop = prepare_axiom sign sP + in (a, Thm{sign=sign, hyps=[], maxidx= maxidx_of_term prop, prop= prop}) + end + handle ERROR => + error("extend_theory: The error above occurred in axiom " ^ a); + +fun mk_axioms sign axpairs = + Symtab.st_of_alist(map (read_ax sign) axpairs, Symtab.null) + handle Symtab.DUPLICATE(a) => error("Two axioms named " ^ a); + +(*Extension of a theory with given classes, types, constants and syntax. + Reads the axioms from strings: axpairs have the form (axname, axiom). *) +fun extend_theory thy thyname ext axpairs = + let val sign = Sign.extend (sign_of thy) thyname ext; + val axioms= mk_axioms sign axpairs + in Extend{sign=sign, axioms= axioms, thy = thy} end; + +(*The union of two theories*) +fun merge_theories (thy1,thy2) = + Merge{sign = Sign.merge(sign_of thy1, sign_of thy2), + thy1 = thy1, thy2 = thy2}; + + +(*** Meta simp sets ***) + +type rrule = {thm:thm, lhs:term}; +datatype meta_simpset = + Mss of {net:rrule Net.net, congs:(string * rrule)list, primes:string, + prems: thm list, mk_rews: thm -> thm list}; + +(*A "mss" contains data needed during conversion: + net: discrimination net of rewrite rules + congs: association list of congruence rules + primes: offset for generating unique new names + for rewriting under lambda abstractions + mk_rews: used when local assumptions are added +*) + +val empty_mss = Mss{net= Net.empty, congs= [], primes="", prems= [], + mk_rews = K[]}; + +exception SIMPLIFIER of string * thm; + +fun prtm a sg t = (writeln a; writeln(Sign.string_of_term sg t)); + +(*simple test for looping rewrite*) +fun loops sign prems (lhs,rhs) = + null(prems) andalso + Pattern.eta_matches (#tsig(Sign.rep_sg sign)) (lhs,rhs); + +fun mk_rrule (thm as Thm{hyps,sign,prop,maxidx,...}) = + let val prems = Logic.strip_imp_prems prop + val concl = Pattern.eta_contract (Logic.strip_imp_concl prop) + val (lhs,rhs) = Logic.dest_equals concl handle TERM _ => + raise SIMPLIFIER("Rewrite rule not a meta-equality",thm) + in if loops sign prems (lhs,rhs) + then (prtm "Warning: ignoring looping rewrite rule" sign prop; None) + else Some{thm=thm,lhs=lhs} + end; + +fun add_simp(mss as Mss{net,congs,primes,prems,mk_rews}, + thm as Thm{sign,prop,...}) = + let fun eq({thm=Thm{prop=p1,...},...}:rrule, + {thm=Thm{prop=p2,...},...}:rrule) = p1 aconv p2 + in case mk_rrule thm of + None => mss + | Some(rrule as {lhs,...}) => + Mss{net= (Net.insert_term((lhs,rrule),net,eq) + handle Net.INSERT => + (prtm "Warning: ignoring duplicate rewrite rule" sign prop; + net)), + congs=congs, primes=primes, prems=prems,mk_rews=mk_rews} + end; + +val add_simps = foldl add_simp; + +fun mss_of thms = add_simps(empty_mss,thms); + +fun add_cong(Mss{net,congs,primes,prems,mk_rews},thm) = + let val (lhs,_) = Logic.dest_equals(concl_of thm) handle TERM _ => + raise SIMPLIFIER("Congruence not a meta-equality",thm) + val lhs = Pattern.eta_contract lhs + val (a,_) = dest_Const (head_of lhs) handle TERM _ => + raise SIMPLIFIER("Congruence must start with a constant",thm) + in Mss{net=net, congs=(a,{lhs=lhs,thm=thm})::congs, primes=primes, + prems=prems, mk_rews=mk_rews} + end; + +val (op add_congs) = foldl add_cong; + +fun add_prems(Mss{net,congs,primes,prems,mk_rews},thms) = + Mss{net=net, congs=congs, primes=primes, prems=thms@prems, mk_rews=mk_rews}; + +fun prems_of_mss(Mss{prems,...}) = prems; + +fun set_mk_rews(Mss{net,congs,primes,prems,...},mk_rews) = + Mss{net=net, congs=congs, primes=primes, prems=prems, mk_rews=mk_rews}; +fun mk_rews_of_mss(Mss{mk_rews,...}) = mk_rews; + + +(*** Meta-level rewriting + uses conversions, omitting proofs for efficiency. See + L C Paulson, A higher-order implementation of rewriting, + Science of Computer Programming 3 (1983), pages 119-149. ***) + +type prover = meta_simpset -> thm -> thm option; +type termrec = (Sign.sg * term list) * term; +type conv = meta_simpset -> termrec -> termrec; + +val trace_simp = ref false; + +fun trace_term a sg t = if !trace_simp then prtm a sg t else (); + +fun trace_thm a (Thm{sign,prop,...}) = trace_term a sign prop; + +fun check_conv(thm as Thm{sign,hyps,prop,...}, prop0) = + let fun err() = (trace_term "Proved wrong thm" sign prop; + error "Check your prover") + val (lhs0,_) = Logic.dest_equals(Logic.strip_imp_concl prop0) + in case prop of + Const("==",_) $ lhs $ rhs => + if (lhs = lhs0) orelse + (lhs aconv (Envir.norm_term (Envir.empty 0) lhs0)) + then (trace_thm "SUCCEEDED" thm; ((sign,hyps),rhs)) + else err() + | _ => err() + end; + +(*Conversion to apply the meta simpset to a term*) +fun rewritec prover (mss as Mss{net,...}) (sghyt as (sgt,hypst),t) = + let val t = Pattern.eta_contract t + fun rew {thm as Thm{sign,hyps,maxidx,prop,...}, lhs} = + let val sign' = Sign.merge(sgt,sign) + val tsig = #tsig(Sign.rep_sg sign') + val insts = Pattern.match tsig (lhs,t) + val prop' = subst_vars insts prop; + val hyps' = hyps union hypst; + val thm' = Thm{sign=sign', hyps=hyps', prop=prop', maxidx=maxidx} + in if nprems_of thm' = 0 + then let val (_,rhs) = Logic.dest_equals prop' + in trace_thm "Rewriting:" thm'; Some((sign',hyps'),rhs) end + else (trace_thm "Trying to rewrite:" thm'; + case prover mss thm' of + None => (trace_thm "FAILED" thm'; None) + | Some(thm2) => Some(check_conv(thm2,prop'))) + end + + fun rewl [] = None + | rewl (rrule::rrules) = + let val opt = rew rrule handle Pattern.MATCH => None + in case opt of None => rewl rrules | some => some end; + + in case t of + Abs(_,_,body) $ u => Some(sghyt,subst_bounds([u], body)) + | _ => rewl (Net.match_term net t) + end; + +(*Conversion to apply a congruence rule to a term*) +fun congc prover {thm=cong,lhs=lhs} (sghyt as (sgt,hypst),t) = + let val Thm{sign,hyps,maxidx,prop,...} = cong + val sign' = Sign.merge(sgt,sign) + val tsig = #tsig(Sign.rep_sg sign') + val insts = Pattern.match tsig (lhs,t) handle Pattern.MATCH => + error("Congruence rule did not match") + val prop' = subst_vars insts prop; + val thm' = Thm{sign=sign', hyps=hyps union hypst, + prop=prop', maxidx=maxidx} + val unit = trace_thm "Applying congruence rule" thm'; + + in case prover thm' of + None => error("Failed congruence proof!") + | Some(thm2) => check_conv(thm2,prop') + end; + + +fun bottomc prover = + let fun botc mss trec = let val trec1 = subc mss trec + in case rewritec prover mss trec1 of + None => trec1 + | Some(trec2) => botc mss trec2 + end + + and subc (mss as Mss{net,congs,primes,prems,mk_rews}) + (trec as (sghy,t)) = + (case t of + Abs(a,T,t) => + let val v = Free(".subc." ^ primes,T) + val mss' = Mss{net=net, congs=congs, primes=primes^"'", + prems=prems,mk_rews=mk_rews} + val (sghy',t') = botc mss' (sghy,subst_bounds([v],t)) + in (sghy', Abs(a, T, abstract_over(v,t'))) end + | t$u => (case t of + Const("==>",_)$s => impc(sghy,s,u,mss) + | Abs(_,_,body) => subc mss (sghy,subst_bounds([u], body)) + | _ => + let fun appc() = let val (sghy1,t1) = botc mss (sghy,t) + val (sghy2,u1) = botc mss (sghy1,u) + in (sghy2,t1$u1) end + val (h,ts) = strip_comb t + in case h of + Const(a,_) => + (case assoc(congs,a) of + None => appc() + | Some(cong) => congc (prover mss) cong trec) + | _ => appc() + end) + | _ => trec) + + and impc(sghy,s,u,mss as Mss{mk_rews,...}) = + let val (sghy1 as (sg1,hyps1),s') = botc mss (sghy,s) + val (rthms,mss) = + if maxidx_of_term s' <> ~1 then ([],mss) + else let val thm = Thm{sign=sg1,hyps=[s'],prop=s',maxidx= ~1} + in (mk_rews thm, add_prems(mss,[thm])) end + val unit = seq (trace_thm "Adding rewrite rule:") rthms + val mss' = add_simps(mss,rthms) + val ((sg2,hyps2),u') = botc mss' (sghy1,u) + in ((sg2,hyps2\s'), Logic.mk_implies(s',u')) end + + in botc end; + + +(*** Meta-rewriting: rewrites t to u and returns the theorem t==u ***) +(* Parameters: + mss: contains equality theorems of the form [|p1,...|] ==> t==u + prover: how to solve premises in conditional rewrites and congruences +*) + +(*** FIXME: check that #primes(mss) does not "occur" in ct alread ***) +fun rewrite_cterm mss prover ct = + let val {sign, t, T, maxidx} = Sign.rep_cterm ct; + val ((sign',hyps),u) = bottomc prover mss ((sign,[]),t); + val prop = Logic.mk_equals(t,u) + in Thm{sign= sign', hyps= hyps, maxidx= maxidx_of_term prop, prop= prop} + end + +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/type.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/type.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,783 @@ +(* Title: Types and Sorts + ID: $Id$ + Author: Tobias Nipkow & Lawrence C Paulson + +Maybe type classes should go in a separate module? +*) + + +signature TYPE = +sig + structure Symtab:SYMTAB + type type_sig + val defaultS: type_sig -> sort + val extend: type_sig * (class * class list)list * sort * + (string list * int)list * + (string list * (sort list * class))list -> type_sig + val freeze: (indexname -> bool) -> term -> term + val freeze_vars: typ -> typ + val infer_types: type_sig * typ Symtab.table * (indexname -> typ option) * + (indexname -> sort option) * (typ -> string) * typ * term + -> term * (indexname*typ)list + val inst_term_tvars: type_sig * (indexname * typ)list -> term -> term + val logical_type: type_sig -> string -> bool + val merge: type_sig * type_sig -> type_sig + val thaw_vars: typ -> typ + val tsig0: type_sig + val type_errors: type_sig * (typ->string) -> typ * string list -> string list + val typ_instance: type_sig * typ * typ -> bool + val typ_match: type_sig -> (indexname*typ)list * (typ*typ) -> + (indexname*typ)list + val unify: type_sig -> (typ*typ) * (indexname*typ)list -> (indexname*typ)list + val varifyT: typ -> typ + val varify: term * string list -> term + exception TUNIFY + exception TYPE_MATCH; +end; + +functor TypeFun(structure Symtab:SYMTAB and Syntax:SYNTAX) : TYPE = +struct +structure Symtab = Symtab + +(* Miscellany *) + +val commas = space_implode ","; +fun str_of_sort S = "{" ^ commas S ^ "}"; +fun str_of_dom dom = "(" ^ commas (map str_of_sort dom) ^ ")"; +fun str_of_decl(t,w,C) = t ^ ": " ^ str_of_dom w ^ C; + + +(* Association list Manipulation *) + + +(* two-fold Association list lookup *) + +fun assoc2 (aal,(key1,key2)) = case assoc (aal,key1) of + Some (al) => assoc (al,key2) + | None => None; + + + +(**** TYPE CLASSES ****) + +type domain = sort list; +type arity = domain * class; + +datatype type_sig = + TySg of {classes: class list, + default: sort, + subclass: (class * class list) list, + args: (string * int) list, + coreg: (string * (class * domain) list) list }; + +(* classes: a list of all declared classes; + default: the default sort attached to all unconstrained TVars + occurring in a term to be type-inferred; + subclass: association list representation of subclass relationship; + (c,cs) is interpreted as "c is a proper subclass of all + elemenst of cs". Note that c itself is not a memeber of cs. + args: an association list of all declared types with the number of their + arguments + coreg: a two-fold association list of all type arities; (t,al) means that + type constructor t has the arities in al; an element (c,ss) of al + represents the arity (ss)c +*) + + +val tsig0 = TySg{classes = [], + default = [], + subclass = [], + args = [], + coreg = []}; + +fun undcl_class (s) = error("Class " ^ s ^ " has not been declared"); + +fun undcl_type(c) = "Undeclared type: " ^ c; +fun undcl_type_err(c) = error(undcl_type(c)); + + +(* 'leq' checks the partial order on classes according to the + statements in the association list 'a' (i.e.'subclass') +*) + +fun less a (C,D) = case assoc (a,C) of + Some(ss) => D mem ss + | None => undcl_class (C) ; + +fun leq a (C,D) = C = D orelse less a (C,D); + + +fun defaultS(TySg{default,...}) = default; + +(* 'logical_type' checks if some type declaration t has as range + a class which is a subclass of "logic" *) + +fun logical_type(tsig as TySg{subclass,coreg,...}) t = + let fun is_log C = leq subclass (C,"logic") + in case assoc (coreg,t) of + Some(ars) => exists (is_log o #1) ars + | None => undcl_type_err(t) + end; + + +(* 'sortorder' checks the ordering on sets of classes,i.e. on sorts: + S1 <= S2 ,iff for every class C2 in S2 there exists a class C1 in S1 + with C1 <= C2 (according to an association list 'a') +*) + +fun sortorder a (S1,S2) = + forall (fn C2 => exists (fn C1 => leq a (C1,C2)) S1) S2; + + +(* 'inj' inserts a new class C into a given class set S (i.e.sort) only if + there exists no class in S which is <= C; + the resulting set is minimal if S was minimal +*) + +fun inj a (C,S) = + let fun inj1 [] = [C] + | inj1 (D::T) = if leq a (D,C) then D::T + else if leq a (C,D) then inj1 T + else D::(inj1 T) + in inj1 S end; + + +(* 'union_sort' forms the minimal union set of two sorts S1 and S2 + under the assumption that S2 is minimal *) + +fun union_sort a = foldr (inj a); + + +(* 'elementwise_union' forms elementwise the minimal union set of two + sort lists under the assumption that the two lists have the same length +*) + +fun elementwise_union a (Ss1,Ss2) = map (union_sort a) (Ss1~~Ss2); + + +(* 'lew' checks for two sort lists the ordering for all corresponding list + elements (i.e. sorts) *) + +fun lew a (w1,w2) = forall (sortorder a) (w1~~w2); + + +(* 'is_min' checks if a class C is minimal in a given sort S under the + assumption that S contains C *) + +fun is_min a S C = not (exists (fn (D) => less a (D,C)) S); + + +(* 'min_sort' reduces a sort to its minimal classes *) + +fun min_sort a S = distinct(filter (is_min a S) S); + + +(* 'min_domain' minimizes the domain sorts of type declarationsl; + the function will be applied on the type declarations in extensions *) + +fun min_domain subclass = + let fun one_min (f,(doms,ran)) = (f, (map (min_sort subclass) doms, ran)) + in map one_min end; + + +(* 'min_filter' filters a list 'ars' consisting of arities (domain * class) + and gives back a list of those range classes whose domains meet the + predicate 'pred' *) + +fun min_filter a pred ars = + let fun filt ([],l) = l + | filt ((c,x)::xs,l) = if pred(x) then filt (xs,inj a (c,l)) + else filt (xs,l) + in filt (ars,[]) end; + + +(* 'cod_above' filters all arities whose domains are elementwise >= than + a given domain 'w' and gives back a list of the corresponding range + classes *) + +fun cod_above (a,w,ars) = min_filter a (fn w' => lew a (w,w')) ars; + + +(* 'least_sort' returns for a given type its maximum sort: + - type variables, free types: the sort brought with + - type constructors: recursive determination of the maximum sort of the + arguments if the type is declared in 'coreg' of the + given type signature *) + +fun least_sort (tsig as TySg{subclass,coreg,...}) = + let fun ls(T as Type(a,Ts)) = + let val ars = case assoc (coreg,a) of Some(ars) => ars + | None => raise TYPE(undcl_type a,[T],[]); + in cod_above(subclass,map ls Ts,ars) end + | ls(TFree(a,S)) = S + | ls(TVar(a,S)) = S + in ls end; + + +fun check_has_sort(tsig as TySg{subclass,coreg,...},T,S) = + if sortorder subclass ((least_sort tsig T),S) then () + else raise TYPE("Type not of sort " ^ (str_of_sort S),[T],[]) + + +(*Instantiation of type variables in types *) +fun inst_typ_tvars(tsig,tye) = + let fun inst(Type(a,Ts)) = Type(a, map inst Ts) + | inst(T as TFree _) = T + | inst(T as TVar(v,S)) = (case assoc(tye,v) of + None => T | Some(U) => (check_has_sort(tsig,U,S); U)) + in inst end; + +(*Instantiation of type variables in terms *) +fun inst_term_tvars(tsig,tye) = map_term_types (inst_typ_tvars(tsig,tye)); + +exception TYPE_MATCH; + +(* Typ matching + typ_match(ts,s,(U,T)) = s' <=> s'(U)=T and s' is an extension of s *) +fun typ_match tsig = +let fun tm(subs, (TVar(v,S), T)) = (case assoc(subs,v) of + None => ( (v, (check_has_sort(tsig,T,S); T))::subs + handle TYPE _ => raise TYPE_MATCH ) + | Some(U) => if U=T then subs else raise TYPE_MATCH) + | tm(subs, (Type(a,Ts), Type(b,Us))) = + if a<>b then raise TYPE_MATCH + else foldl tm (subs, Ts~~Us) + | tm(subs, (TFree(x), TFree(y))) = + if x=y then subs else raise TYPE_MATCH + | tm _ = raise TYPE_MATCH +in tm end; + +fun typ_instance(tsig,T,U) = let val x = typ_match tsig ([],(U,T)) in true end + handle TYPE_MATCH => false; + + +(* EXTENDING AND MERGIN TYPE SIGNATURES *) + +fun not_ident(s) = error("Must be an identifier: " ^ s); + +fun twice(a) = error("Type constructor " ^a^ " has already been declared."); + +(*Is the type valid? Accumulates error messages in "errs".*) +fun type_errors (tsig as TySg{classes,subclass,args,...}, + string_of_typ) (T,errs) = +let fun class_err([],errs) = errs + | class_err(S::Ss,errs) = + if S mem classes then class_err (Ss,errs) + else class_err (Ss,("Class " ^ S ^ " has not been declared") :: errs) + fun errors(Type(c,Us), errs) = + let val errs' = foldr errors (Us,errs) + in case assoc(args,c) of + None => (undcl_type c) :: errs + | Some(n) => if n=length(Us) then errs' + else ("Wrong number of arguments: " ^ c) :: errs + end + | errors(TFree(_,S), errs) = class_err(S,errs) + | errors(TVar(_,S), errs) = class_err(S,errs); +in case errors(T,[]) of + [] => ((least_sort tsig T; errs) + handle TYPE(_,[U],_) => ("Ill-formed type: " ^ string_of_typ U) + :: errs) + | errs' => errs'@errs +end; + + +(* 'add_class' adds a new class to the list of all existing classes *) + +fun add_class (classes,(s,_)) = + if s mem classes then error("Class " ^ s ^ " declared twice.") + else s::classes; + +(* 'add_subclass' adds a tuple consisiting of a new class (the new class + has already been inserted into the 'classes' list) and its + superclasses (they must be declared in 'classes' too) to the 'subclass' + list of the given type signature; + furthermore all inherited superclasses according to the superclasses + brought with are inserted and there is a check that there are no + cycles (i.e. C <= D <= C, with C <> D); *) + +fun add_subclass classes (subclass,(s,ges)) = +let fun upd (subclass,s') = if s' mem classes then + let val Some(ges') = assoc (subclass,s) + in case assoc (subclass,s') of + Some(sups) => if s mem sups + then error(" Cycle :" ^ s^" <= "^ s'^" <= "^ s ) + else overwrite (subclass,(s,sups union ges')) + | None => subclass + end + else undcl_class(s') +in foldl upd (subclass@[(s,ges)],ges) end; + + +(* 'extend_classes' inserts all new classes into the corresponding + lists ('classes','subclass') if possible *) + +fun extend_classes (classes,subclass,newclasses) = + if newclasses = [] then (classes,subclass) else + let val classes' = foldl add_class (classes,newclasses); + val subclass' = foldl (add_subclass classes') (subclass,newclasses); + in (classes',subclass') end; + +(* Corregularity *) + +(* 'is_unique_decl' checks if there exists just one declaration t:(Ss)C *) + +fun is_unique_decl coreg (t,(s,w)) = case assoc2 (coreg,(t,s)) of + Some(w1) => if w = w1 then () else + error("There are two declarations\n" ^ + str_of_decl(t,w,s) ^ " and\n" ^ + str_of_decl(t,w1,s) ^ "\n" ^ + "with the same result class.") + | None => (); + +(* 'restr2' checks if there are two declarations t:(Ss1)C1 and t:(Ss2)C2 + such that C1 >= C2 then Ss1 >= Ss2 (elementwise) *) + +fun subs (classes,subclass) C = + let fun sub (rl,l) = if leq subclass (l,C) then l::rl else rl + in foldl sub ([],classes) end; + +fun coreg_err(t,(w1,C),(w2,D)) = + error("Declarations " ^ str_of_decl(t,w1,C) ^ " and " + ^ str_of_decl(t,w2,D) ^ " are in conflict"); + +fun restr2 classes (subclass,coreg) (t,(s,w)) = + let fun restr ([],test) = () + | restr (s1::Ss,test) = (case assoc2 (coreg,(t,s1)) of + Some (dom) => if lew subclass (test (w,dom)) then restr (Ss,test) + else coreg_err (t,(w,s),(dom,s1)) + | None => restr (Ss,test)) + fun forward (t,(s,w)) = + let val s_sups = case assoc (subclass,s) of + Some(s_sups) => s_sups | None => undcl_class(s); + in restr (s_sups,I) end + fun backward (t,(s,w)) = + let val s_subs = subs (classes,subclass) s + in restr (s_subs,fn (x,y) => (y,x)) end + in (backward (t,(s,w)); forward (t,(s,w))) end; + + +fun varying_decls(t) = + error("Type constructor "^t^" has varying number of arguments."); + + + +(* 'coregular' checks + - the two restriction conditions 'is_unique_decl' and 'restr2' + - if the classes in the new type declarations are known in the + given type signature + - if one type constructor has always the same number of arguments; + if one type declaration has passed all checks it is inserted into + the 'coreg' association list of the given type signatrure *) + +fun coregular (classes,subclass,args) = + let fun ex C = if C mem classes then () else undcl_class(C); + + fun addar(w,C) (coreg,t) = case assoc(args,t) of + Some(n) => if n <> length w then varying_decls(t) else + (is_unique_decl coreg (t,(C,w)); + (seq o seq) ex w; + restr2 classes (subclass,coreg) (t,(C,w)); + let val Some(ars) = assoc(coreg,t) + in overwrite(coreg,(t,(C,w) ins ars)) end) + | None => undcl_type_err(t); + + fun addts(coreg,(ts,ar)) = foldl (addar ar) (coreg,ts) + + in addts end; + + +(* 'close' extends the 'coreg' association list after all new type + declarations have been inserted successfully: + for every declaration t:(Ss)C , for all classses D with C <= D: + if there is no declaration t:(Ss')C' with C < C' and C' <= D + then insert the declaration t:(Ss)D into 'coreg' + this means, if there exists a declaration t:(Ss)C and there is + no declaration t:(Ss')D with C <=D then the declaration holds + for all range classes more general than C *) + +fun close (coreg,subclass) = + let fun check sl (l,(s,dom)) = case assoc (subclass,s) of + Some(sups) => + let fun close_sup (l,sup) = + if exists (fn s'' => less subclass (s,s'') andalso + leq subclass (s'',sup)) sl + then l + else (sup,dom)::l + in foldl close_sup (l,sups) end + | None => l; + fun ext (s,l) = (s, foldl (check (map #1 l)) (l,l)); + in map ext coreg end; + +fun add_types(ac,(ts,n)) = + let fun add_type((args,coreg),t) = case assoc(args,t) of + Some _ => twice(t) | None => ((t,n)::args,(t,[])::coreg) + in if n<0 + then error("Type constructor cannot have negative number of arguments") + else foldl add_type (ac,ts) + end; + +(* 'extend' takes the above described check- and extend-functions to + extend a given type signature with new classes and new type declarations *) + +fun extend (TySg{classes,default,subclass,args,coreg}, + newclasses,newdefault,types,arities) = +let val (classes',subclass') = extend_classes(classes,subclass,newclasses); + val (args',coreg') = foldl add_types ((args,coreg),types); + val old_coreg = map #1 coreg; + fun is_old(c) = if c mem old_coreg then () else undcl_type_err(c); + fun is_new(c) = if c mem old_coreg then twice(c) else (); + val coreg'' = foldl (coregular (classes',subclass',args')) + (coreg',min_domain subclass' arities); + val coreg''' = close (coreg'',subclass'); + val default' = if null newdefault then default else newdefault; +in TySg{classes=classes', default=default',subclass=subclass', args=args', + coreg=coreg'''} end; + + +(* 'assoc_union' merges two association lists if the contents associated + the keys are lists *) + +fun assoc_union (as1,[]) = as1 + | assoc_union (as1,(key,l2)::as2) = case assoc (as1,key) of + Some(l1) => assoc_union (overwrite(as1,(key,l1 union l2)),as2) + | None => assoc_union ((key,l2)::as1,as2); + + +fun trcl r = + let val r' = transitive_closure r + in if exists (op mem) r' then error("Cyclic class structure!") else r' end; + + +(* 'merge_coreg' builds the union of two 'coreg' lists; + it only checks the two restriction conditions and inserts afterwards + all elements of the second list into the first one *) + +fun merge_coreg classes subclass1 = + let fun test_ar classes (t,ars1) (coreg1,(s,w)) = + (is_unique_decl coreg1 (t,(s,w)); + restr2 classes (subclass1,coreg1) (t,(s,w)); + overwrite (coreg1,(t,(s,w) ins ars1))); + + fun merge_c (coreg1,(c as (t,ars2))) = case assoc (coreg1,t) of + Some(ars1) => foldl (test_ar classes (t,ars1)) (coreg1,ars2) + | None => c::coreg1 + in foldl merge_c end; + +fun merge_args(args,(t,n)) = case assoc(args,t) of + Some(m) => if m=n then args else varying_decls(t) + | None => (t,n)::args; + +(* 'merge' takes the above declared functions to merge two type signatures *) + +fun merge(TySg{classes=classes1,default=default1,subclass=subclass1,args=args1, + coreg=coreg1}, + TySg{classes=classes2,default=default2,subclass=subclass2,args=args2, + coreg=coreg2}) = + let val classes' = classes1 union classes2; + val subclass' = trcl (assoc_union (subclass1,subclass2)); + val args' = foldl merge_args (args1,args2) + val coreg' = merge_coreg classes' subclass' (coreg1,coreg2); + val default' = min_sort subclass' (default1 @ default2) + in TySg{classes=classes' , default=default',subclass=subclass', args=args', + coreg=coreg'} + end; + +(**** TYPE INFERENCE ****) + +(* + +Input: +- a 'raw' term which contains only dummy types and some explicit type + constraints encoded as terms. +- the expected type of the term. + +Output: +- the correctly typed term +- the substitution needed to unify the actual type of the term with its + expected type; only the TVars in the expected type are included. + +During type inference all TVars in the term have negative index. This keeps +them apart from normal TVars, which is essential, because at the end the type +of the term is unified with the expected type, which contains normal TVars. + +1. Add initial type information to the term (add_types). + This freezes (freeze_vars) TVars in explicitly provided types (eg + constraints or defaults) by turning them into TFrees. +2. Carry out type inference, possibly introducing new negative TVars. +3. Unify actual and expected type. +4. Turn all (negative) TVars into unique new TFrees (freeze). +5. Thaw all TVars frozen in step 1 (thaw_vars). + +*) + +(*Raised if types are not unifiable*) +exception TUNIFY; + +val tyvar_count = ref(~1); + +fun tyinit() = (tyvar_count := ~1); + +fun new_tvar_inx() = (tyvar_count := !tyvar_count-1; !tyvar_count) + +(* +Generate new TVar. Index is < ~1 to distinguish it from TVars generated from +variable names (see id_type). Name is arbitrary because index is new. +*) + +fun gen_tyvar(S) = TVar(("'a", new_tvar_inx()),S); +fun new_id_type(a) = TVar(("'"^a, new_tvar_inx()),[]); + +(*Occurs check: type variable occurs in type?*) +fun occ v tye = + let fun occ(Type(_,Ts)) = exists occ Ts + | occ(TFree _) = false + | occ(TVar(w,_)) = v=w orelse + (case assoc(tye,w) of + None => false + | Some U => occ U); + in occ end; + +(*Chase variable assignments in tye. + If devar (T,tye) returns a type var then it must be unassigned.*) +fun devar (T as TVar(v,_), tye) = (case assoc(tye,v) of + Some U => devar (U,tye) + | None => T) + | devar (T,tye) = T; + + +(* 'dom' returns for a type constructor t the list of those domains + which deliver a given range class C *) + +fun dom coreg t C = case assoc2 (coreg, (t,C)) of + Some(Ss) => Ss + | None => raise TUNIFY; + + +(* 'Dom' returns the union of all domain lists of 'dom' for a given sort S + (i.e. a set of range classes ); the union is carried out elementwise + for the seperate sorts in the domains *) + +fun Dom (subclass,coreg) (t,S) = + let val domlist = map (dom coreg t) S; + in if null domlist then [] + else foldl (elementwise_union subclass) (hd domlist,tl domlist) + end; + + +fun W ((T,S),tsig as TySg{subclass,coreg,...},tye) = + let fun Wd ((T,S),tye) = W ((devar (T,tye),S),tsig,tye) + fun Wk(T as TVar(v,S')) = + if sortorder subclass (S',S) then tye + else (v,gen_tyvar(union_sort subclass (S',S)))::tye + | Wk(T as TFree(v,S')) = if sortorder subclass (S',S) then tye + else raise TUNIFY + | Wk(T as Type(f,Ts)) = + if null S then tye + else foldr Wd (Ts~~(Dom (subclass,coreg) (f,S)) ,tye) + in Wk(T) end; + + +(* Order-sorted Unification of Types (U) *) + + +(* Precondition: both types are well-formed w.r.t. type constructor arities *) +fun unify (tsig as TySg{subclass,coreg,...}) = + let fun unif ((T,U),tye) = + case (devar(T,tye), devar(U,tye)) of + (T as TVar(v,S1), U as TVar(w,S2)) => + if v=w then tye else + if sortorder subclass (S1,S2) then (w,T)::tye else + if sortorder subclass (S2,S1) then (v,U)::tye + else let val nu = gen_tyvar (union_sort subclass (S1,S2)) + in (v,nu)::(w,nu)::tye end + | (T as TVar(v,S), U) => + if occ v tye U then raise TUNIFY else W ((U,S),tsig,(v,U)::tye) + | (U, T as TVar (v,S)) => + if occ v tye U then raise TUNIFY else W ((U,S),tsig,(v,U)::tye) + | (Type(a,Ts),Type(b,Us)) => + if a<>b then raise TUNIFY else foldr unif (Ts~~Us,tye) + | (T,U) => if T=U then tye else raise TUNIFY + in unif end; + +(*Instantiation of type variables in types*) +(*Pre: instantiations obey restrictions! *) +fun inst_typ tye = + let fun inst(Type(a,Ts)) = Type(a, map inst Ts) + | inst(T as TFree _) = T + | inst(T as TVar(v,_)) = + (case assoc(tye,v) of Some U => inst U | None => T) + in inst end; + +(*Type inference for polymorphic term*) +fun infer tsig = + let fun inf(Ts, Const (_,T), tye) = (T,tye) + | inf(Ts, Free (_,T), tye) = (T,tye) + | inf(Ts, Bound i, tye) = ((nth_elem(i,Ts) , tye) + handle LIST _=> raise TYPE ("loose bound variable", [], [Bound i])) + | inf(Ts, Var (_,T), tye) = (T,tye) + | inf(Ts, Abs (_,T,body), tye) = + let val (U,tye') = inf(T::Ts, body, tye) in (T-->U, tye') end + | inf(Ts, f$u, tye) = + let val (U,tyeU) = inf(Ts, u, tye); + val (T,tyeT) = inf(Ts, f, tyeU); + fun err s = + raise TYPE(s, [inst_typ tyeT T, inst_typ tyeT U], [f$u]) + in case T of + Type("fun",[T1,T2]) => + ( (T2, unify tsig ((T1,U), tyeT)) + handle TUNIFY => err"type mismatch in application" ) + | TVar _ => + let val T2 = gen_tyvar([]) + in (T2, unify tsig ((T, U-->T2), tyeT)) + handle TUNIFY => err"type mismatch in application" + end + | _ => err"rator must have function type" + end + in inf end; + +fun freeze_vars(Type(a,Ts)) = Type(a,map freeze_vars Ts) + | freeze_vars(T as TFree _) = T + | freeze_vars(TVar(v,S)) = TFree(Syntax.string_of_vname v, S); + +(* Attach a type to a constant *) +fun type_const (a,T) = Const(a, incr_tvar (new_tvar_inx()) T); + +(*Find type of ident. If not in table then use ident's name for tyvar + to get consistent typing.*) +fun type_of_ixn(types,ixn as (a,_)) = + case types ixn of Some T => freeze_vars T | None => TVar(("'"^a,~1),[]); + +fun constrain(term,T) = Const(Syntax.constrainC,T-->T) $ term; +fun constrainAbs(Abs(a,_,body),T) = Abs(a,T,body); + +(* + +Attach types to a term. Input is a "parse tree" containing dummy types. +Type constraints are translated and checked for validity wrt tsig. +TVars in constraints are frozen. + +The atoms in the resulting term satisfy the following spec: + +Const(a,T): + T is a renamed copy of the generic type of a; renaming decreases index of + all TVars by new_tvar_inx(), which is less than ~1. The index of all TVars + in the generic type must be 0 for this to work! + +Free(a,T), Var(ixn,T): + T is either the frozen default type of a or TVar(("'"^a,~1),[]) + +Abs(a,T,_): + T is either a type constraint or TVar(("'"^a,i),[]), where i is generated + by new_tvar_inx(). Thus different abstractions can have the bound variables + of the same name but different types. + +*) + +fun add_types (tsig, const_tab, types, sorts, string_of_typ) = + let val S0 = defaultS tsig; + fun defS0 ixn = case sorts ixn of Some S => S | None => S0; + fun prepareT(typ) = + let val T = Syntax.typ_of_term defS0 typ; + val T' = freeze_vars T + in case type_errors (tsig,string_of_typ) (T,[]) of + [] => T' + | errs => raise TYPE(cat_lines errs,[T],[]) + end + fun add (Const(a,_)) = + (case Symtab.lookup(const_tab, a) of + Some T => type_const(a,T) + | None => raise TYPE ("No such constant: "^a, [], [])) + | add (Bound i) = Bound i + | add (Free(a,_)) = + (case Symtab.lookup(const_tab, a) of + Some T => type_const(a,T) + | None => Free(a, type_of_ixn(types,(a,~1)))) + | add (Var(ixn,_)) = Var(ixn, type_of_ixn(types,ixn)) + | add (Abs(a,_,body)) = Abs(a, new_id_type a, add body) + | add ((f as Const(a,_)$t1)$t2) = + if a=Syntax.constrainC then constrain(add t1,prepareT t2) else + if a=Syntax.constrainAbsC then constrainAbs(add t1,prepareT t2) + else add f $ add t2 + | add (f$t) = add f $ add t + in add end; + + +(* Post-Processing *) + + +(*Instantiation of type variables in terms*) +fun inst_types tye = map_term_types (inst_typ tye); + +(*Delete explicit constraints -- occurrences of "_constrain" *) +fun unconstrain (Abs(a,T,t)) = Abs(a, T, unconstrain t) + | unconstrain ((f as Const(a,_)) $ t) = + if a=Syntax.constrainC then unconstrain t + else unconstrain f $ unconstrain t + | unconstrain (f$t) = unconstrain f $ unconstrain t + | unconstrain (t) = t; + + +(* Turn all TVars which satisfy p into new TFrees *) +fun freeze p t = + let val fs = add_term_tfree_names(t,[]); + val inxs = filter p (add_term_tvar_ixns(t,[])); + val vmap = inxs ~~ variantlist(map #1 inxs, fs); + fun free(Type(a,Ts)) = Type(a, map free Ts) + | free(T as TVar(v,S)) = + (case assoc(vmap,v) of None => T | Some(a) => TFree(a,S)) + | free(T as TFree _) = T + in map_term_types free t end; + +(* Thaw all TVars that were frozen in freeze_vars *) +fun thaw_vars(Type(a,Ts)) = Type(a, map thaw_vars Ts) + | thaw_vars(T as TFree(a,S)) = (case explode a of + "?"::"'"::vn => let val ((b,i),_) = Syntax.scan_varname vn + in TVar(("'"^b,i),S) end + | _ => T) + | thaw_vars(T) = T; + + +fun restrict tye = + let fun clean(tye1, ((a,i),T)) = + if i < 0 then tye1 else ((a,i),inst_typ tye T) :: tye1 + in foldl clean ([],tye) end + + +(*Infer types for term t using tables. Check that t's type and T unify *) + +fun infer_term (tsig, const_tab, types, sorts, string_of_typ, T, t) = + let val u = add_types (tsig, const_tab, types, sorts, string_of_typ) t; + val (U,tye) = infer tsig ([], u, []); + val uu = unconstrain u; + val tye' = unify tsig ((T,U),tye) handle TUNIFY => raise TYPE + ("Term does not have expected type", [T, U], [inst_types tye uu]) + val Ttye = restrict tye' (* restriction to TVars in T *) + val all = Const("", Type("", map snd Ttye)) $ (inst_types tye' uu) + (* all is a dummy term which contains all exported TVars *) + val Const(_,Type(_,Ts)) $ u'' = + map_term_types thaw_vars (freeze (fn (_,i) => i<0) all) + (* turn all internally generated TVars into TFrees + and thaw all initially frozen TVars *) + in (u'', (map fst Ttye) ~~ Ts) end; + +fun infer_types args = (tyinit(); infer_term args); + + +(* Turn TFrees into TVars to allow types & axioms to be written without "?" *) +fun varifyT(Type(a,Ts)) = Type(a,map varifyT Ts) + | varifyT(TFree(a,S)) = TVar((a,0),S) + | varifyT(T) = T; + +(* Turn TFrees except those in fixed into new TVars *) +fun varify(t,fixed) = + let val fs = add_term_tfree_names(t,[]) \\ fixed; + val ixns = add_term_tvar_ixns(t,[]); + val fmap = fs ~~ variantlist(fs, map #1 ixns) + fun thaw(Type(a,Ts)) = Type(a, map thaw Ts) + | thaw(T as TVar _) = T + | thaw(T as TFree(a,S)) = + (case assoc(fmap,a) of None => T | Some b => TVar((b,0),S)) + in map_term_types thaw t end; + + +end; diff -r 000000000000 -r a5a9c433f639 src/Pure/unify.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/unify.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,680 @@ +(* Title: unify + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright Cambridge University 1992 + +Higher-Order Unification + +Potential problem: type of Vars is often ignored, so two Vars with same +indexname but different types can cause errors! +*) + + +signature UNIFY = +sig + structure Sign: SIGN + structure Envir : ENVIR + structure Sequence : SEQUENCE + (*references for control and tracing*) + val trace_bound: int ref + val trace_simp: bool ref + val trace_types: bool ref + val search_bound: int ref + (*other exports*) + val combound : (term*int*int) -> term + val rlist_abs: (string*typ)list * term -> term + val smash_unifiers : Sign.sg * Envir.env * (term*term)list + -> (Envir.env Sequence.seq) + val unifiers: Sign.sg * Envir.env * ((term*term)list) + -> (Envir.env * (term * term)list) Sequence.seq +end; + +functor UnifyFun (structure Sign: SIGN and Envir: ENVIR and Sequence: SEQUENCE + and Pattern:PATTERN + sharing type Sign.sg = Pattern.sg + and type Envir.env = Pattern.env) + : UNIFY = +struct + +structure Sign = Sign; +structure Envir = Envir; +structure Sequence = Sequence; +structure Pretty = Sign.Syntax.Pretty; + +(*Unification options*) + +val trace_bound = ref 10 (*tracing starts above this depth, 0 for full*) +and search_bound = ref 20 (*unification quits above this depth*) +and trace_simp = ref false (*print dpairs before calling SIMPL*) +and trace_types = ref false (*announce potential incompleteness + of type unification*) + +val sgr = ref(Sign.pure); + +type binderlist = (string*typ) list; + +type dpair = binderlist * term * term; + +fun body_type(Envir.Envir{iTs,...}) = +let fun bT(Type("fun",[_,T])) = bT T + | bT(T as TVar(ixn,_)) = (case assoc(iTs,ixn) of + None => T | Some(T') => bT T') + | bT T = T +in bT end; + +fun binder_types(Envir.Envir{iTs,...}) = +let fun bTs(Type("fun",[T,U])) = T :: bTs U + | bTs(T as TVar(ixn,_)) = (case assoc(iTs,ixn) of + None => [] | Some(T') => bTs T') + | bTs _ = [] +in bTs end; + +fun strip_type env T = (binder_types env T, body_type env T); + + +(*Put a term into head normal form for unification. + Operands need not be in normal form. Does eta-expansions on the head, + which involves renumbering (thus copying) the args. To avoid this + inefficiency, avoid partial application: if an atom is applied to + any arguments at all, apply it to its full number of arguments. + For + rbinder = [(x1,T),...,(xm,Tm)] (user's var names preserved!) + args = [arg1,...,argn] + the value of + (xm,...,x1)(head(arg1,...,argn)) remains invariant. +*) + +local exception SAME +in + fun head_norm (env,t) : term = + let fun hnorm (Var (v,T)) = + (case Envir.lookup (env,v) of + Some u => head_norm (env, u) + | None => raise SAME) + | hnorm (Abs(a,T,body)) = Abs(a, T, hnorm body) + | hnorm (Abs(_,_,body) $ t) = + head_norm (env, subst_bounds([t], body)) + | hnorm (f $ t) = + (case hnorm f of + Abs(_,_,body) => + head_norm (env, subst_bounds([t], body)) + | nf => nf $ t) + | hnorm _ = raise SAME + in hnorm t handle SAME=> t end +end; + + +(*finds type of term without checking that combinations are consistent + rbinder holds types of bound variables*) +fun fastype (Envir.Envir{iTs,...}) = +let val funerr = "fastype: expected function type"; + fun fast(rbinder, f$u) = + (case (fast (rbinder, f)) of + Type("fun",[_,T]) => T + | TVar(ixn,_) => + (case assoc(iTs,ixn) of + Some(Type("fun",[_,T])) => T + | _ => raise TERM(funerr, [f$u])) + | _ => raise TERM(funerr, [f$u])) + | fast (rbinder, Const (_,T)) = T + | fast (rbinder, Free (_,T)) = T + | fast (rbinder, Bound i) = + (#2 (nth_elem (i,rbinder)) + handle LIST _=> raise TERM("fastype: Bound", [Bound i])) + | fast (rbinder, Var (_,T)) = T + | fast (rbinder, Abs (_,T,u)) = T --> fast (("",T) :: rbinder, u) +in fast end; + + +(*Eta normal form*) +fun eta_norm(env as Envir.Envir{iTs,...}) = + let fun etif (Type("fun",[T,U]), t) = + Abs("", T, etif(U, incr_boundvars 1 t $ Bound 0)) + | etif (TVar(ixn,_),t) = + (case assoc(iTs,ixn) of + None => t | Some(T) => etif(T,t)) + | etif (_,t) = t; + fun eta_nm (rbinder, Abs(a,T,body)) = + Abs(a, T, eta_nm ((a,T)::rbinder, body)) + | eta_nm (rbinder, t) = etif(fastype env (rbinder,t), t) + in eta_nm end; + + +(*OCCURS CHECK + Does the uvar occur in the term t? + two forms of search, for whether there is a rigid path to the current term. + "seen" is list of variables passed thru, is a memo variable for sharing. + This version searches for nonrigid occurrence, returns true if found. *) +fun occurs_terms (seen: (indexname list) ref, + env: Envir.env, v: indexname, ts: term list): bool = + let fun occurs [] = false + | occurs (t::ts) = occur t orelse occurs ts + and occur (Const _) = false + | occur (Bound _) = false + | occur (Free _) = false + | occur (Var (w,_)) = + if w mem !seen then false + else if v=w then true + (*no need to lookup: v has no assignment*) + else (seen := w:: !seen; + case Envir.lookup(env,w) of + None => false + | Some t => occur t) + | occur (Abs(_,_,body)) = occur body + | occur (f$t) = occur t orelse occur f + in occurs ts end; + + + +(* f(a1,...,an) ----> (f, [a1,...,an]) using the assignments*) +fun head_of_in (env,t) : term = case t of + f$_ => head_of_in(env,f) + | Var (v,_) => (case Envir.lookup(env,v) of + Some u => head_of_in(env,u) | None => t) + | _ => t; + + +datatype occ = NoOcc | Nonrigid | Rigid; + +(* Rigid occur check +Returns Rigid if it finds a rigid occurrence of the variable, + Nonrigid if it finds a nonrigid path to the variable. + NoOcc otherwise. + Continues searching for a rigid occurrence even if it finds a nonrigid one. + +Condition for detecting non-unifable terms: [ section 5.3 of Huet (1975) ] + a rigid path to the variable, appearing with no arguments. +Here completeness is sacrificed in order to reduce danger of divergence: + reject ALL rigid paths to the variable. +Could check for rigid paths to bound variables that are out of scope. +Not necessary because the assignment test looks at variable's ENTIRE rbinder. + +Treatment of head(arg1,...,argn): +If head is a variable then no rigid path, switch to nonrigid search +for arg1,...,argn. +If head is an abstraction then possibly no rigid path (head could be a + constant function) so again use nonrigid search. Happens only if + term is not in normal form. + +Warning: finds a rigid occurrence of ?f in ?f(t). + Should NOT be called in this case: there is a flex-flex unifier +*) +fun rigid_occurs_term (seen: (indexname list)ref, env, v: indexname, t) = + let fun nonrigid t = if occurs_terms(seen,env,v,[t]) then Nonrigid + else NoOcc + fun occurs [] = NoOcc + | occurs (t::ts) = + (case occur t of + Rigid => Rigid + | oc => (case occurs ts of NoOcc => oc | oc2 => oc2)) + and occomb (f$t) = + (case occur t of + Rigid => Rigid + | oc => (case occomb f of NoOcc => oc | oc2 => oc2)) + | occomb t = occur t + and occur (Const _) = NoOcc + | occur (Bound _) = NoOcc + | occur (Free _) = NoOcc + | occur (Var (w,_)) = + if w mem !seen then NoOcc + else if v=w then Rigid + else (seen := w:: !seen; + case Envir.lookup(env,w) of + None => NoOcc + | Some t => occur t) + | occur (Abs(_,_,body)) = occur body + | occur (t as f$_) = (*switch to nonrigid search?*) + (case head_of_in (env,f) of + Var (w,_) => (*w is not assigned*) + if v=w then Rigid + else nonrigid t + | Abs(_,_,body) => nonrigid t (*not in normal form*) + | _ => occomb t) + in occur t end; + + +exception CANTUNIFY; (*Signals non-unifiability. Does not signal errors!*) +exception ASSIGN; (*Raised if not an assignment*) + + +fun unify_types(T,U, env as Envir.Envir{asol,iTs,maxidx}) = + if T=U then env else + let val iTs' = Sign.Type.unify (#tsig(Sign.rep_sg (!sgr))) ((U,T),iTs) + in Envir.Envir{asol=asol,maxidx=maxidx,iTs=iTs'} + end handle Sign.Type.TUNIFY => raise CANTUNIFY; + +fun test_unify_types(args as (T,U,_)) = +let val sot = Sign.string_of_typ (!sgr); + fun warn() = writeln("Potential loss of completeness: "^sot U^" = "^sot T); + val env' = unify_types(args) +in if is_TVar(T) orelse is_TVar(U) then warn() else (); + env' +end; + +(*Is the term eta-convertible to a single variable with the given rbinder? + Examples: ?a ?f(B.0) ?g(B.1,B.0) + Result is var a for use in SIMPL. *) +fun get_eta_var ([], _, Var vT) = vT + | get_eta_var (_::rbinder, n, f $ Bound i) = + if n=i then get_eta_var (rbinder, n+1, f) + else raise ASSIGN + | get_eta_var _ = raise ASSIGN; + + +(* ([xn,...,x1], t) ======> (x1,...,xn)t *) +fun rlist_abs ([], body) = body + | rlist_abs ((a,T)::pairs, body) = rlist_abs(pairs, Abs(a, T, body)); + + +(*Solve v=u by assignment -- "fixedpoint" to Huet -- if v not in u. + If v occurs rigidly then nonunifiable. + If v occurs nonrigidly then must use full algorithm. *) +fun assignment (env, rbinder, t, u) = + let val (v,T) = get_eta_var(rbinder,0,t) + in case rigid_occurs_term (ref[], env, v, u) of + NoOcc => let val env = unify_types(body_type env T, + fastype env (rbinder,u),env) + in Envir.update ((v, rlist_abs(rbinder,u)), env) end + | Nonrigid => raise ASSIGN + | Rigid => raise CANTUNIFY + end; + + +(*Extends an rbinder with a new disagreement pair, if both are abstractions. + Tries to unify types of the bound variables! + Checks that binders have same length, since terms should be eta-normal; + if not, raises TERM, probably indicating type mismatch. + Uses variable a (unless the null string) to preserve user's naming.*) +fun new_dpair (rbinder, Abs(a,T,body1), Abs(b,U,body2), env) = + let val env' = unify_types(T,U,env) + val c = if a="" then b else a + in new_dpair((c,T) :: rbinder, body1, body2, env') end + | new_dpair (_, Abs _, _, _) = raise TERM ("new_dpair", []) + | new_dpair (_, _, Abs _, _) = raise TERM ("new_dpair", []) + | new_dpair (rbinder, t1, t2, env) = ((rbinder, t1, t2), env); + + +fun head_norm_dpair (env, (rbinder,t,u)) : dpair * Envir.env = + new_dpair (rbinder, + eta_norm env (rbinder, head_norm(env,t)), + eta_norm env (rbinder, head_norm(env,u)), env); + + + +(*flexflex: the flex-flex pairs, flexrigid: the flex-rigid pairs + Does not perform assignments for flex-flex pairs: + may create nonrigid paths, which prevent other assignments*) +fun SIMPL0 (dp0, (env,flexflex,flexrigid)) + : Envir.env * dpair list * dpair list = + let val (dp as (rbinder,t,u), env) = head_norm_dpair(env,dp0); + fun SIMRANDS(f$t, g$u, env) = + SIMPL0((rbinder,t,u), SIMRANDS(f,g,env)) + | SIMRANDS (t as _$_, _, _) = + raise TERM ("SIMPL: operands mismatch", [t,u]) + | SIMRANDS (t, u as _$_, _) = + raise TERM ("SIMPL: operands mismatch", [t,u]) + | SIMRANDS(_,_,env) = (env,flexflex,flexrigid); + in case (head_of t, head_of u) of + (Var(_,T), Var(_,U)) => + let val T' = body_type env T and U' = body_type env U; + val env = unify_types(T',U',env) + in (env, dp::flexflex, flexrigid) end + | (Var _, _) => + ((assignment (env,rbinder,t,u), flexflex, flexrigid) + handle ASSIGN => (env, flexflex, dp::flexrigid)) + | (_, Var _) => + ((assignment (env,rbinder,u,t), flexflex, flexrigid) + handle ASSIGN => (env, flexflex, (rbinder,u,t)::flexrigid)) + | (Const(a,T), Const(b,U)) => + if a=b then SIMRANDS(t,u, unify_types(T,U,env)) + else raise CANTUNIFY + | (Bound i, Bound j) => + if i=j then SIMRANDS(t,u,env) else raise CANTUNIFY + | (Free(a,T), Free(b,U)) => + if a=b then SIMRANDS(t,u, unify_types(T,U,env)) + else raise CANTUNIFY + | _ => raise CANTUNIFY + end; + + +(* changed(env,t) checks whether the head of t is a variable assigned in env*) +fun changed (env, f$_) = changed (env,f) + | changed (env, Var (v,_)) = + (case Envir.lookup(env,v) of None=>false | _ => true) + | changed _ = false; + + +(*Recursion needed if any of the 'head variables' have been updated + Clever would be to re-do just the affected dpairs*) +fun SIMPL (env,dpairs) : Envir.env * dpair list * dpair list = + let val all as (env',flexflex,flexrigid) = + foldr SIMPL0 (dpairs, (env,[],[])); + val dps = flexrigid@flexflex + in if exists (fn ((_,t,u)) => changed(env',t) orelse changed(env',u)) dps + then SIMPL(env',dps) else all + end; + + +(*computes t(Bound(n+k-1),...,Bound(n)) *) +fun combound (t, n, k) = + if k>0 then combound (t,n+1,k-1) $ (Bound n) else t; + + +(*Makes the terms E1,...,Em, where Ts = [T...Tm]. + Each Ei is ?Gi(B.(n-1),...,B.0), and has type Ti + The B.j are bound vars of binder. + The terms are not made in eta-normal-form, SIMPL does that later. + If done here, eta-expansion must be recursive in the arguments! *) +fun make_args name (binder: typ list, env, []) = (env, []) (*frequent case*) + | make_args name (binder: typ list, env, Ts) : Envir.env * term list = + let fun funtype T = binder--->T; + val (env', vars) = Envir.genvars name (env, map funtype Ts) + in (env', map (fn var=> combound(var, 0, length binder)) vars) end; + + +(*Abstraction over a list of types, like list_abs*) +fun types_abs ([],u) = u + | types_abs (T::Ts, u) = Abs("", T, types_abs(Ts,u)); + +(*Abstraction over the binder of a type*) +fun type_abs (env,T,t) = types_abs(binder_types env T, t); + + +(*MATCH taking "big steps". + Copies u into the Var v, using projection on targs or imitation. + A projection is allowed unless SIMPL raises an exception. + Allocates new variables in projection on a higher-order argument, + or if u is a variable (flex-flex dpair). + Returns long sequence of every way of copying u, for backtracking + For example, projection in ?b'(?a) may be wrong if other dpairs constrain ?a. + The order for trying projections is crucial in ?b'(?a) + NB "vname" is only used in the call to make_args!! *) +fun matchcopy vname = let fun mc(rbinder, targs, u, ed as (env,dpairs)) + : (term * (Envir.env * dpair list))Sequence.seq = +let (*Produce copies of uarg and cons them in front of uargs*) + fun copycons uarg (uargs, (env, dpairs)) = + Sequence.maps(fn (uarg', ed') => (uarg'::uargs, ed')) + (mc (rbinder, targs,eta_norm env (rbinder,head_norm(env,uarg)), + (env, dpairs))); + (*Produce sequence of all possible ways of copying the arg list*) + fun copyargs [] = Sequence.cons( ([],ed), Sequence.null) + | copyargs (uarg::uargs) = + Sequence.flats (Sequence.maps (copycons uarg) (copyargs uargs)); + val (uhead,uargs) = strip_comb u; + val base = body_type env (fastype env (rbinder,uhead)); + fun joinargs (uargs',ed') = (list_comb(uhead,uargs'), ed'); + (*attempt projection on argument with given typ*) + val Ts = map (curry (fastype env) rbinder) targs; + fun projenv (head, (Us,bary), targ, tail) = + let val env = if !trace_types then test_unify_types(base,bary,env) + else unify_types(base,bary,env) + in Sequence.seqof (fn () => + let val (env',args) = make_args vname (Ts,env,Us); + (*higher-order projection: plug in targs for bound vars*) + fun plugin arg = list_comb(head_of arg, targs); + val dp = (rbinder, list_comb(targ, map plugin args), u); + val (env2,frigid,fflex) = SIMPL (env', dp::dpairs) + (*may raise exception CANTUNIFY*) + in Some ((list_comb(head,args), (env2, frigid@fflex)), + tail) + end handle CANTUNIFY => Sequence.pull tail) + end handle CANTUNIFY => tail; + (*make a list of projections*) + fun make_projs (T::Ts, targ::targs) = + (Bound(length Ts), T, targ) :: make_projs (Ts,targs) + | make_projs ([],[]) = [] + | make_projs _ = raise TERM ("make_projs", u::targs); + (*try projections and imitation*) + fun matchfun ((bvar,T,targ)::projs) = + (projenv(bvar, strip_type env T, targ, matchfun projs)) + | matchfun [] = (*imitation last of all*) + (case uhead of + Const _ => Sequence.maps joinargs (copyargs uargs) + | Free _ => Sequence.maps joinargs (copyargs uargs) + | _ => Sequence.null) (*if Var, would be a loop!*) +in case uhead of + Abs(a, T, body) => + Sequence.maps(fn (body', ed') => (Abs (a,T,body'), ed')) + (mc ((a,T)::rbinder, + (map (incr_boundvars 1) targs) @ [Bound 0], body, ed)) + | Var (w,uary) => + (*a flex-flex dpair: make variable for t*) + let val (env', newhd) = Envir.genvar (#1 w) (env, Ts---> base) + val tabs = combound(newhd, 0, length Ts) + val tsub = list_comb(newhd,targs) + in Sequence.single (tabs, (env', (rbinder,tsub,u):: dpairs)) + end + | _ => matchfun(rev(make_projs(Ts, targs))) +end +in mc end; + + +(*Call matchcopy to produce assignments to the variable in the dpair*) +fun MATCH (env, (rbinder,t,u), dpairs) + : (Envir.env * dpair list)Sequence.seq = + let val (Var(v,T), targs) = strip_comb t; + val Ts = binder_types env T; + fun new_dset (u', (env',dpairs')) = + (*if v was updated to s, must unify s with u' *) + case Envir.lookup(env',v) of + None => (Envir.update ((v, types_abs(Ts, u')), env'), dpairs') + | Some s => (env', ([], s, types_abs(Ts, u'))::dpairs') + in Sequence.maps new_dset + (matchcopy (#1 v) (rbinder, targs, u, (env,dpairs))) + end; + + + +(**** Flex-flex processing ****) + +(*At end of unification, do flex-flex assignments like ?a -> ?f(?b) + Attempts to update t with u, raising ASSIGN if impossible*) +fun ff_assign(env, rbinder, t, u) : Envir.env = +let val (v,T) = get_eta_var(rbinder,0,t) +in if occurs_terms (ref[], env, v, [u]) then raise ASSIGN + else let val env = unify_types(body_type env T,fastype env (rbinder,u),env) + in Envir.vupdate ((v, rlist_abs(rbinder, u)), env) end +end; + + +(*Flex argument: a term, its type, and the index that refers to it.*) +type flarg = {t: term, T: typ, j: int}; + + +(*Form the arguments into records for deletion/sorting.*) +fun flexargs ([],[],[]) = [] : flarg list + | flexargs (j::js, t::ts, T::Ts) = {j=j, t=t, T=T} :: flexargs(js,ts,Ts) + | flexargs _ = error"flexargs"; + + +(*If an argument contains a banned Bound, then it should be deleted. + But if the path is flexible, this is difficult; the code gives up!*) +exception CHANGE and CHANGE_FAIL; (*rigid and flexible occurrences*) + + +(*Squash down indices at level >=lev to delete the js from a term. + flex should initially be false, since the empty path is rigid.*) +fun change_bnos (lev, js, flex) t = + let val (head,args) = strip_comb t + val flex' = flex orelse is_Var head + val head' = case head of + Bound i => + if i j < i-lev) js)) + | Abs (a,T,t) => Abs (a, T, change_bnos(lev+1, js, flex) t) + | _ => head + in list_comb (head', map (change_bnos (lev, js, flex')) args) + end; + + +(*Change indices, delete the argument if it contains a banned Bound*) +fun change_arg js ({j,t,T}, args) : flarg list = + {j=j, t= change_bnos(0,js,false)t, T=T} :: args handle CHANGE => args; + + +(*Sort the arguments to create assignments if possible: + create eta-terms like ?g(B.1,B.0) *) +fun arg_less ({t= Bound i1,...}, {t= Bound i2,...}) = (i2 U) + val body = list_comb(v', map (Bound o #j) args) + val env2 = Envir.vupdate (((v, types_abs(Ts, body)), env')) + (*the vupdate affects ts' if they contain v*) + in + (env2, Envir.norm_term env2 (list_comb(v',ts'))) + end + end; + + +(*Add tpair if not trivial or already there. + Should check for swapped pairs??*) +fun add_tpair (rbinder, (t0,u0), tpairs) : (term*term) list = + if t0 aconv u0 then tpairs + else + let val t = rlist_abs(rbinder, t0) and u = rlist_abs(rbinder, u0); + fun same(t',u') = (t aconv t') andalso (u aconv u') + in if exists same tpairs then tpairs else (t,u)::tpairs end; + + +(*Simplify both terms and check for assignments. + Bound vars in the binder are "banned" unless used in both t AND u *) +fun clean_ffpair ((rbinder, t, u), (env,tpairs)) = + let val loot = loose_bnos t and loou = loose_bnos u + fun add_index (((a,T), j), (bnos, newbinder)) = + if j mem loot andalso j mem loou + then (bnos, (a,T)::newbinder) + else (j::bnos, newbinder); + val indices = 0 upto (length rbinder - 1); + val (banned,rbin') = foldr add_index (rbinder~~indices, ([],[])); + val (env', t') = clean_term banned (env, t); + val (env'',u') = clean_term banned (env',u) + in (ff_assign(env'', rbin', t', u'), tpairs) + handle ASSIGN => (ff_assign(env'', rbin', u', t'), tpairs) + handle ASSIGN => (env'', add_tpair(rbin', (t',u'), tpairs)) + end + handle CHANGE_FAIL => (env, add_tpair(rbinder, (t,u), tpairs)); + + +(*IF the flex-flex dpair is an assignment THEN do it ELSE put in tpairs + eliminates trivial tpairs like t=t, as well as repeated ones + trivial tpairs can easily escape SIMPL: ?A=t, ?A=?B, ?B=t gives t=t + Resulting tpairs MAY NOT be in normal form: assignments may occur here.*) +fun add_ffpair ((rbinder,t0,u0), (env,tpairs)) + : Envir.env * (term*term)list = + let val t = Envir.norm_term env t0 and u = Envir.norm_term env u0 + in case (head_of t, head_of u) of + (Var(v,T), Var(w,U)) => (*Check for identical variables...*) + if v=w then (*...occur check would falsely return true!*) + if T=U then (env, add_tpair (rbinder, (t,u), tpairs)) + else raise TERM ("add_ffpair: Var name confusion", [t,u]) + else if xless(v,w) then (*prefer to update the LARGER variable*) + clean_ffpair ((rbinder, u, t), (env,tpairs)) + else clean_ffpair ((rbinder, t, u), (env,tpairs)) + | _ => raise TERM ("add_ffpair: Vars expected", [t,u]) + end; + + +(*Print a tracing message + list of dpairs. + In t==u print u first because it may be rigid or flexible -- + t is always flexible.*) +fun print_dpairs msg (env,dpairs) = + let fun pdp (rbinder,t,u) = + let fun termT t = Sign.pretty_term (!sgr) + (Envir.norm_term env (rlist_abs(rbinder,t))) + val bsymbs = [termT u, Pretty.str" =?=", Pretty.brk 1, + termT t]; + in writeln(Pretty.string_of(Pretty.blk(0,bsymbs))) end; + in writeln msg; seq pdp dpairs end; + + +(*Unify the dpairs in the environment. + Returns flex-flex disagreement pairs NOT IN normal form. + SIMPL may raise exception CANTUNIFY. *) +fun hounifiers (sg,env, tus : (term*term)list) + : (Envir.env * (term*term)list)Sequence.seq = + let fun add_unify tdepth ((env,dpairs), reseq) = + Sequence.seqof (fn()=> + let val (env',flexflex,flexrigid) = + (if tdepth> !trace_bound andalso !trace_simp + then print_dpairs "Enter SIMPL" (env,dpairs) else (); + SIMPL (env,dpairs)) + in case flexrigid of + [] => Some (foldr add_ffpair (flexflex, (env',[])), reseq) + | dp::frigid' => + if tdepth > !search_bound then + (prs"***Unification bound exceeded\n"; Sequence.pull reseq) + else + (if tdepth > !trace_bound then + print_dpairs "Enter MATCH" (env',flexrigid@flexflex) + else (); + Sequence.pull (Sequence.its_right (add_unify (tdepth+1)) + (MATCH (env',dp, frigid'@flexflex), reseq))) + end + handle CANTUNIFY => + (if tdepth > !trace_bound then writeln"Failure node" else (); + Sequence.pull reseq)); + val dps = map (fn(t,u)=> ([],t,u)) tus + in sgr := sg; + add_unify 1 ((env,dps), Sequence.null) + end; + +fun unifiers(params) = + Sequence.cons((Pattern.unify(params), []), Sequence.null) + handle Pattern.Unif => Sequence.null + | Pattern.Pattern => hounifiers(params); + + +(*For smash_flexflex1*) +fun var_head_of (env,t) : indexname * typ = + case head_of (strip_abs_body (Envir.norm_term env t)) of + Var(v,T) => (v,T) + | _ => raise CANTUNIFY; (*not flexible, cannot use trivial substitution*) + + +(*Eliminate a flex-flex pair by the trivial substitution, see Huet (1975) + Unifies ?f(t1...rm) with ?g(u1...un) by ?f -> %x1...xm.?a, ?g -> %x1...xn.?a + Unfortunately, unifies ?f(t,u) with ?g(t,u) by ?f, ?g -> %(x,y)?a, + though just ?g->?f is a more general unifier. + Unlike Huet (1975), does not smash together all variables of same type -- + requires more work yet gives a less general unifier (fewer variables). + Handles ?f(t1...rm) with ?f(u1...um) to avoid multiple updates. *) +fun smash_flexflex1 ((t,u), env) : Envir.env = + let val (v,T) = var_head_of (env,t) + and (w,U) = var_head_of (env,u); + val (env', var) = Envir.genvar (#1v) (env, body_type env T) + val env'' = Envir.vupdate((w, type_abs(env',U,var)), env') + in if (v,T)=(w,U) then env'' (*the other update would be identical*) + else Envir.vupdate((v, type_abs(env',T,var)), env'') + end; + + +(*Smash all flex-flexpairs. Should allow selection of pairs by a predicate?*) +fun smash_flexflex (env,tpairs) : Envir.env = + foldr smash_flexflex1 (tpairs, env); + +(*Returns unifiers with no remaining disagreement pairs*) +fun smash_unifiers (sg, env, tus) : Envir.env Sequence.seq = + Sequence.maps smash_flexflex (unifiers(sg,env,tus)); + +end; diff -r 000000000000 -r a5a9c433f639 src/Tools/agrep --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/agrep Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,2 @@ +#! /bin/csh +grep "$*" {Pure/Syntax,Pure/Thy}/*ML */*ML */ex/*ML diff -r 000000000000 -r a5a9c433f639 src/Tools/expandshort --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/expandshort Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,31 @@ +#! /bin/sh +# +#expandshort - shell script to expand shorthand goal commands +# ALSO contracts uses of resolve_tac, dresolve_tac, eresolve_tac, +# rewrite_goals_tac on 1-element lists +# +# Usage: +# expandshort FILE1 ... FILEn +# +# leaves previous versions as XXX~~ +# +for f in $* +do +echo Expanding shorthands in $f. \ Backup file is $f~~ +mv $f $f~~; sed -e ' +s/^ba \([0-9]*\); *$/by (assume_tac \1);/ +s/^br \(.*\) \([0-9]*\); *$/by (rtac \1 \2);/ +s/^brs \(.*\) \([0-9]*\); *$/by (resolve_tac \1 \2);/ +s/^bd \(.*\) \([0-9]*\); *$/by (dtac \1 \2);/ +s/^bds \(.*\) \([0-9]*\); *$/by (dresolve_tac \1 \2);/ +s/^be \(.*\) \([0-9]*\); *$/by (etac \1 \2);/ +s/^bes \(.*\) \([0-9]*\); *$/by (eresolve_tac \1 \2);/ +s/^bw \(.*\); *$/by (rewtac \1);/ +s/^bws \(.*\); *$/by (rewrite_goals_tac \1);/ +s/dresolve_tac *\[\([a-zA-Z0-9_]*\)\] */dtac \1 /g +s/eresolve_tac *\[\([a-zA-Z0-9_]*\)\] */etac \1 /g +s/resolve_tac *\[\([a-zA-Z0-9_]*\)\] */rtac \1 /g +s/rewrite_goals_tac *\[\([a-zA-Z0-9_]*\)\]\( *\)/rewtac \1\2/g +' $f~~ > $f +done +echo Finished. diff -r 000000000000 -r a5a9c433f639 src/Tools/make-all --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/make-all Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,169 @@ +#! /bin/sh +# +#make-all: make all systems afresh + +# Creates gzipped log files called makeNNNN.log.gz on each subdirectory and +# displays the last few lines of these files -- this indicates whether +# the "make" failed (whether it terminated due to an error) + +# switches are +# -noforce don't delete old databases/images first +# -clean delete databases/images after use (leaving Pure) +# -notest make databases/images w/o running the examples +# -noexec don't execute, just check settings and Makefiles + +#Environment variables required: +# ISABELLEBIN: the directory to hold Poly/ML databases or New Jersey ML images +# ISABELLECOMP: the ML compiler + +# A typical shell script for /bin/sh is... +# ML_DBASE=/usr/groups/theory/poly2.04/`arch`/ML_dbase +# ISABELLEBIN=/homes/`whoami`/bin +# ISABELLECOMP="poly -noDisplay" +# export ML_DBASE ISABELLEBIN ISABELLECOMP +# nohup make-all $* + +set -e #fail immediately upon errors + +# process command line switches +CLEAN="off"; +FORCE="on"; +TEST="test"; +EXEC="on"; +NO=""; +for A in $* +do + case $A in + -clean) CLEAN="on" ;; + -noforce) FORCE="off" ;; + -notest) TEST="" ;; + -noexec) EXEC="off" + NO="-n" ;; + *) echo "Bad flag for make-all: $A" + echo "Usage: make-all [-noforce] [-clean] [-notest] [-noexec]" + exit ;; + esac +done + +echo Started at `date` +echo Source=`pwd` +echo Destination=${ISABELLEBIN?'No destination directory specified'} +echo force=$FORCE ' ' clean=$CLEAN ' ' +echo Compiler=${ISABELLECOMP?'No compiler specified'} +echo Running on `hostname` +echo Log files will be called make$$.log.gz + +case $FORCE.$EXEC in + on.on) (cd $ISABELLEBIN; rm -f Pure FOL ZF CCL LCF CTT LK Modal HOL Cube FOLP) +esac + +echo +echo +echo '*****Pure Isabelle*****' +(cd Pure; make $NO > make$$.log) +tail Pure/make$$.log +gzip Pure/make$$.log + +echo +echo +echo '*****First-Order Logic (FOL)*****' +(cd FOL; make $NO $TEST > make$$.log) +tail FOL/make$$.log +gzip FOL/make$$.log +#cannot delete FOL yet... it is needed for ZF, CCL and LCF! + +echo +echo +echo '*****Set theory (ZF)*****' +(cd ZF; make $NO $TEST > make$$.log) +tail ZF/make$$.log +gzip ZF/make$$.log +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/ZF +esac + +echo +echo +echo '*****Classical Computational Logic (CCL)*****' +(cd CCL; make $NO $TEST > make$$.log) +tail CCL/make$$.log +gzip CCL/make$$.log +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/CCL +esac + +echo +echo +echo '*****Domain Theory (LCF)*****' +(cd LCF; make $NO $TEST > make$$.log) +tail LCF/make$$.log +gzip LCF/make$$.log +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/FOL $ISABELLEBIN/LCF +esac + +echo +echo +echo '*****Constructive Type Theory (CTT)*****' +(cd CTT; make $NO $TEST > make$$.log) +tail CTT/make$$.log +gzip CTT/make$$.log +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/CTT +esac + +echo +echo +echo '*****Classical Sequent Calculus (LK)*****' +(cd LK; make $NO $TEST > make$$.log) +tail LK/make$$.log +gzip LK/make$$.log +#cannot delete LK yet... it is needed for Modal! + +echo +echo +echo '*****Modal logic (Modal)*****' +(cd Modal; make $NO $TEST > make$$.log) +tail Modal/make$$.log +gzip Modal/make$$.log +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/LK $ISABELLEBIN/Modal +esac + +echo +echo +echo '*****Higher-Order Logic (HOL)*****' +(cd HOL; make $NO $TEST > make$$.log) +tail HOL/make$$.log +gzip HOL/make$$.log +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/HOL +esac + +echo +echo +echo '*****The Lambda-Cube (Cube)*****' +(cd Cube; make $NO $TEST > make$$.log) +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/Cube +esac +tail Cube/make$$.log +gzip Cube/make$$.log + +echo +echo +echo '*****First-Order Logic with Proof Terms (FOLP)*****' +(cd FOLP; make $NO $TEST > make$$.log) +case $CLEAN.$EXEC in + on.on) rm $ISABELLEBIN/FOLP +esac +tail FOLP/make$$.log +gzip FOLP/make$$.log + +case $TEST.$EXEC in + test.on) echo + echo '***** Now check the dates on the "test" files *****' + ls -lrt FOL/test ZF/test CCL/test LCF/test CTT/test\ + LK/test Modal/test HOL/test Cube/test FOLP/test +esac +echo Finished at `date` diff -r 000000000000 -r a5a9c433f639 src/Tools/rm-logfiles --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/rm-logfiles Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,7 @@ +#! /bin/sh +#rm-logfiles: remove useless files from subdirectories +rm log */make*.log */make*.log.gz */make*.log.z +rm */test +rm */.*.thy.ML +rm */ex/.*.thy.ML +rm HOL/Subst/.*.thy.ML diff -r 000000000000 -r a5a9c433f639 src/Tools/teeinput --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/teeinput Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,4 @@ +#! /bin/sh +# teeinput -- start a program and log all inputs to a file +# environment variable $LISTEN specifies the file name +tee -a -i $LISTEN | $* \ No newline at end of file diff -r 000000000000 -r a5a9c433f639 src/Tools/xlisten --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/xlisten Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,10 @@ +#! /bin/sh +# xlisten -- start a program in one window and create a listener window +# environment variable $LISTEN specifies the file name + +#create the file! +date > $LISTEN + +xterm -geo 80x10+0+0 -T Listener -n Listener -e tail -f $LISTEN & +sleep 2 +xterm -geo 80x45+0-0 -e teeinput $* & diff -r 000000000000 -r a5a9c433f639 src/ZF/Arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Arith.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,356 @@ +(* Title: ZF/arith.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For arith.thy. Arithmetic operators and their definitions + +Proofs about elementary arithmetic: addition, multiplication, etc. + +Could prove def_rec_0, def_rec_succ... +*) + +open Arith; + +(*"Difference" is subtraction of natural numbers. + There are no negative numbers; we have + m #- n = 0 iff m<=n and m #- n = succ(k) iff m>n. + Also, rec(m, 0, %z w.z) is pred(m). +*) + +(** rec -- better than nat_rec; the succ case has no type requirement! **) + +val rec_trans = rec_def RS def_transrec RS trans; + +goal Arith.thy "rec(0,a,b) = a"; +by (rtac rec_trans 1); +by (rtac nat_case_0 1); +val rec_0 = result(); + +goal Arith.thy "rec(succ(m),a,b) = b(m, rec(m,a,b))"; +val rec_ss = ZF_ss + addcongs (mk_typed_congs Arith.thy [("b", "[i,i]=>i")]) + addrews [nat_case_succ, nat_succI]; +by (rtac rec_trans 1); +by (SIMP_TAC rec_ss 1); +val rec_succ = result(); + +val major::prems = goal Arith.thy + "[| n: nat; \ +\ a: C(0); \ +\ !!m z. [| m: nat; z: C(m) |] ==> b(m,z): C(succ(m)) \ +\ |] ==> rec(n,a,b) : C(n)"; +by (rtac (major RS nat_induct) 1); +by (ALLGOALS + (ASM_SIMP_TAC (ZF_ss addrews (prems@[rec_0,rec_succ])))); +val rec_type = result(); + +val prems = goalw Arith.thy [rec_def] + "[| n=n'; a=a'; !!m z. b(m,z)=b'(m,z) \ +\ |] ==> rec(n,a,b)=rec(n',a',b')"; +by (SIMP_TAC (ZF_ss addcongs [transrec_cong,nat_case_cong] + addrews (prems RL [sym])) 1); +val rec_cong = result(); + +val nat_typechecks = [rec_type,nat_0I,nat_1I,nat_succI,Ord_nat]; + +val nat_ss = ZF_ss addcongs [nat_case_cong,rec_cong] + addrews ([rec_0,rec_succ] @ nat_typechecks); + + +(** Addition **) + +val add_type = prove_goalw Arith.thy [add_def] + "[| m:nat; n:nat |] ==> m #+ n : nat" + (fn prems=> [ (typechk_tac (prems@nat_typechecks@ZF_typechecks)) ]); + +val add_0 = prove_goalw Arith.thy [add_def] + "0 #+ n = n" + (fn _ => [ (rtac rec_0 1) ]); + +val add_succ = prove_goalw Arith.thy [add_def] + "succ(m) #+ n = succ(m #+ n)" + (fn _=> [ (rtac rec_succ 1) ]); + +(** Multiplication **) + +val mult_type = prove_goalw Arith.thy [mult_def] + "[| m:nat; n:nat |] ==> m #* n : nat" + (fn prems=> + [ (typechk_tac (prems@[add_type]@nat_typechecks@ZF_typechecks)) ]); + +val mult_0 = prove_goalw Arith.thy [mult_def] + "0 #* n = 0" + (fn _ => [ (rtac rec_0 1) ]); + +val mult_succ = prove_goalw Arith.thy [mult_def] + "succ(m) #* n = n #+ (m #* n)" + (fn _ => [ (rtac rec_succ 1) ]); + +(** Difference **) + +val diff_type = prove_goalw Arith.thy [diff_def] + "[| m:nat; n:nat |] ==> m #- n : nat" + (fn prems=> [ (typechk_tac (prems@nat_typechecks@ZF_typechecks)) ]); + +val diff_0 = prove_goalw Arith.thy [diff_def] + "m #- 0 = m" + (fn _ => [ (rtac rec_0 1) ]); + +val diff_0_eq_0 = prove_goalw Arith.thy [diff_def] + "n:nat ==> 0 #- n = 0" + (fn [prem]=> + [ (rtac (prem RS nat_induct) 1), + (ALLGOALS (ASM_SIMP_TAC nat_ss)) ]); + +(*Must simplify BEFORE the induction!! (Else we get a critical pair) + succ(m) #- succ(n) rewrites to pred(succ(m) #- n) *) +val diff_succ_succ = prove_goalw Arith.thy [diff_def] + "[| m:nat; n:nat |] ==> succ(m) #- succ(n) = m #- n" + (fn prems=> + [ (ASM_SIMP_TAC (nat_ss addrews prems) 1), + (nat_ind_tac "n" prems 1), + (ALLGOALS (ASM_SIMP_TAC (nat_ss addrews prems))) ]); + +val prems = goal Arith.thy + "[| m:nat; n:nat |] ==> m #- n : succ(m)"; +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +by (etac succE 3); +by (ALLGOALS + (ASM_SIMP_TAC + (nat_ss addrews (prems@[diff_0,diff_0_eq_0,diff_succ_succ])))); +val diff_leq = result(); + +(*** Simplification over add, mult, diff ***) + +val arith_typechecks = [add_type, mult_type, diff_type]; +val arith_rews = [add_0, add_succ, + mult_0, mult_succ, + diff_0, diff_0_eq_0, diff_succ_succ]; + +val arith_congs = mk_congs Arith.thy ["op #+", "op #-", "op #*"]; + +val arith_ss = nat_ss addcongs arith_congs + addrews (arith_rews@arith_typechecks); + +(*** Addition ***) + +(*Associative law for addition*) +val add_assoc = prove_goal Arith.thy + "m:nat ==> (m #+ n) #+ k = m #+ (n #+ k)" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews prems))) ]); + +(*The following two lemmas are used for add_commute and sometimes + elsewhere, since they are safe for rewriting.*) +val add_0_right = prove_goal Arith.thy + "m:nat ==> m #+ 0 = m" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews prems))) ]); + +val add_succ_right = prove_goal Arith.thy + "m:nat ==> m #+ succ(n) = succ(m #+ n)" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews prems))) ]); + +(*Commutative law for addition*) +val add_commute = prove_goal Arith.thy + "[| m:nat; n:nat |] ==> m #+ n = n #+ m" + (fn prems=> + [ (nat_ind_tac "n" prems 1), + (ALLGOALS + (ASM_SIMP_TAC + (arith_ss addrews (prems@[add_0_right, add_succ_right])))) ]); + +(*Cancellation law on the left*) +val [knat,eqn] = goal Arith.thy + "[| k:nat; k #+ m = k #+ n |] ==> m=n"; +by (rtac (eqn RS rev_mp) 1); +by (nat_ind_tac "k" [knat] 1); +by (ALLGOALS (SIMP_TAC arith_ss)); +by (fast_tac ZF_cs 1); +val add_left_cancel = result(); + +(*** Multiplication ***) + +(*right annihilation in product*) +val mult_0_right = prove_goal Arith.thy + "m:nat ==> m #* 0 = 0" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews prems))) ]); + +(*right successor law for multiplication*) +val mult_succ_right = prove_goal Arith.thy + "[| m:nat; n:nat |] ==> m #* succ(n) = m #+ (m #* n)" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews ([add_assoc RS sym]@prems)))), + (*The final goal requires the commutative law for addition*) + (REPEAT (ares_tac (prems@[refl,add_commute]@ZF_congs@arith_congs) 1)) ]); + +(*Commutative law for multiplication*) +val mult_commute = prove_goal Arith.thy + "[| m:nat; n:nat |] ==> m #* n = n #* m" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC + (arith_ss addrews (prems@[mult_0_right, mult_succ_right])))) ]); + +(*addition distributes over multiplication*) +val add_mult_distrib = prove_goal Arith.thy + "[| m:nat; k:nat |] ==> (m #+ n) #* k = (m #* k) #+ (n #* k)" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews ([add_assoc RS sym]@prems)))) ]); + +(*Distributive law on the left; requires an extra typing premise*) +val add_mult_distrib_left = prove_goal Arith.thy + "[| m:nat; n:nat; k:nat |] ==> k #* (m #+ n) = (k #* m) #+ (k #* n)" + (fn prems=> + let val mult_commute' = read_instantiate [("m","k")] mult_commute + val ss = arith_ss addrews ([mult_commute',add_mult_distrib]@prems) + in [ (SIMP_TAC ss 1) ] + end); + +(*Associative law for multiplication*) +val mult_assoc = prove_goal Arith.thy + "[| m:nat; n:nat; k:nat |] ==> (m #* n) #* k = m #* (n #* k)" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews (prems@[add_mult_distrib])))) ]); + + +(*** Difference ***) + +val diff_self_eq_0 = prove_goal Arith.thy + "m:nat ==> m #- m = 0" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews prems))) ]); + +(*Addition is the inverse of subtraction: if n<=m then n+(m-n) = m. *) +val notless::prems = goal Arith.thy + "[| ~m:n; m:nat; n:nat |] ==> n #+ (m#-n) = m"; +by (rtac (notless RS rev_mp) 1); +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +by (ALLGOALS (ASM_SIMP_TAC + (arith_ss addrews (prems@[succ_mem_succ_iff, Ord_0_mem_succ, + naturals_are_ordinals])))); +val add_diff_inverse = result(); + + +(*Subtraction is the inverse of addition. *) +val [mnat,nnat] = goal Arith.thy + "[| m:nat; n:nat |] ==> (n#+m) #-n = m"; +by (rtac (nnat RS nat_induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews [mnat]))); +val diff_add_inverse = result(); + +val [mnat,nnat] = goal Arith.thy + "[| m:nat; n:nat |] ==> n #- (n#+m) = 0"; +by (rtac (nnat RS nat_induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews [mnat]))); +val diff_add_0 = result(); + + +(*** Remainder ***) + +(*In ordinary notation: if 0 m #- n : m"; +by (cut_facts_tac prems 1); +by (etac rev_mp 1); +by (etac rev_mp 1); +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +by (ALLGOALS (ASM_SIMP_TAC + (nat_ss addrews (prems@[diff_leq,diff_succ_succ])))); +val div_termination = result(); + +val div_rls = + [Ord_transrec_type, apply_type, div_termination, if_type] @ + nat_typechecks; + +(*Type checking depends upon termination!*) +val prems = goalw Arith.thy [mod_def] + "[| 0:n; m:nat; n:nat |] ==> m mod n : nat"; +by (REPEAT (ares_tac (prems @ div_rls) 1 ORELSE etac Ord_trans 1)); +val mod_type = result(); + +val div_ss = ZF_ss addrews [naturals_are_ordinals,div_termination]; + +val prems = goal Arith.thy "[| 0:n; m:n; m:nat; n:nat |] ==> m mod n = m"; +by (rtac (mod_def RS def_transrec RS trans) 1); +by (SIMP_TAC (div_ss addrews prems) 1); +val mod_less = result(); + +val prems = goal Arith.thy + "[| 0:n; ~m:n; m:nat; n:nat |] ==> m mod n = (m#-n) mod n"; +by (rtac (mod_def RS def_transrec RS trans) 1); +by (SIMP_TAC (div_ss addrews prems) 1); +val mod_geq = result(); + +(*** Quotient ***) + +(*Type checking depends upon termination!*) +val prems = goalw Arith.thy [div_def] + "[| 0:n; m:nat; n:nat |] ==> m div n : nat"; +by (REPEAT (ares_tac (prems @ div_rls) 1 ORELSE etac Ord_trans 1)); +val div_type = result(); + +val prems = goal Arith.thy + "[| 0:n; m:n; m:nat; n:nat |] ==> m div n = 0"; +by (rtac (div_def RS def_transrec RS trans) 1); +by (SIMP_TAC (div_ss addrews prems) 1); +val div_less = result(); + +val prems = goal Arith.thy + "[| 0:n; ~m:n; m:nat; n:nat |] ==> m div n = succ((m#-n) div n)"; +by (rtac (div_def RS def_transrec RS trans) 1); +by (SIMP_TAC (div_ss addrews prems) 1); +val div_geq = result(); + +(*Main Result.*) +val prems = goal Arith.thy + "[| 0:n; m:nat; n:nat |] ==> (m div n)#*n #+ m mod n = m"; +by (res_inst_tac [("i","m")] complete_induct 1); +by (resolve_tac prems 1); +by (res_inst_tac [("Q","x:n")] (excluded_middle RS disjE) 1); +by (ALLGOALS + (ASM_SIMP_TAC + (arith_ss addrews ([mod_type,div_type] @ prems @ + [mod_less,mod_geq, div_less, div_geq, + add_assoc, add_diff_inverse, div_termination])))); +val mod_div_equality = result(); + + +(**** Additional theorems about "less than" ****) + +val [mnat,nnat] = goal Arith.thy + "[| m:nat; n:nat |] ==> ~ (m #+ n) : n"; +by (rtac (mnat RS nat_induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews [mem_not_refl]))); +by (rtac notI 1); +by (etac notE 1); +by (etac (succI1 RS Ord_trans) 1); +by (rtac (nnat RS naturals_are_ordinals) 1); +val add_not_less_self = result(); + +val [mnat,nnat] = goal Arith.thy + "[| m:nat; n:nat |] ==> m : succ(m #+ n)"; +by (rtac (mnat RS nat_induct) 1); +(*May not simplify even with ZF_ss because it would expand m:succ(...) *) +by (rtac (add_0 RS ssubst) 1); +by (rtac (add_succ RS ssubst) 2); +by (REPEAT (ares_tac [nnat, Ord_0_mem_succ, succ_mem_succI, + naturals_are_ordinals, nat_succI, add_type] 1)); +val add_less_succ_self = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/Arith.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Arith.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,26 @@ +(* Title: ZF/arith.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Arithmetic operators and their definitions +*) + +Arith = Epsilon + +consts + rec :: "[i, i, [i,i]=>i]=>i" + "#*" :: "[i,i]=>i" (infixl 70) + div :: "[i,i]=>i" (infixl 70) + mod :: "[i,i]=>i" (infixl 70) + "#+" :: "[i,i]=>i" (infixl 65) + "#-" :: "[i,i]=>i" (infixl 65) + +rules + rec_def "rec(k,a,b) == transrec(k, %n f. nat_case(n, a, %m. b(m, f`m)))" + + add_def "m#+n == rec(m, n, %u v.succ(v))" + diff_def "m#-n == rec(n, m, %u v. rec(v, 0, %x y.x))" + mult_def "m#*n == rec(m, 0, %u v. n #+ v)" + mod_def "m mod n == transrec(m, %j f. if(j:n, j, f`(j#-n)))" + div_def "m div n == transrec(m, %j f. if(j:n, 0, succ(f`(j#-n))))" +end diff -r 000000000000 -r a5a9c433f639 src/ZF/Bool.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Bool.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,99 @@ +(* Title: ZF/bool + ID: $Id$ + Author: Martin D Coen, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For ZF/bool.thy. Booleans in Zermelo-Fraenkel Set Theory +*) + +open Bool; + +val bool_defs = [bool_def,one_def,cond_def]; + +(* Introduction rules *) + +goalw Bool.thy bool_defs "1 : bool"; +by (rtac (consI1 RS consI2) 1); +val bool_1I = result(); + +goalw Bool.thy bool_defs "0 : bool"; +by (rtac consI1 1); +val bool_0I = result(); + +goalw Bool.thy bool_defs "~ 1=0"; +by (rtac succ_not_0 1); +val one_not_0 = result(); + +(** 1=0 ==> R **) +val one_neq_0 = one_not_0 RS notE; + +val prems = goalw Bool.thy bool_defs "[| c: bool; P(1); P(0) |] ==> P(c)"; +by (cut_facts_tac prems 1); +by (fast_tac ZF_cs 1); +val boolE = result(); + +(** cond **) + +(*1 means true*) +goalw Bool.thy bool_defs "cond(1,c,d) = c"; +by (rtac (refl RS if_P) 1); +val cond_1 = result(); + +(*0 means false*) +goalw Bool.thy bool_defs "cond(0,c,d) = d"; +by (rtac (succ_not_0 RS not_sym RS if_not_P) 1); +val cond_0 = result(); + +val major::prems = goal Bool.thy + "[| b: bool; c: A(1); d: A(0) |] ==> cond(b,c,d): A(b)"; +by (rtac (major RS boolE) 1); +by (rtac (cond_0 RS ssubst) 2); +by (resolve_tac prems 2); +by (rtac (cond_1 RS ssubst) 1); +by (resolve_tac prems 1); +val cond_type = result(); + +val [cond_cong] = mk_congs Bool.thy ["cond"]; +val bool_congs = mk_congs Bool.thy ["cond","not","op and","op or","op xor"]; + +val [rew] = goal Bool.thy "[| !!b. j(b)==cond(b,c,d) |] ==> j(1) = c"; +by (rewtac rew); +by (rtac cond_1 1); +val def_cond_1 = result(); + +val [rew] = goal Bool.thy "[| !!b. j(b)==cond(b,c,d) |] ==> j(0) = d"; +by (rewtac rew); +by (rtac cond_0 1); +val def_cond_0 = result(); + +fun conds def = [standard (def RS def_cond_1), standard (def RS def_cond_0)]; + +val [not_1,not_0] = conds not_def; + +val [and_1,and_0] = conds and_def; + +val [or_1,or_0] = conds or_def; + +val [xor_1,xor_0] = conds xor_def; + +val not_type = prove_goalw Bool.thy [not_def] + "a:bool ==> not(a) : bool" + (fn prems=> [ (typechk_tac (prems@[bool_1I, bool_0I, cond_type])) ]); + +val and_type = prove_goalw Bool.thy [and_def] + "[| a:bool; b:bool |] ==> a and b : bool" + (fn prems=> [ (typechk_tac (prems@[bool_1I, bool_0I, cond_type])) ]); + +val or_type = prove_goalw Bool.thy [or_def] + "[| a:bool; b:bool |] ==> a or b : bool" + (fn prems=> [ (typechk_tac (prems@[bool_1I, bool_0I, cond_type])) ]); + +val xor_type = prove_goalw Bool.thy [xor_def] + "[| a:bool; b:bool |] ==> a xor b : bool" + (fn prems=> [ (typechk_tac(prems@[bool_1I, bool_0I, cond_type, not_type])) ]); + +val bool_typechecks = [bool_1I, bool_0I, cond_type, not_type, and_type, + or_type, xor_type] + +val bool_rews = [cond_1,cond_0,not_1,not_0,and_1,and_0,or_1,or_0,xor_1,xor_0]; + diff -r 000000000000 -r a5a9c433f639 src/ZF/Bool.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Bool.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,27 @@ +(* Title: ZF/bool.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Booleans in Zermelo-Fraenkel Set Theory +*) + +Bool = ZF + +consts + "1" :: "i" ("1") + bool :: "i" + cond :: "[i,i,i]=>i" + not :: "i=>i" + and :: "[i,i]=>i" (infixl 70) + or :: "[i,i]=>i" (infixl 65) + xor :: "[i,i]=>i" (infixl 65) + +rules + one_def "1 == succ(0)" + bool_def "bool == {0,1}" + cond_def "cond(b,c,d) == if(b=1,c,d)" + not_def "not(b) == cond(b,0,1)" + and_def "a and b == cond(a,b,0)" + or_def "a or b == cond(a,1,b)" + xor_def "a xor b == cond(a,not(b),b)" +end diff -r 000000000000 -r a5a9c433f639 src/ZF/Datatype.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Datatype.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,66 @@ +(* Title: ZF/datatype.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +(Co-) Datatype Definitions for Zermelo-Fraenkel Set Theory + +*) + + +(*Datatype definitions use least fixedpoints, standard products and sums*) +functor Datatype_Fun (Const: CONSTRUCTOR) + : sig include CONSTRUCTOR_RESULT INTR_ELIM INDRULE end = +struct +structure Constructor = Constructor_Fun (structure Const=Const and + Pr=Standard_Prod and Su=Standard_Sum); +open Const Constructor; + +structure Inductive = Inductive_Fun + (val thy = con_thy; + val rec_doms = (map #1 rec_specs) ~~ (map #2 rec_specs); + val sintrs = sintrs; + val monos = monos; + val con_defs = con_defs; + val type_intrs = type_intrs; + val type_elims = type_elims); + +open Inductive +end; + + +(*Co-datatype definitions use greatest fixedpoints, Quine products and sums*) +functor Co_Datatype_Fun (Const: CONSTRUCTOR) + : sig include CONSTRUCTOR_RESULT INTR_ELIM CO_INDRULE end = +struct +structure Constructor = Constructor_Fun (structure Const=Const and + Pr=Quine_Prod and Su=Quine_Sum); +open Const Constructor; + +structure Co_Inductive = Co_Inductive_Fun + (val thy = con_thy; + val rec_doms = (map #1 rec_specs) ~~ (map #2 rec_specs); + val sintrs = sintrs; + val monos = monos; + val con_defs = con_defs; + val type_intrs = type_intrs; + val type_elims = type_elims); + +open Co_Inductive +end; + + + +(*For most datatypes involving univ*) +val data_typechecks = + [SigmaI, InlI, InrI, + Pair_in_univ, Inl_in_univ, Inr_in_univ, + zero_in_univ, A_into_univ, nat_into_univ, sum_univ RS subsetD]; + + +(*For most co-datatypes involving quniv*) +val co_data_typechecks = + [QSigmaI, QInlI, QInrI, + QPair_in_quniv, QInl_in_quniv, QInr_in_quniv, + zero_in_quniv, A_into_quniv, nat_into_quniv, qsum_quniv RS subsetD]; + diff -r 000000000000 -r a5a9c433f639 src/ZF/Epsilon.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Epsilon.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,325 @@ +(* Title: ZF/epsilon.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For epsilon.thy. Epsilon induction and recursion +*) + +open Epsilon; + +(*** Basic closure properties ***) + +goalw Epsilon.thy [eclose_def] "A <= eclose(A)"; +by (rtac (nat_rec_0 RS equalityD2 RS subset_trans) 1); +br (nat_0I RS UN_upper) 1; +val arg_subset_eclose = result(); + +val arg_into_eclose = arg_subset_eclose RS subsetD; + +goalw Epsilon.thy [eclose_def,Transset_def] "Transset(eclose(A))"; +by (rtac (subsetI RS ballI) 1); +by (etac UN_E 1); +by (rtac (nat_succI RS UN_I) 1); +by (assume_tac 1); +by (etac (nat_rec_succ RS ssubst) 1); +by (etac UnionI 1); +by (assume_tac 1); +val Transset_eclose = result(); + +(* x : eclose(A) ==> x <= eclose(A) *) +val eclose_subset = + standard (rewrite_rule [Transset_def] Transset_eclose RS bspec); + +(* [| A : eclose(B); c : A |] ==> c : eclose(B) *) +val ecloseD = standard (eclose_subset RS subsetD); + +val arg_in_eclose_sing = arg_subset_eclose RS singleton_subsetD; +val arg_into_eclose_sing = arg_in_eclose_sing RS ecloseD; + +(* This is epsilon-induction for eclose(A); see also eclose_induct_down... + [| a: eclose(A); !!x. [| x: eclose(A); ALL y:x. P(y) |] ==> P(x) + |] ==> P(a) +*) +val eclose_induct = standard (Transset_eclose RSN (2, Transset_induct)); + +(*Epsilon induction*) +val prems = goal Epsilon.thy + "[| !!x. ALL y:x. P(y) ==> P(x) |] ==> P(a)"; +by (rtac (arg_in_eclose_sing RS eclose_induct) 1); +by (eresolve_tac prems 1); +val eps_induct = result(); + +(*Perform epsilon-induction on i. *) +fun eps_ind_tac a = + EVERY' [res_inst_tac [("a",a)] eps_induct, + rename_last_tac a ["1"]]; + + +(*** Leastness of eclose ***) + +(** eclose(A) is the least transitive set including A as a subset. **) + +goalw Epsilon.thy [Transset_def] + "!!X A n. [| Transset(X); A<=X; n: nat |] ==> \ +\ nat_rec(n, A, %m r. Union(r)) <= X"; +by (etac nat_induct 1); +by (ASM_SIMP_TAC (ZF_ss addrews [nat_rec_0]) 1); +by (ASM_SIMP_TAC (ZF_ss addrews [nat_rec_succ]) 1); +by (fast_tac ZF_cs 1); +val eclose_least_lemma = result(); + +goalw Epsilon.thy [eclose_def] + "!!X A. [| Transset(X); A<=X |] ==> eclose(A) <= X"; +br (eclose_least_lemma RS UN_least) 1; +by (REPEAT (assume_tac 1)); +val eclose_least = result(); + +(*COMPLETELY DIFFERENT induction principle from eclose_induct!!*) +val [major,base,step] = goal Epsilon.thy + "[| a: eclose(b); \ +\ !!y. [| y: b |] ==> P(y); \ +\ !!y z. [| y: eclose(b); P(y); z: y |] ==> P(z) \ +\ |] ==> P(a)"; +by (rtac (major RSN (3, eclose_least RS subsetD RS CollectD2)) 1); +by (rtac (CollectI RS subsetI) 2); +by (etac (arg_subset_eclose RS subsetD) 2); +by (etac base 2); +by (rewtac Transset_def); +by (fast_tac (ZF_cs addEs [step,ecloseD]) 1); +val eclose_induct_down = result(); + +goal Epsilon.thy "!!X. Transset(X) ==> eclose(X) = X"; +be ([eclose_least, arg_subset_eclose] MRS equalityI) 1; +br subset_refl 1; +val Transset_eclose_eq_arg = result(); + + +(*** Epsilon recursion ***) + +(*Unused...*) +goal Epsilon.thy "!!A B C. [| A: eclose(B); B: eclose(C) |] ==> A: eclose(C)"; +by (rtac ([Transset_eclose, eclose_subset] MRS eclose_least RS subsetD) 1); +by (REPEAT (assume_tac 1)); +val mem_eclose_trans = result(); + +(*Variant of the previous lemma in a useable form for the sequel*) +goal Epsilon.thy + "!!A B C. [| A: eclose({B}); B: eclose({C}) |] ==> A: eclose({C})"; +by (rtac ([Transset_eclose, singleton_subsetI] MRS eclose_least RS subsetD) 1); +by (REPEAT (assume_tac 1)); +val mem_eclose_sing_trans = result(); + +goalw Epsilon.thy [Transset_def] + "!!i j. [| Transset(i); j:i |] ==> Memrel(i)-``{j} = j"; +by (fast_tac (eq_cs addSIs [MemrelI] addSEs [MemrelE]) 1); +val under_Memrel = result(); + +(* j : eclose(A) ==> Memrel(eclose(A)) -`` j = j *) +val under_Memrel_eclose = Transset_eclose RS under_Memrel; + +val wfrec_ssubst = standard (wf_Memrel RS wfrec RS ssubst); + +val [kmemj,jmemi] = goal Epsilon.thy + "[| k:eclose({j}); j:eclose({i}) |] ==> \ +\ wfrec(Memrel(eclose({i})), k, H) = wfrec(Memrel(eclose({j})), k, H)"; +by (rtac (kmemj RS eclose_induct) 1); +by (rtac wfrec_ssubst 1); +by (rtac wfrec_ssubst 1); +by (ASM_SIMP_TAC (wf_ss addrews [under_Memrel_eclose, + jmemi RSN (2,mem_eclose_sing_trans)]) 1); +val wfrec_eclose_eq = result(); + +val [prem] = goal Epsilon.thy + "k: i ==> wfrec(Memrel(eclose({i})),k,H) = wfrec(Memrel(eclose({k})),k,H)"; +by (rtac (arg_in_eclose_sing RS wfrec_eclose_eq) 1); +by (rtac (prem RS arg_into_eclose_sing) 1); +val wfrec_eclose_eq2 = result(); + +goalw Epsilon.thy [transrec_def] + "transrec(a,H) = H(a, lam x:a. transrec(x,H))"; +by (rtac wfrec_ssubst 1); +by (SIMP_TAC (wf_ss addrews [wfrec_eclose_eq2, + arg_in_eclose_sing, under_Memrel_eclose]) 1); +val transrec = result(); + +(*Avoids explosions in proofs; resolve it with a meta-level definition.*) +val rew::prems = goal Epsilon.thy + "[| !!x. f(x)==transrec(x,H) |] ==> f(a) = H(a, lam x:a. f(x))"; +by (rewtac rew); +by (REPEAT (resolve_tac (prems@[transrec]) 1)); +val def_transrec = result(); + +val prems = goal Epsilon.thy + "[| !!x u. [| x:eclose({a}); u: Pi(x,B) |] ==> H(x,u) : B(x) \ +\ |] ==> transrec(a,H) : B(a)"; +by (res_inst_tac [("i", "a")] (arg_in_eclose_sing RS eclose_induct) 1); +by (rtac (transrec RS ssubst) 1); +by (REPEAT (ares_tac (prems @ [lam_type]) 1 ORELSE etac bspec 1)); +val transrec_type = result(); + +goal Epsilon.thy "!!i. Ord(i) ==> eclose({i}) <= succ(i)"; +by (etac (Ord_is_Transset RS Transset_succ RS eclose_least) 1); +by (rtac (succI1 RS singleton_subsetI) 1); +val eclose_sing_Ord = result(); + +val prems = goal Epsilon.thy + "[| j: i; Ord(i); \ +\ !!x u. [| x: i; u: Pi(x,B) |] ==> H(x,u) : B(x) \ +\ |] ==> transrec(j,H) : B(j)"; +by (rtac transrec_type 1); +by (resolve_tac prems 1); +by (rtac (Ord_in_Ord RS eclose_sing_Ord RS subsetD RS succE) 1); +by (DEPTH_SOLVE (ares_tac prems 1 ORELSE eresolve_tac [ssubst,Ord_trans] 1)); +val Ord_transrec_type = result(); + +(*Congruence*) +val prems = goalw Epsilon.thy [transrec_def,Memrel_def] + "[| a=a'; !!x u. H(x,u)=H'(x,u) |] ==> transrec(a,H)=transrec(a',H')"; +val transrec_ss = + ZF_ss addcongs ([wfrec_cong] @ mk_congs Epsilon.thy ["eclose"]) + addrews (prems RL [sym]); +by (SIMP_TAC transrec_ss 1); +val transrec_cong = result(); + +(*** Rank ***) + +val ord_ss = ZF_ss addcongs (mk_congs Ord.thy ["Ord"]); + +(*NOT SUITABLE FOR REWRITING -- RECURSIVE!*) +goal Epsilon.thy "rank(a) = (UN y:a. succ(rank(y)))"; +by (rtac (rank_def RS def_transrec RS ssubst) 1); +by (SIMP_TAC ZF_ss 1); +val rank = result(); + +goal Epsilon.thy "Ord(rank(a))"; +by (eps_ind_tac "a" 1); +by (rtac (rank RS ssubst) 1); +by (rtac (Ord_succ RS Ord_UN) 1); +by (etac bspec 1); +by (assume_tac 1); +val Ord_rank = result(); + +val [major] = goal Epsilon.thy "Ord(i) ==> rank(i) = i"; +by (rtac (major RS trans_induct) 1); +by (rtac (rank RS ssubst) 1); +by (ASM_SIMP_TAC (ord_ss addrews [Ord_equality]) 1); +val rank_of_Ord = result(); + +val [prem] = goal Epsilon.thy "a:b ==> rank(a) : rank(b)"; +by (res_inst_tac [("a1","b")] (rank RS ssubst) 1); +by (rtac (prem RS UN_I) 1); +by (rtac succI1 1); +val rank_lt = result(); + +val [major] = goal Epsilon.thy "a: eclose(b) ==> rank(a) : rank(b)"; +by (rtac (major RS eclose_induct_down) 1); +by (etac rank_lt 1); +by (etac (rank_lt RS Ord_trans) 1); +by (assume_tac 1); +by (rtac Ord_rank 1); +val eclose_rank_lt = result(); + +goal Epsilon.thy "!!a b. a<=b ==> rank(a) <= rank(b)"; +by (rtac (rank RS ssubst) 1); +by (rtac (rank RS ssubst) 1); +by (etac UN_mono 1); +by (rtac subset_refl 1); +val rank_mono = result(); + +goal Epsilon.thy "rank(Pow(a)) = succ(rank(a))"; +by (rtac (rank RS trans) 1); +by (rtac equalityI 1); +by (fast_tac ZF_cs 2); +by (rtac UN_least 1); +by (etac (PowD RS rank_mono RS Ord_succ_mono) 1); +by (rtac Ord_rank 1); +by (rtac Ord_rank 1); +val rank_Pow = result(); + +goal Epsilon.thy "rank(0) = 0"; +by (rtac (rank RS trans) 1); +by (fast_tac (ZF_cs addSIs [equalityI]) 1); +val rank_0 = result(); + +goal Epsilon.thy "rank(succ(x)) = succ(rank(x))"; +by (rtac (rank RS trans) 1); +br ([UN_least, succI1 RS UN_upper] MRS equalityI) 1; +be succE 1; +by (fast_tac ZF_cs 1); +by (REPEAT (ares_tac [Ord_succ_mono,Ord_rank,OrdmemD,rank_lt] 1)); +val rank_succ = result(); + +goal Epsilon.thy "rank(Union(A)) = (UN x:A. rank(x))"; +by (rtac equalityI 1); +by (rtac (rank_mono RS UN_least) 2); +by (etac Union_upper 2); +by (rtac (rank RS ssubst) 1); +by (rtac UN_least 1); +by (etac UnionE 1); +by (rtac subset_trans 1); +by (etac (RepFunI RS Union_upper) 2); +by (etac (rank_lt RS Ord_succ_subsetI) 1); +by (rtac Ord_rank 1); +val rank_Union = result(); + +goal Epsilon.thy "rank(eclose(a)) = rank(a)"; +by (rtac equalityI 1); +by (rtac (arg_subset_eclose RS rank_mono) 2); +by (res_inst_tac [("a1","eclose(a)")] (rank RS ssubst) 1); +by (rtac UN_least 1); +by (etac ([eclose_rank_lt, Ord_rank] MRS Ord_succ_subsetI) 1); +val rank_eclose = result(); + +(* [| i: j; j: rank(a) |] ==> i: rank(a) *) +val rank_trans = Ord_rank RSN (3, Ord_trans); + +goalw Epsilon.thy [Pair_def] "rank(a) : rank()"; +by (rtac (consI1 RS rank_lt RS Ord_trans) 1); +by (rtac (consI1 RS consI2 RS rank_lt) 1); +by (rtac Ord_rank 1); +val rank_pair1 = result(); + +goalw Epsilon.thy [Pair_def] "rank(b) : rank()"; +by (rtac (consI1 RS consI2 RS rank_lt RS Ord_trans) 1); +by (rtac (consI1 RS consI2 RS rank_lt) 1); +by (rtac Ord_rank 1); +val rank_pair2 = result(); + +goalw (merge_theories(Epsilon.thy,Sum.thy)) [Inl_def] "rank(a) : rank(Inl(a))"; +by (rtac rank_pair2 1); +val rank_Inl = result(); + +goalw (merge_theories(Epsilon.thy,Sum.thy)) [Inr_def] "rank(a) : rank(Inr(a))"; +by (rtac rank_pair2 1); +val rank_Inr = result(); + +val [major] = goal Epsilon.thy "i: rank(a) ==> (EX x:a. i<=rank(x))"; +by (resolve_tac ([major] RL [rank RS subst] RL [UN_E]) 1); +by (rtac bexI 1); +by (etac member_succD 1); +by (rtac Ord_rank 1); +by (assume_tac 1); +val rank_implies_mem = result(); + + +(*** Corollaries of leastness ***) + +goal Epsilon.thy "!!A B. A:B ==> eclose(A)<=eclose(B)"; +by (rtac (Transset_eclose RS eclose_least) 1); +by (etac (arg_into_eclose RS eclose_subset) 1); +val mem_eclose_subset = result(); + +goal Epsilon.thy "!!A B. A<=B ==> eclose(A) <= eclose(B)"; +by (rtac (Transset_eclose RS eclose_least) 1); +by (etac subset_trans 1); +by (rtac arg_subset_eclose 1); +val eclose_mono = result(); + +(** Idempotence of eclose **) + +goal Epsilon.thy "eclose(eclose(A)) = eclose(A)"; +by (rtac equalityI 1); +by (rtac ([Transset_eclose, subset_refl] MRS eclose_least) 1); +by (rtac arg_subset_eclose 1); +val eclose_idem = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/Epsilon.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Epsilon.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,18 @@ +(* Title: ZF/epsilon.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Epsilon induction and recursion +*) + +Epsilon = Nat + +consts + eclose,rank :: "i=>i" + transrec :: "[i, [i,i]=>i] =>i" + +rules + eclose_def "eclose(A) == UN n:nat. nat_rec(n, A, %m r. Union(r))" + transrec_def "transrec(a,H) == wfrec(Memrel(eclose({a})), a, H)" + rank_def "rank(a) == transrec(a, %x f. UN y:x. succ(f`y))" +end diff -r 000000000000 -r a5a9c433f639 src/ZF/Fixedpt.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Fixedpt.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,317 @@ +(* Title: ZF/fixedpt.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For fixedpt.thy. Least and greatest fixed points; the Knaster-Tarski Theorem + +Proved in the lattice of subsets of D, namely Pow(D), with Inter as glb +*) + +open Fixedpt; + +(*** Monotone operators ***) + +val prems = goalw Fixedpt.thy [bnd_mono_def] + "[| h(D)<=D; \ +\ !!W X. [| W<=D; X<=D; W<=X |] ==> h(W) <= h(X) \ +\ |] ==> bnd_mono(D,h)"; +by (REPEAT (ares_tac (prems@[conjI,allI,impI]) 1 + ORELSE etac subset_trans 1)); +val bnd_monoI = result(); + +val [major] = goalw Fixedpt.thy [bnd_mono_def] "bnd_mono(D,h) ==> h(D) <= D"; +by (rtac (major RS conjunct1) 1); +val bnd_monoD1 = result(); + +val major::prems = goalw Fixedpt.thy [bnd_mono_def] + "[| bnd_mono(D,h); W<=X; X<=D |] ==> h(W) <= h(X)"; +by (rtac (major RS conjunct2 RS spec RS spec RS mp RS mp) 1); +by (REPEAT (resolve_tac prems 1)); +val bnd_monoD2 = result(); + +val [major,minor] = goal Fixedpt.thy + "[| bnd_mono(D,h); X<=D |] ==> h(X) <= D"; +by (rtac (major RS bnd_monoD2 RS subset_trans) 1); +by (rtac (major RS bnd_monoD1) 3); +by (rtac minor 1); +by (rtac subset_refl 1); +val bnd_mono_subset = result(); + +goal Fixedpt.thy "!!A B. [| bnd_mono(D,h); A <= D; B <= D |] ==> \ +\ h(A) Un h(B) <= h(A Un B)"; +by (REPEAT (ares_tac [Un_upper1, Un_upper2, Un_least] 1 + ORELSE etac bnd_monoD2 1)); +val bnd_mono_Un = result(); + +(*Useful??*) +goal Fixedpt.thy "!!A B. [| bnd_mono(D,h); A <= D; B <= D |] ==> \ +\ h(A Int B) <= h(A) Int h(B)"; +by (REPEAT (ares_tac [Int_lower1, Int_lower2, Int_greatest] 1 + ORELSE etac bnd_monoD2 1)); +val bnd_mono_Int = result(); + +(**** Proof of Knaster-Tarski Theorem for the lfp ****) + +(*lfp is contained in each pre-fixedpoint*) +val prems = goalw Fixedpt.thy [lfp_def] + "[| h(A) <= A; A<=D |] ==> lfp(D,h) <= A"; +by (rtac (PowI RS CollectI RS Inter_lower) 1); +by (REPEAT (resolve_tac prems 1)); +val lfp_lowerbound = result(); + +(*Unfolding the defn of Inter dispenses with the premise bnd_mono(D,h)!*) +goalw Fixedpt.thy [lfp_def,Inter_def] "lfp(D,h) <= D"; +by (fast_tac ZF_cs 1); +val lfp_subset = result(); + +(*Used in datatype package*) +val [rew] = goal Fixedpt.thy "A==lfp(D,h) ==> A <= D"; +by (rewtac rew); +by (rtac lfp_subset 1); +val def_lfp_subset = result(); + +val subset0_cs = FOL_cs + addSIs [ballI, InterI, CollectI, PowI, empty_subsetI] + addIs [bexI, UnionI, ReplaceI, RepFunI] + addSEs [bexE, make_elim PowD, UnionE, ReplaceE, RepFunE, + CollectE, emptyE] + addEs [rev_ballE, InterD, make_elim InterD, subsetD]; + +val subset_cs = subset0_cs + addSIs [subset_refl,cons_subsetI,subset_consI,Union_least,UN_least,Un_least, + Inter_greatest,Int_greatest,RepFun_subset] + addSIs [Un_upper1,Un_upper2,Int_lower1,Int_lower2] + addIs [Union_upper,Inter_lower] + addSEs [cons_subsetE]; + +val prems = goalw Fixedpt.thy [lfp_def] + "[| h(D) <= D; !!X. [| h(X) <= X; X<=D |] ==> A<=X |] ==> \ +\ A <= lfp(D,h)"; +br (Pow_top RS CollectI RS Inter_greatest) 1; +by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [CollectE,PowD] 1)); +val lfp_greatest = result(); + +val hmono::prems = goal Fixedpt.thy + "[| bnd_mono(D,h); h(A)<=A; A<=D |] ==> h(lfp(D,h)) <= A"; +by (rtac (hmono RS bnd_monoD2 RS subset_trans) 1); +by (rtac lfp_lowerbound 1); +by (REPEAT (resolve_tac prems 1)); +val lfp_lemma1 = result(); + +val [hmono] = goal Fixedpt.thy + "bnd_mono(D,h) ==> h(lfp(D,h)) <= lfp(D,h)"; +by (rtac (bnd_monoD1 RS lfp_greatest) 1); +by (rtac lfp_lemma1 2); +by (REPEAT (ares_tac [hmono] 1)); +val lfp_lemma2 = result(); + +val [hmono] = goal Fixedpt.thy + "bnd_mono(D,h) ==> lfp(D,h) <= h(lfp(D,h))"; +by (rtac lfp_lowerbound 1); +by (rtac (hmono RS bnd_monoD2) 1); +by (rtac (hmono RS lfp_lemma2) 1); +by (rtac (hmono RS bnd_mono_subset) 2); +by (REPEAT (rtac lfp_subset 1)); +val lfp_lemma3 = result(); + +val prems = goal Fixedpt.thy + "bnd_mono(D,h) ==> lfp(D,h) = h(lfp(D,h))"; +by (REPEAT (resolve_tac (prems@[equalityI,lfp_lemma2,lfp_lemma3]) 1)); +val lfp_Tarski = result(); + +(*Definition form, to control unfolding*) +val [rew,mono] = goal Fixedpt.thy + "[| A==lfp(D,h); bnd_mono(D,h) |] ==> A = h(A)"; +by (rewtac rew); +by (rtac (mono RS lfp_Tarski) 1); +val def_lfp_Tarski = result(); + +(*** General induction rule for least fixedpoints ***) + +val [hmono,indstep] = goal Fixedpt.thy + "[| bnd_mono(D,h); !!x. x : h(Collect(lfp(D,h),P)) ==> P(x) \ +\ |] ==> h(Collect(lfp(D,h),P)) <= Collect(lfp(D,h),P)"; +by (rtac subsetI 1); +by (rtac CollectI 1); +by (etac indstep 2); +by (rtac (hmono RS lfp_lemma2 RS subsetD) 1); +by (rtac (hmono RS bnd_monoD2 RS subsetD) 1); +by (REPEAT (ares_tac [Collect_subset, lfp_subset] 1)); +val Collect_is_pre_fixedpt = result(); + +(*This rule yields an induction hypothesis in which the components of a + data structure may be assumed to be elements of lfp(D,h)*) +val prems = goal Fixedpt.thy + "[| bnd_mono(D,h); a : lfp(D,h); \ +\ !!x. x : h(Collect(lfp(D,h),P)) ==> P(x) \ +\ |] ==> P(a)"; +by (rtac (Collect_is_pre_fixedpt RS lfp_lowerbound RS subsetD RS CollectD2) 1); +by (rtac (lfp_subset RS (Collect_subset RS subset_trans)) 3); +by (REPEAT (ares_tac prems 1)); +val induct = result(); + +(*Definition form, to control unfolding*) +val rew::prems = goal Fixedpt.thy + "[| A == lfp(D,h); bnd_mono(D,h); a:A; \ +\ !!x. x : h(Collect(A,P)) ==> P(x) \ +\ |] ==> P(a)"; +by (rtac induct 1); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems) 1)); +val def_induct = result(); + +(*This version is useful when "A" is not a subset of D; + second premise could simply be h(D Int A) <= D or !!X. X<=D ==> h(X)<=D *) +val [hsub,hmono] = goal Fixedpt.thy + "[| h(D Int A) <= A; bnd_mono(D,h) |] ==> lfp(D,h) <= A"; +by (rtac (lfp_lowerbound RS subset_trans) 1); +by (rtac (hmono RS bnd_mono_subset RS Int_greatest) 1); +by (REPEAT (resolve_tac [hsub,Int_lower1,Int_lower2] 1)); +val lfp_Int_lowerbound = result(); + +(*Monotonicity of lfp, where h precedes i under a domain-like partial order + monotonicity of h is not strictly necessary; h must be bounded by D*) +val [hmono,imono,subhi] = goal Fixedpt.thy + "[| bnd_mono(D,h); bnd_mono(E,i); \ +\ !!X. X<=D ==> h(X) <= i(X) |] ==> lfp(D,h) <= lfp(E,i)"; +br (bnd_monoD1 RS lfp_greatest) 1; +br imono 1; +by (rtac (hmono RSN (2, lfp_Int_lowerbound)) 1); +by (rtac (Int_lower1 RS subhi RS subset_trans) 1); +by (rtac (imono RS bnd_monoD2 RS subset_trans) 1); +by (REPEAT (ares_tac [Int_lower2] 1)); +val lfp_mono = result(); + +(*This (unused) version illustrates that monotonicity is not really needed, + but both lfp's must be over the SAME set D; Inter is anti-monotonic!*) +val [isubD,subhi] = goal Fixedpt.thy + "[| i(D) <= D; !!X. X<=D ==> h(X) <= i(X) |] ==> lfp(D,h) <= lfp(D,i)"; +br lfp_greatest 1; +br isubD 1; +by (rtac lfp_lowerbound 1); +be (subhi RS subset_trans) 1; +by (REPEAT (assume_tac 1)); +val lfp_mono2 = result(); + + +(**** Proof of Knaster-Tarski Theorem for the gfp ****) + +(*gfp contains each post-fixedpoint that is contained in D*) +val prems = goalw Fixedpt.thy [gfp_def] + "[| A <= h(A); A<=D |] ==> A <= gfp(D,h)"; +by (rtac (PowI RS CollectI RS Union_upper) 1); +by (REPEAT (resolve_tac prems 1)); +val gfp_upperbound = result(); + +goalw Fixedpt.thy [gfp_def] "gfp(D,h) <= D"; +by (fast_tac ZF_cs 1); +val gfp_subset = result(); + +(*Used in datatype package*) +val [rew] = goal Fixedpt.thy "A==gfp(D,h) ==> A <= D"; +by (rewtac rew); +by (rtac gfp_subset 1); +val def_gfp_subset = result(); + +val hmono::prems = goalw Fixedpt.thy [gfp_def] + "[| bnd_mono(D,h); !!X. [| X <= h(X); X<=D |] ==> X<=A |] ==> \ +\ gfp(D,h) <= A"; +by (fast_tac (subset_cs addIs ((hmono RS bnd_monoD1)::prems)) 1); +val gfp_least = result(); + +val hmono::prems = goal Fixedpt.thy + "[| bnd_mono(D,h); A<=h(A); A<=D |] ==> A <= h(gfp(D,h))"; +by (rtac (hmono RS bnd_monoD2 RSN (2,subset_trans)) 1); +by (rtac gfp_subset 3); +by (rtac gfp_upperbound 2); +by (REPEAT (resolve_tac prems 1)); +val gfp_lemma1 = result(); + +val [hmono] = goal Fixedpt.thy + "bnd_mono(D,h) ==> gfp(D,h) <= h(gfp(D,h))"; +by (rtac gfp_least 1); +by (rtac gfp_lemma1 2); +by (REPEAT (ares_tac [hmono] 1)); +val gfp_lemma2 = result(); + +val [hmono] = goal Fixedpt.thy + "bnd_mono(D,h) ==> h(gfp(D,h)) <= gfp(D,h)"; +by (rtac gfp_upperbound 1); +by (rtac (hmono RS bnd_monoD2) 1); +by (rtac (hmono RS gfp_lemma2) 1); +by (REPEAT (rtac ([hmono, gfp_subset] MRS bnd_mono_subset) 1)); +val gfp_lemma3 = result(); + +val prems = goal Fixedpt.thy + "bnd_mono(D,h) ==> gfp(D,h) = h(gfp(D,h))"; +by (REPEAT (resolve_tac (prems@[equalityI,gfp_lemma2,gfp_lemma3]) 1)); +val gfp_Tarski = result(); + +(*Definition form, to control unfolding*) +val [rew,mono] = goal Fixedpt.thy + "[| A==gfp(D,h); bnd_mono(D,h) |] ==> A = h(A)"; +by (rewtac rew); +by (rtac (mono RS gfp_Tarski) 1); +val def_gfp_Tarski = result(); + + +(*** Coinduction rules for greatest fixed points ***) + +(*weak version*) +goal Fixedpt.thy "!!X h. [| a: X; X <= h(X); X <= D |] ==> a : gfp(D,h)"; +by (REPEAT (ares_tac [gfp_upperbound RS subsetD] 1)); +val weak_coinduct = result(); + +val [subs_h,subs_D,mono] = goal Fixedpt.thy + "[| X <= h(X Un gfp(D,h)); X <= D; bnd_mono(D,h) |] ==> \ +\ X Un gfp(D,h) <= h(X Un gfp(D,h))"; +by (rtac (subs_h RS Un_least) 1); +by (rtac (mono RS gfp_lemma2 RS subset_trans) 1); +by (rtac (Un_upper2 RS subset_trans) 1); +by (rtac ([mono, subs_D, gfp_subset] MRS bnd_mono_Un) 1); +val coinduct_lemma = result(); + +(*strong version*) +goal Fixedpt.thy + "!!X D. [| bnd_mono(D,h); a: X; X <= h(X Un gfp(D,h)); X <= D |] ==> \ +\ a : gfp(D,h)"; +by (rtac (coinduct_lemma RSN (2, weak_coinduct)) 1); +by (REPEAT (ares_tac [gfp_subset, UnI1, Un_least] 1)); +val coinduct = result(); + +(*Definition form, to control unfolding*) +val rew::prems = goal Fixedpt.thy + "[| A == gfp(D,h); bnd_mono(D,h); a: X; X <= h(X Un A); X <= D |] ==> \ +\ a : A"; +by (rewtac rew); +by (rtac coinduct 1); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems) 1)); +val def_coinduct = result(); + +(*Lemma used immediately below!*) +val [subsA,XimpP] = goal ZF.thy + "[| X <= A; !!z. z:X ==> P(z) |] ==> X <= Collect(A,P)"; +by (rtac (subsA RS subsetD RS CollectI RS subsetI) 1); +by (assume_tac 1); +by (etac XimpP 1); +val subset_Collect = result(); + +(*The version used in the induction/coinduction package*) +val prems = goal Fixedpt.thy + "[| A == gfp(D, %w. Collect(D,P(w))); bnd_mono(D, %w. Collect(D,P(w))); \ +\ a: X; X <= D; !!z. z: X ==> P(X Un A, z) |] ==> \ +\ a : A"; +by (rtac def_coinduct 1); +by (REPEAT (ares_tac (subset_Collect::prems) 1)); +val def_Collect_coinduct = result(); + +(*Monotonicity of gfp!*) +val [hmono,subde,subhi] = goal Fixedpt.thy + "[| bnd_mono(D,h); D <= E; \ +\ !!X. X<=D ==> h(X) <= i(X) |] ==> gfp(D,h) <= gfp(E,i)"; +by (rtac gfp_upperbound 1); +by (rtac (hmono RS gfp_lemma2 RS subset_trans) 1); +by (rtac (gfp_subset RS subhi) 1); +by (rtac ([gfp_subset, subde] MRS subset_trans) 1); +val gfp_mono = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/Fixedpt.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Fixedpt.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,23 @@ +(* Title: ZF/fixedpt.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Least and greatest fixed points +*) + +Fixedpt = ZF + +consts + bnd_mono :: "[i,i=>i]=>o" + lfp, gfp :: "[i,i=>i]=>i" + +rules + (*monotone operator from Pow(D) to itself*) + bnd_mono_def + "bnd_mono(D,h) == h(D)<=D & (ALL W X. W<=X --> X<=D --> h(W) <= h(X))" + + lfp_def "lfp(D,h) == Inter({X: Pow(D). h(X) <= X})" + + gfp_def "gfp(D,h) == Union({X: Pow(D). X <= h(X)})" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/List.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/List.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,80 @@ +(* Title: ZF/list.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Datatype definition of Lists +*) + +structure List = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("list", "univ(A)", + [(["Nil"], "i"), + (["Cons"], "[i,i]=>i")])]; + val rec_styp = "i=>i"; + val ext = None + val sintrs = + ["Nil : list(A)", + "[| a: A; l: list(A) |] ==> Cons(a,l) : list(A)"]; + val monos = []; + val type_intrs = data_typechecks + val type_elims = []); + +val [NilI, ConsI] = List.intrs; + +(*An elimination rule, for type-checking*) +val ConsE = List.mk_cases List.con_defs "Cons(a,l) : list(A)"; + +(*Proving freeness results*) +val Cons_iff = List.mk_free "Cons(a,l)=Cons(a',l') <-> a=a' & l=l'"; +val Nil_Cons_iff = List.mk_free "~ Nil=Cons(a,l)"; + +(*Perform induction on l, then prove the major premise using prems. *) +fun list_ind_tac a prems i = + EVERY [res_inst_tac [("x",a)] List.induct i, + rename_last_tac a ["1"] (i+2), + ares_tac prems i]; + +(** Lemmas to justify using "list" in other recursive type definitions **) + +goalw List.thy List.defs "!!A B. A<=B ==> list(A) <= list(B)"; +by (rtac lfp_mono 1); +by (REPEAT (rtac List.bnd_mono 1)); +by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); +val list_mono = result(); + +(*There is a similar proof by list induction.*) +goalw List.thy (List.defs@List.con_defs) "list(univ(A)) <= univ(A)"; +by (rtac lfp_lowerbound 1); +by (rtac (A_subset_univ RS univ_mono) 2); +by (fast_tac (ZF_cs addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ, + Pair_in_univ]) 1); +val list_univ = result(); + +val list_subset_univ = standard + (list_mono RS (list_univ RSN (2,subset_trans))); + +(***** +val major::prems = goal List.thy + "[| l: list(A); \ +\ c: C(0); \ +\ !!x y. [| x: A; y: list(A) |] ==> h(x,y): C() \ +\ |] ==> list_case(l,c,h) : C(l)"; +by (rtac (major RS list_induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (ZF_ss addrews + ([list_case_0,list_case_Pair]@prems)))); +val list_case_type = result(); +****) + + +(** For recursion **) + +goalw List.thy List.con_defs "rank(a) : rank(Cons(a,l))"; +by (SIMP_TAC rank_ss 1); +val rank_Cons1 = result(); + +goalw List.thy List.con_defs "rank(l) : rank(Cons(a,l))"; +by (SIMP_TAC rank_ss 1); +val rank_Cons2 = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/ListFn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ListFn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,306 @@ +(* Title: ZF/list-fn.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For list-fn.thy. Lists in Zermelo-Fraenkel Set Theory +*) + +open ListFn; + + +(** list_rec -- by Vset recursion **) + +(*Used to verify list_rec*) +val list_rec_ss = ZF_ss + addcongs (mk_typed_congs ListFn.thy [("h", "[i,i,i]=>i")]) + addrews List.case_eqns; + +goal ListFn.thy "list_rec(Nil,c,h) = c"; +by (rtac (list_rec_def RS def_Vrec RS trans) 1); +by (SIMP_TAC list_rec_ss 1); +val list_rec_Nil = result(); + +goal ListFn.thy "list_rec(Cons(a,l), c, h) = h(a, l, list_rec(l,c,h))"; +by (rtac (list_rec_def RS def_Vrec RS trans) 1); +by (SIMP_TAC (list_rec_ss addrews [Vset_rankI, rank_Cons2]) 1); +val list_rec_Cons = result(); + +(*Type checking -- proved by induction, as usual*) +val prems = goal ListFn.thy + "[| l: list(A); \ +\ c: C(Nil); \ +\ !!x y r. [| x:A; y: list(A); r: C(y) |] ==> h(x,y,r): C(Cons(x,y)) \ +\ |] ==> list_rec(l,c,h) : C(l)"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS (ASM_SIMP_TAC + (ZF_ss addrews (prems@[list_rec_Nil,list_rec_Cons])))); +val list_rec_type = result(); + +(** Versions for use with definitions **) + +val [rew] = goal ListFn.thy + "[| !!l. j(l)==list_rec(l, c, h) |] ==> j(Nil) = c"; +by (rewtac rew); +by (rtac list_rec_Nil 1); +val def_list_rec_Nil = result(); + +val [rew] = goal ListFn.thy + "[| !!l. j(l)==list_rec(l, c, h) |] ==> j(Cons(a,l)) = h(a,l,j(l))"; +by (rewtac rew); +by (rtac list_rec_Cons 1); +val def_list_rec_Cons = result(); + +fun list_recs def = map standard + ([def] RL [def_list_rec_Nil, def_list_rec_Cons]); + +(** map **) + +val [map_Nil,map_Cons] = list_recs map_def; + +val prems = goalw ListFn.thy [map_def] + "[| l: list(A); !!x. x: A ==> h(x): B |] ==> map(h,l) : list(B)"; +by (REPEAT (ares_tac (prems@[list_rec_type, NilI, ConsI]) 1)); +val map_type = result(); + +val [major] = goal ListFn.thy "l: list(A) ==> map(h,l) : list({h(u). u:A})"; +by (rtac (major RS map_type) 1); +by (etac RepFunI 1); +val map_type2 = result(); + +(** length **) + +val [length_Nil,length_Cons] = list_recs length_def; + +val prems = goalw ListFn.thy [length_def] + "l: list(A) ==> length(l) : nat"; +by (REPEAT (ares_tac (prems @ [list_rec_type, nat_0I, nat_succI]) 1)); +val length_type = result(); + +(** app **) + +val [app_Nil,app_Cons] = list_recs app_def; + +val prems = goalw ListFn.thy [app_def] + "[| xs: list(A); ys: list(A) |] ==> xs@ys : list(A)"; +by (REPEAT (ares_tac (prems @ [list_rec_type, ConsI]) 1)); +val app_type = result(); + +(** rev **) + +val [rev_Nil,rev_Cons] = list_recs rev_def; + +val prems = goalw ListFn.thy [rev_def] + "xs: list(A) ==> rev(xs) : list(A)"; +by (REPEAT (ares_tac (prems @ [list_rec_type, NilI, ConsI, app_type]) 1)); +val rev_type = result(); + + +(** flat **) + +val [flat_Nil,flat_Cons] = list_recs flat_def; + +val prems = goalw ListFn.thy [flat_def] + "ls: list(list(A)) ==> flat(ls) : list(A)"; +by (REPEAT (ares_tac (prems @ [list_rec_type, NilI, ConsI, app_type]) 1)); +val flat_type = result(); + + +(** list_add **) + +val [list_add_Nil,list_add_Cons] = list_recs list_add_def; + +val prems = goalw ListFn.thy [list_add_def] + "xs: list(nat) ==> list_add(xs) : nat"; +by (REPEAT (ares_tac (prems @ [list_rec_type, nat_0I, add_type]) 1)); +val list_add_type = result(); + +(** ListFn simplification **) + +val list_typechecks = + [NilI, ConsI, list_rec_type, + map_type, map_type2, app_type, length_type, rev_type, flat_type, + list_add_type]; + +val list_congs = + List.congs @ + mk_congs ListFn.thy + ["list_rec","map","op @","length","rev","flat","list_add"]; + +val list_ss = arith_ss + addcongs list_congs + addrews List.case_eqns + addrews list_typechecks + addrews [list_rec_Nil, list_rec_Cons, + map_Nil, map_Cons, + app_Nil, app_Cons, + length_Nil, length_Cons, + rev_Nil, rev_Cons, + flat_Nil, flat_Cons, + list_add_Nil, list_add_Cons]; + +(*** theorems about map ***) + +val prems = goal ListFn.thy + "l: list(A) ==> map(%u.u, l) = l"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val map_ident = result(); + +val prems = goal ListFn.thy + "l: list(A) ==> map(h, map(j,l)) = map(%u.h(j(u)), l)"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val map_compose = result(); + +val prems = goal ListFn.thy + "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)"; +by (list_ind_tac "xs" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val map_app_distrib = result(); + +val prems = goal ListFn.thy + "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))"; +by (list_ind_tac "ls" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [map_app_distrib]))); +val map_flat = result(); + +val prems = goal ListFn.thy + "l: list(A) ==> \ +\ list_rec(map(h,l), c, d) = \ +\ list_rec(l, c, %x xs r. d(h(x), map(h,xs), r))"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS + (ASM_SIMP_TAC + (list_ss addcongs (mk_typed_congs ListFn.thy [("d", "[i,i,i]=>i")])))); +val list_rec_map = result(); + +(** theorems about list(Collect(A,P)) -- used in ex/term.ML **) + +(* c : list(Collect(B,P)) ==> c : list(B) *) +val list_CollectD = standard (Collect_subset RS list_mono RS subsetD); + +val prems = goal ListFn.thy + "l: list({x:A. h(x)=j(x)}) ==> map(h,l) = map(j,l)"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val map_list_Collect = result(); + +(*** theorems about length ***) + +val prems = goal ListFn.thy + "xs: list(A) ==> length(map(h,xs)) = length(xs)"; +by (list_ind_tac "xs" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val length_map = result(); + +val prems = goal ListFn.thy + "xs: list(A) ==> length(xs@ys) = length(xs) #+ length(ys)"; +by (list_ind_tac "xs" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val length_app = result(); + +(* [| m: nat; n: nat |] ==> m #+ succ(n) = succ(n) #+ m + Used for rewriting below*) +val add_commute_succ = nat_succI RSN (2,add_commute); + +val prems = goal ListFn.thy + "xs: list(A) ==> length(rev(xs)) = length(xs)"; +by (list_ind_tac "xs" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [length_app, add_commute_succ]))); +val length_rev = result(); + +val prems = goal ListFn.thy + "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))"; +by (list_ind_tac "ls" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [length_app]))); +val length_flat = result(); + +(*** theorems about app ***) + +val [major] = goal ListFn.thy "xs: list(A) ==> xs@Nil=xs"; +by (rtac (major RS List.induct) 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val app_right_Nil = result(); + +val prems = goal ListFn.thy "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)"; +by (list_ind_tac "xs" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val app_assoc = result(); + +val prems = goal ListFn.thy + "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)"; +by (list_ind_tac "ls" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [app_assoc]))); +val flat_app_distrib = result(); + +(*** theorems about rev ***) + +val prems = goal ListFn.thy "l: list(A) ==> rev(map(h,l)) = map(h,rev(l))"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [map_app_distrib]))); +val rev_map_distrib = result(); + +(*Simplifier needs the premises as assumptions because rewriting will not + instantiate the variable ?A in the rules' typing conditions; note that + rev_type does not instantiate ?A. Only the premises do. +*) +val prems = goal ListFn.thy + "[| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)"; +by (cut_facts_tac prems 1); +by (etac List.induct 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [app_right_Nil,app_assoc]))); +val rev_app_distrib = result(); + +val prems = goal ListFn.thy "l: list(A) ==> rev(rev(l))=l"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [rev_app_distrib]))); +val rev_rev_ident = result(); + +val prems = goal ListFn.thy + "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))"; +by (list_ind_tac "ls" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews + [map_app_distrib, flat_app_distrib, rev_app_distrib, app_right_Nil]))); +val rev_flat = result(); + + +(*** theorems about list_add ***) + +val prems = goal ListFn.thy + "[| xs: list(nat); ys: list(nat) |] ==> \ +\ list_add(xs@ys) = list_add(ys) #+ list_add(xs)"; +by (cut_facts_tac prems 1); +by (list_ind_tac "xs" prems 1); +by (ALLGOALS + (ASM_SIMP_TAC (list_ss addrews [add_0_right, add_assoc RS sym]))); +by (resolve_tac arith_congs 1); +by (REPEAT (ares_tac [refl, list_add_type, add_commute] 1)); +val list_add_app = result(); + +val prems = goal ListFn.thy + "l: list(nat) ==> list_add(rev(l)) = list_add(l)"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS + (ASM_SIMP_TAC (list_ss addrews [list_add_app, add_0_right]))); +val list_add_rev = result(); + +val prems = goal ListFn.thy + "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))"; +by (list_ind_tac "ls" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [list_add_app]))); +by (REPEAT (ares_tac [refl, list_add_type, map_type, add_commute] 1)); +val list_add_flat = result(); + +(** New induction rule **) + +val major::prems = goal ListFn.thy + "[| l: list(A); \ +\ P(Nil); \ +\ !!x y. [| x: A; y: list(A); P(y) |] ==> P(y @ [x]) \ +\ |] ==> P(l)"; +by (rtac (major RS rev_rev_ident RS subst) 1); +by (rtac (major RS rev_type RS List.induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews prems))); +val list_append_induct = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/ListFn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ListFn.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,44 @@ +(* Title: ZF/list-fn + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Functions for Lists in Zermelo-Fraenkel Set Theory + +map is a binding operator -- it applies to meta-level functions, not +object-level functions. This simplifies the final form of term_rec_conv, +although complicating its derivation. +*) + +ListFn = List + +consts + "@" :: "[i,i]=>i" (infixr 60) + list_rec :: "[i, i, [i,i,i]=>i] => i" + map :: "[i=>i, i] => i" + length,rev :: "i=>i" + flat :: "i=>i" + list_add :: "i=>i" + + (* List Enumeration *) + "[]" :: "i" ("[]") + "@List" :: "args => i" ("[(_)]") + + +translations + "[x, xs]" == "Cons(x, [xs])" + "[x]" == "Cons(x, [])" + "[]" == "Nil" + + +rules + list_rec_def + "list_rec(l,c,h) == Vrec(l, %l g.list_case(c, %x xs. h(x, xs, g`xs), l))" + + map_def "map(f,l) == list_rec(l, Nil, %x xs r. Cons(f(x), r))" + length_def "length(l) == list_rec(l, 0, %x xs r. succ(r))" + app_def "xs@ys == list_rec(xs, ys, %x xs r. Cons(x,r))" + rev_def "rev(l) == list_rec(l, Nil, %x xs r. r @ [x])" + flat_def "flat(ls) == list_rec(ls, Nil, %l ls r. l @ r)" + list_add_def "list_add(l) == list_rec(l, 0, %x xs r. x#+r)" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Makefile Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,50 @@ +######################################################################### +# # +# Makefile for Isabelle (ZF) # +# # +######################################################################### + +#To make the system, cd to this directory and type +# make -f Makefile +#To make the system and test it on standard examples, type +# make -f Makefile test + +#Environment variable ISABELLECOMP specifies the compiler. +#Environment variable ISABELLEBIN specifies the destination directory. +#For Poly/ML, ISABELLEBIN must begin with a / + +#Makes FOL if this file is ABSENT -- but not +#if it is out of date, since this Makefile does not know its dependencies! + +BIN = $(ISABELLEBIN) +COMP = $(ISABELLECOMP) +FILES = ROOT.ML zf.thy zf.ML upair.ML subset.ML pair.ML domrange.ML \ + func.ML simpdata.ML bool.thy bool.ML \ + sum.thy sum.ML qpair.thy qpair.ML mono.ML fixedpt.thy fixedpt.ML \ + ind-syntax.ML intr-elim.ML indrule.ML inductive.ML co-inductive.ML \ + equalities.ML perm.thy perm.ML trancl.thy trancl.ML \ + wf.thy wf.ML ordinal.thy ordinal.ML nat.thy nat.ML \ + epsilon.thy epsilon.ML arith.thy arith.ML univ.thy univ.ML \ + quniv.thy quniv.ML constructor.ML datatype.ML \ + fin.ML list.ML list-fn.thy list-fn.ML + +#Uses cp rather than make_database because Poly/ML allows only 3 levels +$(BIN)/ZF: $(BIN)/FOL $(FILES) + case "$(COMP)" in \ + poly*) cp $(BIN)/FOL $(BIN)/ZF;\ + echo 'open PolyML; use"ROOT";' | $(COMP) $(BIN)/ZF ;;\ + sml*) echo 'use"ROOT.ML"; xML"$(BIN)/ZF" banner;' | $(BIN)/FOL;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +$(BIN)/FOL: + cd ../FOL; $(MAKE) + +test: ex/ROOT.ML $(BIN)/ZF + case "$(COMP)" in \ + poly*) echo 'use"ex/ROOT.ML"; quit();' | $(COMP) $(BIN)/ZF ;;\ + sml*) echo 'use"ex/ROOT.ML";' | $(BIN)/ZF;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +.PRECIOUS: $(BIN)/FOL $(BIN)/ZF diff -r 000000000000 -r a5a9c433f639 src/ZF/Nat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Nat.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,182 @@ +(* Title: ZF/nat.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For nat.thy. Natural numbers in Zermelo-Fraenkel Set Theory +*) + +open Nat; + +goal Nat.thy "bnd_mono(Inf, %X. {0} Un {succ(i). i:X})"; +by (rtac bnd_monoI 1); +by (REPEAT (ares_tac [subset_refl, RepFun_mono, Un_mono] 2)); +by (cut_facts_tac [infinity] 1); +by (fast_tac ZF_cs 1); +val nat_bnd_mono = result(); + +(* nat = {0} Un {succ(x). x:nat} *) +val nat_unfold = nat_bnd_mono RS (nat_def RS def_lfp_Tarski); + +(** Type checking of 0 and successor **) + +goal Nat.thy "0 : nat"; +by (rtac (nat_unfold RS ssubst) 1); +by (rtac (singletonI RS UnI1) 1); +val nat_0I = result(); + +val prems = goal Nat.thy "n : nat ==> succ(n) : nat"; +by (rtac (nat_unfold RS ssubst) 1); +by (rtac (RepFunI RS UnI2) 1); +by (resolve_tac prems 1); +val nat_succI = result(); + +goalw Nat.thy [one_def] "1 : nat"; +by (rtac (nat_0I RS nat_succI) 1); +val nat_1I = result(); + +goal Nat.thy "bool <= nat"; +by (REPEAT (ares_tac [subsetI,nat_0I,nat_1I] 1 ORELSE etac boolE 1)); +val bool_subset_nat = result(); + +val bool_into_nat = bool_subset_nat RS subsetD; + + +(** Injectivity properties and induction **) + +(*Mathematical induction*) +val major::prems = goal Nat.thy + "[| n: nat; P(0); !!x. [| x: nat; P(x) |] ==> P(succ(x)) |] ==> P(n)"; +by (rtac ([nat_def, nat_bnd_mono, major] MRS def_induct) 1); +by (fast_tac (ZF_cs addIs prems) 1); +val nat_induct = result(); + +(*Perform induction on n, then prove the n:nat subgoal using prems. *) +fun nat_ind_tac a prems i = + EVERY [res_inst_tac [("n",a)] nat_induct i, + rename_last_tac a ["1"] (i+2), + ares_tac prems i]; + +val major::prems = goal Nat.thy + "[| n: nat; n=0 ==> P; !!x. [| x: nat; n=succ(x) |] ==> P |] ==> P"; +br (major RS (nat_unfold RS equalityD1 RS subsetD) RS UnE) 1; +by (DEPTH_SOLVE (eresolve_tac [singletonE,RepFunE] 1 + ORELSE ares_tac prems 1)); +val natE = result(); + +val prems = goal Nat.thy "n: nat ==> Ord(n)"; +by (nat_ind_tac "n" prems 1); +by (REPEAT (ares_tac [Ord_0, Ord_succ] 1)); +val naturals_are_ordinals = result(); + +goal Nat.thy "!!n. n: nat ==> n=0 | 0:n"; +by (etac nat_induct 1); +by (fast_tac ZF_cs 1); +by (fast_tac (ZF_cs addIs [naturals_are_ordinals RS Ord_0_mem_succ]) 1); +val natE0 = result(); + +goal Nat.thy "Ord(nat)"; +by (rtac OrdI 1); +by (etac (naturals_are_ordinals RS Ord_is_Transset) 2); +by (rewtac Transset_def); +by (rtac ballI 1); +by (etac nat_induct 1); +by (REPEAT (ares_tac [empty_subsetI,succ_subsetI] 1)); +val Ord_nat = result(); + +(** Variations on mathematical induction **) + +(*complete induction*) +val complete_induct = Ord_nat RSN (2, Ord_induct); + +val prems = goal Nat.thy + "[| m: nat; n: nat; \ +\ !!x. [| x: nat; m<=x; P(x) |] ==> P(succ(x)) \ +\ |] ==> m <= n --> P(m) --> P(n)"; +by (nat_ind_tac "n" prems 1); +by (ALLGOALS + (ASM_SIMP_TAC + (ZF_ss addrews (prems@distrib_rews@[subset_empty_iff, subset_succ_iff, + Ord_nat RS Ord_in_Ord])))); +val nat_induct_from_lemma = result(); + +(*Induction starting from m rather than 0*) +val prems = goal Nat.thy + "[| m <= n; m: nat; n: nat; \ +\ P(m); \ +\ !!x. [| x: nat; m<=x; P(x) |] ==> P(succ(x)) \ +\ |] ==> P(n)"; +by (rtac (nat_induct_from_lemma RS mp RS mp) 1); +by (REPEAT (ares_tac prems 1)); +val nat_induct_from = result(); + +(*Induction suitable for subtraction and less-than*) +val prems = goal Nat.thy + "[| m: nat; n: nat; \ +\ !!x. [| x: nat |] ==> P(x,0); \ +\ !!y. [| y: nat |] ==> P(0,succ(y)); \ +\ !!x y. [| x: nat; y: nat; P(x,y) |] ==> P(succ(x),succ(y)) \ +\ |] ==> P(m,n)"; +by (res_inst_tac [("x","m")] bspec 1); +by (resolve_tac prems 2); +by (nat_ind_tac "n" prems 1); +by (rtac ballI 2); +by (nat_ind_tac "x" [] 2); +by (REPEAT (ares_tac (prems@[ballI]) 1 ORELSE etac bspec 1)); +val diff_induct = result(); + +(** nat_case **) + +goalw Nat.thy [nat_case_def] "nat_case(0,a,b) = a"; +by (fast_tac (ZF_cs addIs [the_equality]) 1); +val nat_case_0 = result(); + +goalw Nat.thy [nat_case_def] "nat_case(succ(m),a,b) = b(m)"; +by (fast_tac (ZF_cs addIs [the_equality]) 1); +val nat_case_succ = result(); + +val major::prems = goal Nat.thy + "[| n: nat; a: C(0); !!m. m: nat ==> b(m): C(succ(m)) \ +\ |] ==> nat_case(n,a,b) : C(n)"; +by (rtac (major RS nat_induct) 1); +by (REPEAT (resolve_tac [nat_case_0 RS ssubst, + nat_case_succ RS ssubst] 1 + THEN resolve_tac prems 1)); +by (assume_tac 1); +val nat_case_type = result(); + +val prems = goalw Nat.thy [nat_case_def] + "[| n=n'; a=a'; !!m z. b(m)=b'(m) \ +\ |] ==> nat_case(n,a,b)=nat_case(n',a',b')"; +by (REPEAT (resolve_tac [the_cong,disj_cong,ex_cong] 1 + ORELSE EVERY1 (map rtac ((prems RL [ssubst]) @ [iff_refl])))); +val nat_case_cong = result(); + + +(** nat_rec -- used to define eclose and transrec, then obsolete **) + +val nat_rec_trans = wf_Memrel RS (nat_rec_def RS def_wfrec RS trans); + +goal Nat.thy "nat_rec(0,a,b) = a"; +by (rtac nat_rec_trans 1); +by (rtac nat_case_0 1); +val nat_rec_0 = result(); + +val [prem] = goal Nat.thy + "m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))"; +val nat_rec_ss = ZF_ss + addcongs (mk_typed_congs Nat.thy [("b", "[i,i]=>i")]) + addrews [prem, nat_case_succ, nat_succI, Memrel_iff, + vimage_singleton_iff]; +by (rtac nat_rec_trans 1); +by (SIMP_TAC nat_rec_ss 1); +val nat_rec_succ = result(); + +(** The union of two natural numbers is a natural number -- their maximum **) + +(* [| ?i : nat; ?j : nat |] ==> ?i Un ?j : nat *) +val Un_nat_type = standard (Ord_nat RSN (3,Ord_member_UnI)); + +(* [| ?i : nat; ?j : nat |] ==> ?i Int ?j : nat *) +val Int_nat_type = standard (Ord_nat RSN (3,Ord_member_IntI)); + diff -r 000000000000 -r a5a9c433f639 src/ZF/Nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Nat.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,26 @@ +(* Title: ZF/nat.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Natural numbers in Zermelo-Fraenkel Set Theory +*) + +Nat = Ord + Bool + +consts + nat :: "i" + nat_case :: "[i, i, i=>i]=>i" + nat_rec :: "[i, i, [i,i]=>i]=>i" + +rules + + nat_def "nat == lfp(Inf, %X. {0} Un {succ(i). i:X})" + + nat_case_def + "nat_case(k,a,b) == THE y. k=0 & y=a | (EX x. k=succ(x) & y=b(x))" + + nat_rec_def + "nat_rec(k,a,b) == \ +\ wfrec(Memrel(nat), k, %n f. nat_case(n, a, %m. b(m, f`m)))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/Ord.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Ord.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,410 @@ +(* Title: ZF/ordinal.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For ordinal.thy. Ordinals in Zermelo-Fraenkel Set Theory +*) + +open Ord; + +(*** Rules for Transset ***) + +(** Two neat characterisations of Transset **) + +goalw Ord.thy [Transset_def] "Transset(A) <-> A<=Pow(A)"; +by (fast_tac ZF_cs 1); +val Transset_iff_Pow = result(); + +goalw Ord.thy [Transset_def] "Transset(A) <-> Union(succ(A)) = A"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val Transset_iff_Union_succ = result(); + +(** Consequences of downwards closure **) + +goalw Ord.thy [Transset_def] + "!!C a b. [| Transset(C); {a,b}: C |] ==> a:C & b: C"; +by (fast_tac ZF_cs 1); +val Transset_doubleton_D = result(); + +val [prem1,prem2] = goalw Ord.thy [Pair_def] + "[| Transset(C); : C |] ==> a:C & b: C"; +by (cut_facts_tac [prem2] 1); +by (fast_tac (ZF_cs addSDs [prem1 RS Transset_doubleton_D]) 1); +val Transset_Pair_D = result(); + +val prem1::prems = goal Ord.thy + "[| Transset(C); A*B <= C; b: B |] ==> A <= C"; +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addSDs [prem1 RS Transset_Pair_D]) 1); +val Transset_includes_domain = result(); + +val prem1::prems = goal Ord.thy + "[| Transset(C); A*B <= C; a: A |] ==> B <= C"; +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addSDs [prem1 RS Transset_Pair_D]) 1); +val Transset_includes_range = result(); + +val [prem1,prem2] = goalw (merge_theories(Ord.thy,Sum.thy)) [sum_def] + "[| Transset(C); A+B <= C |] ==> A <= C & B <= C"; +br (prem2 RS (Un_subset_iff RS iffD1) RS conjE) 1; +by (REPEAT (etac (prem1 RS Transset_includes_range) 1 + ORELSE resolve_tac [conjI, singletonI] 1)); +val Transset_includes_summands = result(); + +val [prem] = goalw (merge_theories(Ord.thy,Sum.thy)) [sum_def] + "Transset(C) ==> (A+B) Int C <= (A Int C) + (B Int C)"; +br (Int_Un_distrib RS ssubst) 1; +by (fast_tac (ZF_cs addSDs [prem RS Transset_Pair_D]) 1); +val Transset_sum_Int_subset = result(); + +(** Closure properties **) + +goalw Ord.thy [Transset_def] "Transset(0)"; +by (fast_tac ZF_cs 1); +val Transset_0 = result(); + +goalw Ord.thy [Transset_def] + "!!i j. [| Transset(i); Transset(j) |] ==> Transset(i Un j)"; +by (fast_tac ZF_cs 1); +val Transset_Un = result(); + +goalw Ord.thy [Transset_def] + "!!i j. [| Transset(i); Transset(j) |] ==> Transset(i Int j)"; +by (fast_tac ZF_cs 1); +val Transset_Int = result(); + +goalw Ord.thy [Transset_def] "!!i. Transset(i) ==> Transset(succ(i))"; +by (fast_tac ZF_cs 1); +val Transset_succ = result(); + +goalw Ord.thy [Transset_def] "!!i. Transset(i) ==> Transset(Pow(i))"; +by (fast_tac ZF_cs 1); +val Transset_Pow = result(); + +goalw Ord.thy [Transset_def] "!!A. Transset(A) ==> Transset(Union(A))"; +by (fast_tac ZF_cs 1); +val Transset_Union = result(); + +val [Transprem] = goalw Ord.thy [Transset_def] + "[| !!i. i:A ==> Transset(i) |] ==> Transset(Union(A))"; +by (fast_tac (ZF_cs addEs [Transprem RS bspec RS subsetD]) 1); +val Transset_Union_family = result(); + +val [prem,Transprem] = goalw Ord.thy [Transset_def] + "[| j:A; !!i. i:A ==> Transset(i) |] ==> Transset(Inter(A))"; +by (cut_facts_tac [prem] 1); +by (fast_tac (ZF_cs addEs [Transprem RS bspec RS subsetD]) 1); +val Transset_Inter_family = result(); + +(*** Natural Deduction rules for Ord ***) + +val prems = goalw Ord.thy [Ord_def] + "[| Transset(i); !!x. x:i ==> Transset(x) |] ==> Ord(i) "; +by (REPEAT (ares_tac (prems@[ballI,conjI]) 1)); +val OrdI = result(); + +val [major] = goalw Ord.thy [Ord_def] + "Ord(i) ==> Transset(i)"; +by (rtac (major RS conjunct1) 1); +val Ord_is_Transset = result(); + +val [major,minor] = goalw Ord.thy [Ord_def] + "[| Ord(i); j:i |] ==> Transset(j) "; +by (rtac (minor RS (major RS conjunct2 RS bspec)) 1); +val Ord_contains_Transset = result(); + +(*** Lemmas for ordinals ***) + +goalw Ord.thy [Ord_def,Transset_def] "!!i j. [| Ord(i); j:i |] ==> Ord(j) "; +by (fast_tac ZF_cs 1); +val Ord_in_Ord = result(); + +goal Ord.thy "!!i j. [| Ord(i); Transset(j); j<=i |] ==> Ord(j)"; +by (REPEAT (ares_tac [OrdI] 1 + ORELSE eresolve_tac [Ord_contains_Transset, subsetD] 1)); +val Ord_subset_Ord = result(); + +goalw Ord.thy [Ord_def,Transset_def] "!!i j. [| j:i; Ord(i) |] ==> j<=i"; +by (fast_tac ZF_cs 1); +val OrdmemD = result(); + +goal Ord.thy "!!i j k. [| i:j; j:k; Ord(k) |] ==> i:k"; +by (REPEAT (ares_tac [OrdmemD RS subsetD] 1)); +val Ord_trans = result(); + +goal Ord.thy "!!i j. [| i:j; Ord(j) |] ==> succ(i) <= j"; +by (REPEAT (ares_tac [OrdmemD RSN (2,succ_subsetI)] 1)); +val Ord_succ_subsetI = result(); + + +(*** The construction of ordinals: 0, succ, Union ***) + +goal Ord.thy "Ord(0)"; +by (REPEAT (ares_tac [OrdI,Transset_0] 1 ORELSE etac emptyE 1)); +val Ord_0 = result(); + +goal Ord.thy "!!i. Ord(i) ==> Ord(succ(i))"; +by (REPEAT (ares_tac [OrdI,Transset_succ] 1 + ORELSE eresolve_tac [succE,ssubst,Ord_is_Transset, + Ord_contains_Transset] 1)); +val Ord_succ = result(); + +val nonempty::prems = goal Ord.thy + "[| j:A; !!i. i:A ==> Ord(i) |] ==> Ord(Inter(A))"; +by (rtac (nonempty RS Transset_Inter_family RS OrdI) 1); +by (rtac Ord_is_Transset 1); +by (REPEAT (ares_tac ([Ord_contains_Transset,nonempty]@prems) 1 + ORELSE etac InterD 1)); +val Ord_Inter = result(); + +val jmemA::prems = goal Ord.thy + "[| j:A; !!x. x:A ==> Ord(B(x)) |] ==> Ord(INT x:A. B(x))"; +by (rtac (jmemA RS RepFunI RS Ord_Inter) 1); +by (etac RepFunE 1); +by (etac ssubst 1); +by (eresolve_tac prems 1); +val Ord_INT = result(); + + +(*** Natural Deduction rules for Memrel ***) + +goalw Ord.thy [Memrel_def] " : Memrel(A) <-> a:b & a:A & b:A"; +by (fast_tac ZF_cs 1); +val Memrel_iff = result(); + +val prems = goal Ord.thy "[| a: b; a: A; b: A |] ==> : Memrel(A)"; +by (REPEAT (resolve_tac (prems@[conjI, Memrel_iff RS iffD2]) 1)); +val MemrelI = result(); + +val [major,minor] = goal Ord.thy + "[| : Memrel(A); \ +\ [| a: A; b: A; a:b |] ==> P \ +\ |] ==> P"; +by (rtac (major RS (Memrel_iff RS iffD1) RS conjE) 1); +by (etac conjE 1); +by (rtac minor 1); +by (REPEAT (assume_tac 1)); +val MemrelE = result(); + +(*The membership relation (as a set) is well-founded. + Proof idea: show A<=B by applying the foundation axiom to A-B *) +goalw Ord.thy [wf_def] "wf(Memrel(A))"; +by (EVERY1 [rtac (foundation RS disjE RS allI), + etac disjI1, + etac bexE, + rtac (impI RS allI RS bexI RS disjI2), + etac MemrelE, + etac bspec, + REPEAT o assume_tac]); +val wf_Memrel = result(); + +(*** Transfinite induction ***) + +(*Epsilon induction over a transitive set*) +val major::prems = goalw Ord.thy [Transset_def] + "[| i: k; Transset(k); \ +\ !!x.[| x: k; ALL y:x. P(y) |] ==> P(x) \ +\ |] ==> P(i)"; +by (rtac (major RS (wf_Memrel RS wf_induct2)) 1); +by (fast_tac (ZF_cs addEs [MemrelE]) 1); +by (resolve_tac prems 1); +by (assume_tac 1); +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addIs [MemrelI]) 1); +val Transset_induct = result(); + +(*Induction over an ordinal*) +val Ord_induct = Ord_is_Transset RSN (2, Transset_induct); + +(*Induction over the class of ordinals -- a useful corollary of Ord_induct*) +val [major,indhyp] = goal Ord.thy + "[| Ord(i); \ +\ !!x.[| Ord(x); ALL y:x. P(y) |] ==> P(x) \ +\ |] ==> P(i)"; +by (rtac (major RS Ord_succ RS (succI1 RS Ord_induct)) 1); +by (rtac indhyp 1); +by (rtac (major RS Ord_succ RS Ord_in_Ord) 1); +by (REPEAT (assume_tac 1)); +val trans_induct = result(); + +(*Perform induction on i, then prove the Ord(i) subgoal using prems. *) +fun trans_ind_tac a prems i = + EVERY [res_inst_tac [("i",a)] trans_induct i, + rename_last_tac a ["1"] (i+1), + ares_tac prems i]; + + +(*** Fundamental properties of the epsilon ordering (< on ordinals) ***) + +(*Finds contradictions for the following proof*) +val Ord_trans_tac = EVERY' [etac notE, etac Ord_trans, REPEAT o atac]; + +(** Proving that "member" is a linear ordering on the ordinals **) + +val prems = goal Ord.thy + "Ord(i) ==> (ALL j. Ord(j) --> i:j | i=j | j:i)"; +by (trans_ind_tac "i" prems 1); +by (rtac (impI RS allI) 1); +by (trans_ind_tac "j" [] 1); +by (DEPTH_SOLVE (swap_res_tac [disjCI,equalityI,subsetI] 1 + ORELSE ball_tac 1 + ORELSE eresolve_tac [impE,disjE,allE] 1 + ORELSE hyp_subst_tac 1 + ORELSE Ord_trans_tac 1)); +val Ord_linear_lemma = result(); + +val ordi::ordj::prems = goal Ord.thy + "[| Ord(i); Ord(j); i:j ==> P; i=j ==> P; j:i ==> P |] \ +\ ==> P"; +by (rtac (ordi RS Ord_linear_lemma RS spec RS impE) 1); +by (rtac ordj 1); +by (REPEAT (eresolve_tac (prems@[asm_rl,disjE]) 1)); +val Ord_linear = result(); + +val prems = goal Ord.thy + "[| Ord(i); Ord(j); i<=j ==> P; j<=i ==> P |] \ +\ ==> P"; +by (res_inst_tac [("i","i"),("j","j")] Ord_linear 1); +by (DEPTH_SOLVE (ares_tac (prems@[subset_refl]) 1 + ORELSE eresolve_tac [asm_rl,OrdmemD,ssubst] 1)); +val Ord_subset = result(); + +goal Ord.thy "!!i j. [| j<=i; ~ i<=j; Ord(i); Ord(j) |] ==> j:i"; +by (etac Ord_linear 1); +by (REPEAT (ares_tac [subset_refl] 1 + ORELSE eresolve_tac [notE,OrdmemD,ssubst] 1)); +val Ord_member = result(); + +val [prem] = goal Ord.thy "Ord(i) ==> 0: succ(i)"; +by (rtac (empty_subsetI RS Ord_member) 1); +by (fast_tac ZF_cs 1); +by (rtac (prem RS Ord_succ) 1); +by (rtac Ord_0 1); +val Ord_0_mem_succ = result(); + +goal Ord.thy "!!i j. [| Ord(i); Ord(j) |] ==> j:i <-> j<=i & ~(i<=j)"; +by (rtac iffI 1); +by (rtac conjI 1); +by (etac OrdmemD 1); +by (rtac (mem_anti_refl RS notI) 2); +by (etac subsetD 2); +by (REPEAT (eresolve_tac [asm_rl, conjE, Ord_member] 1)); +val Ord_member_iff = result(); + +goal Ord.thy "!!i. Ord(i) ==> 0:i <-> ~ i=0"; +be (Ord_0 RSN (2,Ord_member_iff) RS iff_trans) 1; +by (fast_tac eq_cs 1); +val Ord_0_member_iff = result(); + +(** For ordinals, i: succ(j) means 'less-than or equals' **) + +goal Ord.thy "!!i j. [| j<=i; Ord(i); Ord(j) |] ==> j : succ(i)"; +by (rtac Ord_member 1); +by (REPEAT (ares_tac [Ord_succ] 3)); +by (rtac (mem_anti_refl RS notI) 2); +by (etac subsetD 2); +by (ALLGOALS (fast_tac ZF_cs)); +val member_succI = result(); + +goalw Ord.thy [Transset_def,Ord_def] + "!!i j. [| i : succ(j); Ord(j) |] ==> i<=j"; +by (fast_tac ZF_cs 1); +val member_succD = result(); + +goal Ord.thy "!!i j. [| Ord(i); Ord(j) |] ==> j:succ(i) <-> j<=i"; +by (fast_tac (subset_cs addSEs [member_succI, member_succD]) 1); +val member_succ_iff = result(); + +goal Ord.thy + "!!i j. [| Ord(i); Ord(j) |] ==> i<=succ(j) <-> i<=j | i=succ(j)"; +by (ASM_SIMP_TAC (ZF_ss addrews [member_succ_iff RS iff_sym, Ord_succ]) 1); +by (fast_tac ZF_cs 1); +val subset_succ_iff = result(); + +goal Ord.thy "!!i j. [| i:succ(j); j:k; Ord(k) |] ==> i:k"; +by (fast_tac (ZF_cs addEs [Ord_trans]) 1); +val Ord_trans1 = result(); + +goal Ord.thy "!!i j. [| i:j; j:succ(k); Ord(k) |] ==> i:k"; +by (fast_tac (ZF_cs addEs [Ord_trans]) 1); +val Ord_trans2 = result(); + +goal Ord.thy "!!i jk. [| i:j; j<=k; Ord(k) |] ==> i:k"; +by (fast_tac (ZF_cs addEs [Ord_trans]) 1); +val Ord_transsub2 = result(); + +goal Ord.thy "!!i j. [| i:j; Ord(j) |] ==> succ(i) : succ(j)"; +by (rtac member_succI 1); +by (REPEAT (ares_tac [subsetI,Ord_succ,Ord_in_Ord] 1 + ORELSE eresolve_tac [succE,Ord_trans,ssubst] 1)); +val succ_mem_succI = result(); + +goal Ord.thy "!!i j. [| succ(i) : succ(j); Ord(j) |] ==> i:j"; +by (REPEAT (eresolve_tac [asm_rl, make_elim member_succD, succ_subsetE] 1)); +val succ_mem_succE = result(); + +goal Ord.thy "!!i j. Ord(j) ==> succ(i) : succ(j) <-> i:j"; +by (REPEAT (ares_tac [iffI,succ_mem_succI,succ_mem_succE] 1)); +val succ_mem_succ_iff = result(); + +goal Ord.thy "!!i j. [| i<=j; Ord(i); Ord(j) |] ==> succ(i) <= succ(j)"; +by (rtac (member_succI RS succ_mem_succI RS member_succD) 1); +by (REPEAT (ares_tac [Ord_succ] 1)); +val Ord_succ_mono = result(); + +goal Ord.thy "!!i j k. [| i:k; j:k; Ord(k) |] ==> i Un j : k"; +by (res_inst_tac [("i","i"),("j","j")] Ord_subset 1); +by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 1)); +by (ASM_SIMP_TAC (ZF_ss addrews [subset_Un_iff RS iffD1]) 1); +by (rtac (Un_commute RS ssubst) 1); +by (ASM_SIMP_TAC (ZF_ss addrews [subset_Un_iff RS iffD1]) 1); +val Ord_member_UnI = result(); + +goal Ord.thy "!!i j k. [| i:k; j:k; Ord(k) |] ==> i Int j : k"; +by (res_inst_tac [("i","i"),("j","j")] Ord_subset 1); +by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 1)); +by (ASM_SIMP_TAC (ZF_ss addrews [subset_Int_iff RS iffD1]) 1); +by (rtac (Int_commute RS ssubst) 1); +by (ASM_SIMP_TAC (ZF_ss addrews [subset_Int_iff RS iffD1]) 1); +val Ord_member_IntI = result(); + + +(*** Results about limits ***) + +val prems = goal Ord.thy "[| !!i. i:A ==> Ord(i) |] ==> Ord(Union(A))"; +by (rtac (Ord_is_Transset RS Transset_Union_family RS OrdI) 1); +by (REPEAT (etac UnionE 1 ORELSE ares_tac ([Ord_contains_Transset]@prems) 1)); +val Ord_Union = result(); + +val prems = goal Ord.thy "[| !!x. x:A ==> Ord(B(x)) |] ==> Ord(UN x:A. B(x))"; +by (rtac Ord_Union 1); +by (etac RepFunE 1); +by (etac ssubst 1); +by (eresolve_tac prems 1); +val Ord_UN = result(); + +(*The upper bound must be a successor ordinal -- + consider that (UN i:nat.i)=nat although nat is an upper bound of each i*) +val [ordi,limit] = goal Ord.thy + "[| Ord(i); !!y. y:A ==> f(y): succ(i) |] ==> (UN y:A. f(y)) : succ(i)"; +by (rtac (member_succD RS UN_least RS member_succI) 1); +by (REPEAT (ares_tac [ordi, Ord_UN, ordi RS Ord_succ RS Ord_in_Ord, + limit] 1)); +val sup_least = result(); + +val [jmemi,ordi,limit] = goal Ord.thy + "[| j: i; Ord(i); !!y. y:A ==> f(y): j |] ==> (UN y:A. succ(f(y))) : i"; +by (cut_facts_tac [jmemi RS (ordi RS Ord_in_Ord)] 1); +by (rtac (sup_least RS Ord_trans2) 1); +by (REPEAT (ares_tac [jmemi, ordi, succ_mem_succI, limit] 1)); +val sup_least2 = result(); + +goal Ord.thy "!!i. Ord(i) ==> (UN y:i. succ(y)) = i"; +by (fast_tac (eq_cs addSEs [Ord_trans1]) 1); +val Ord_equality = result(); + +(*Holds for all transitive sets, not just ordinals*) +goal Ord.thy "!!i. Ord(i) ==> Union(i) <= i"; +by (fast_tac (ZF_cs addSEs [Ord_trans]) 1); +val Ord_Union_subset = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/Ord.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Ord.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,19 @@ +(* Title: ZF/ordinal.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Ordinals in Zermelo-Fraenkel Set Theory +*) + +Ord = WF + +consts + Memrel :: "i=>i" + Transset,Ord :: "i=>o" + +rules + Memrel_def "Memrel(A) == {z: A*A . EX x y. z= & x:y }" + Transset_def "Transset(i) == ALL x:i. x<=i" + Ord_def "Ord(i) == Transset(i) & (ALL x:i. Transset(x))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/Pair.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Pair.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,153 @@ +(* Title: ZF/pair + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Ordered pairs in Zermelo-Fraenkel Set Theory +*) + +(** Lemmas for showing that uniquely determines a and b **) + +val doubleton_iff = prove_goal ZF.thy + "{a,b} = {c,d} <-> (a=c & b=d) | (a=d & b=c)" + (fn _=> [ (resolve_tac [extension RS iff_trans] 1), + (fast_tac upair_cs 1) ]); + +val Pair_iff = prove_goalw ZF.thy [Pair_def] + " = <-> a=c & b=d" + (fn _=> [ (SIMP_TAC (FOL_ss addrews [doubleton_iff]) 1), + (fast_tac FOL_cs 1) ]); + +val Pair_inject = standard (Pair_iff RS iffD1 RS conjE); + +val Pair_inject1 = prove_goal ZF.thy " = ==> a=c" + (fn [major]=> + [ (rtac (major RS Pair_inject) 1), (assume_tac 1) ]); + +val Pair_inject2 = prove_goal ZF.thy " = ==> b=d" + (fn [major]=> + [ (rtac (major RS Pair_inject) 1), (assume_tac 1) ]); + +val Pair_neq_0 = prove_goalw ZF.thy [Pair_def] "=0 ==> P" + (fn [major]=> + [ (rtac (major RS equalityD1 RS subsetD RS emptyE) 1), + (rtac consI1 1) ]); + +val Pair_neq_fst = prove_goalw ZF.thy [Pair_def] "=a ==> P" + (fn [major]=> + [ (rtac (consI1 RS mem_anti_sym RS FalseE) 1), + (rtac (major RS subst) 1), + (rtac consI1 1) ]); + +val Pair_neq_snd = prove_goalw ZF.thy [Pair_def] "=b ==> P" + (fn [major]=> + [ (rtac (consI1 RS consI2 RS mem_anti_sym RS FalseE) 1), + (rtac (major RS subst) 1), + (rtac (consI1 RS consI2) 1) ]); + + +(*** Sigma: Disjoint union of a family of sets + Generalizes Cartesian product ***) + +val SigmaI = prove_goalw ZF.thy [Sigma_def] + "[| a:A; b:B(a) |] ==> : Sigma(A,B)" + (fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]); + +(*The general elimination rule*) +val SigmaE = prove_goalw ZF.thy [Sigma_def] + "[| c: Sigma(A,B); \ +\ !!x y.[| x:A; y:B(x); c= |] ==> P \ +\ |] ==> P" + (fn major::prems=> + [ (cut_facts_tac [major] 1), + (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]); + +(** Elimination of :A*B -- introduces no eigenvariables **) +val SigmaD1 = prove_goal ZF.thy " : Sigma(A,B) ==> a : A" + (fn [major]=> + [ (rtac (major RS SigmaE) 1), + (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); + +val SigmaD2 = prove_goal ZF.thy " : Sigma(A,B) ==> b : B(a)" + (fn [major]=> + [ (rtac (major RS SigmaE) 1), + (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); + +(*Also provable via + rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac) + THEN prune_params_tac) + (read_instantiate [("c","")] SigmaE); *) +val SigmaE2 = prove_goal ZF.thy + "[| : Sigma(A,B); \ +\ [| a:A; b:B(a) |] ==> P \ +\ |] ==> P" + (fn [major,minor]=> + [ (rtac minor 1), + (rtac (major RS SigmaD1) 1), + (rtac (major RS SigmaD2) 1) ]); + +val Sigma_cong = prove_goalw ZF.thy [Sigma_def] + "[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> \ +\ Sigma(A,B) = Sigma(A',B')" + (fn prems=> [ (prove_cong_tac (prems@[RepFun_cong]) 1) ]); + +val Sigma_empty1 = prove_goal ZF.thy "Sigma(0,B) = 0" + (fn _ => [ (fast_tac (lemmas_cs addIs [equalityI] addSEs [SigmaE]) 1) ]); + +val Sigma_empty2 = prove_goal ZF.thy "A*0 = 0" + (fn _ => [ (fast_tac (lemmas_cs addIs [equalityI] addSEs [SigmaE]) 1) ]); + + +(*** Eliminator - split ***) + +val split = prove_goalw ZF.thy [split_def] + "split(%x y.c(x,y), ) = c(a,b)" + (fn _ => + [ (fast_tac (upair_cs addIs [the_equality] addEs [Pair_inject]) 1) ]); + +val split_type = prove_goal ZF.thy + "[| p:Sigma(A,B); \ +\ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C() \ +\ |] ==> split(%x y.c(x,y), p) : C(p)" + (fn major::prems=> + [ (rtac (major RS SigmaE) 1), + (etac ssubst 1), + (REPEAT (ares_tac (prems @ [split RS ssubst]) 1)) ]); + +(*This congruence rule uses NO typing information...*) +val split_cong = prove_goalw ZF.thy [split_def] + "[| p=p'; !!x y.c(x,y) = c'(x,y) |] ==> \ +\ split(%x y.c(x,y), p) = split(%x y.c'(x,y), p')" + (fn prems=> [ (prove_cong_tac (prems@[the_cong]) 1) ]); + + +(*** conversions for fst and snd ***) + +val fst_conv = prove_goalw ZF.thy [fst_def] "fst() = a" + (fn _=> [ (rtac split 1) ]); + +val snd_conv = prove_goalw ZF.thy [snd_def] "snd() = b" + (fn _=> [ (rtac split 1) ]); + + +(*** split for predicates: result type o ***) + +goalw ZF.thy [fsplit_def] "!!R a b. R(a,b) ==> fsplit(R, )"; +by (REPEAT (ares_tac [refl,exI,conjI] 1)); +val fsplitI = result(); + +val major::prems = goalw ZF.thy [fsplit_def] + "[| fsplit(R,z); !!x y. [| z = ; R(x,y) |] ==> P |] ==> P"; +by (cut_facts_tac [major] 1); +by (REPEAT (eresolve_tac (prems@[asm_rl,exE,conjE]) 1)); +val fsplitE = result(); + +goal ZF.thy "!!R a b. fsplit(R,) ==> R(a,b)"; +by (REPEAT (eresolve_tac [asm_rl,fsplitE,Pair_inject,ssubst] 1)); +val fsplitD = result(); + +val pair_cs = upair_cs + addSIs [SigmaI] + addSEs [SigmaE2, SigmaE, Pair_inject, make_elim succ_inject, + Pair_neq_0, sym RS Pair_neq_0, succ_neq_0, sym RS succ_neq_0]; + diff -r 000000000000 -r a5a9c433f639 src/ZF/Perm.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Perm.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,433 @@ +(* Title: ZF/perm.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For perm.thy. The theory underlying permutation groups + -- Composition of relations, the identity relation + -- Injections, surjections, bijections + -- Lemmas for the Schroeder-Bernstein Theorem +*) + +open Perm; + +(** Surjective function space **) + +goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> f: A->B"; +by (etac CollectD1 1); +val surj_is_fun = result(); + +goalw Perm.thy [surj_def] "!!f A B. f : Pi(A,B) ==> f: surj(A,range(f))"; +by (fast_tac (ZF_cs addIs [apply_equality] + addEs [range_of_fun,domain_type]) 1); +val fun_is_surj = result(); + +goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> range(f)=B"; +by (best_tac (ZF_cs addIs [equalityI,apply_Pair] addEs [range_type]) 1); +val surj_range = result(); + + +(** Injective function space **) + +goalw Perm.thy [inj_def] "!!f A B. f: inj(A,B) ==> f: A->B"; +by (etac CollectD1 1); +val inj_is_fun = result(); + +goalw Perm.thy [inj_def] + "!!f A B. [| :f; :f; f: inj(A,B) |] ==> a=c"; +by (REPEAT (eresolve_tac [asm_rl, Pair_mem_PiE, CollectE] 1)); +by (fast_tac ZF_cs 1); +val inj_equality = result(); + +(** Bijections -- simple lemmas but no intro/elim rules -- use unfolding **) + +goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> f: inj(A,B)"; +by (etac IntD1 1); +val bij_is_inj = result(); + +goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> f: surj(A,B)"; +by (etac IntD2 1); +val bij_is_surj = result(); + +(* f: bij(A,B) ==> f: A->B *) +val bij_is_fun = standard (bij_is_inj RS inj_is_fun); + +(** Identity function **) + +val [prem] = goalw Perm.thy [id_def] "a:A ==> : id(A)"; +by (rtac (prem RS lamI) 1); +val idI = result(); + +val major::prems = goalw Perm.thy [id_def] + "[| p: id(A); !!x.[| x:A; p= |] ==> P \ +\ |] ==> P"; +by (rtac (major RS lamE) 1); +by (REPEAT (ares_tac prems 1)); +val idE = result(); + +goalw Perm.thy [id_def] "id(A) : A->A"; +by (rtac lam_type 1); +by (assume_tac 1); +val id_type = result(); + +val [prem] = goalw Perm.thy [id_def] "A<=B ==> id(A) <= id(B)"; +by (rtac (prem RS lam_mono) 1); +val id_mono = result(); + +goalw Perm.thy [inj_def,id_def] "id(A): inj(A,A)"; +by (REPEAT (ares_tac [CollectI,lam_type] 1)); +by (SIMP_TAC ZF_ss 1); +val id_inj = result(); + +goalw Perm.thy [id_def,surj_def] "id(A): surj(A,A)"; +by (fast_tac (ZF_cs addIs [lam_type,beta]) 1); +val id_surj = result(); + +goalw Perm.thy [bij_def] "id(A): bij(A,A)"; +by (fast_tac (ZF_cs addIs [id_inj,id_surj]) 1); +val id_bij = result(); + + +(** Converse of a relation **) + +val [prem] = goal Perm.thy "f: inj(A,B) ==> converse(f) : range(f)->A"; +by (rtac (prem RS inj_is_fun RS PiE) 1); +by (rtac (converse_type RS PiI) 1); +by (fast_tac ZF_cs 1); +by (fast_tac (ZF_cs addEs [prem RSN (3,inj_equality)]) 1); +by flexflex_tac; +val inj_converse_fun = result(); + +val prems = goalw Perm.thy [surj_def] + "f: inj(A,B) ==> converse(f): surj(range(f), A)"; +by (fast_tac (ZF_cs addIs (prems@[inj_converse_fun,apply_Pair,apply_equality, + converseI,inj_is_fun])) 1); +val inj_converse_surj = result(); + +(*The premises are equivalent to saying that f is injective...*) +val prems = goal Perm.thy + "[| f: A->B; converse(f): C->A; a: A |] ==> converse(f)`(f`a) = a"; +by (fast_tac (ZF_cs addIs (prems@[apply_Pair,apply_equality,converseI])) 1); +val left_inverse_lemma = result(); + +val prems = goal Perm.thy + "[| f: inj(A,B); a: A |] ==> converse(f)`(f`a) = a"; +by (fast_tac (ZF_cs addIs (prems@ + [left_inverse_lemma,inj_converse_fun,inj_is_fun])) 1); +val left_inverse = result(); + +val prems = goal Perm.thy + "[| f: A->B; converse(f): C->A; b: C |] ==> f`(converse(f)`b) = b"; +by (rtac (apply_Pair RS (converseD RS apply_equality)) 1); +by (REPEAT (resolve_tac prems 1)); +val right_inverse_lemma = result(); + +val prems = goal Perm.thy + "[| f: inj(A,B); b: range(f) |] ==> f`(converse(f)`b) = b"; +by (rtac right_inverse_lemma 1); +by (REPEAT (resolve_tac (prems@ [inj_converse_fun,inj_is_fun]) 1)); +val right_inverse = result(); + +val prems = goal Perm.thy + "f: inj(A,B) ==> converse(f): inj(range(f), A)"; +bw inj_def; (*rewrite subgoal but not prems!!*) +by (cut_facts_tac prems 1); +by (safe_tac ZF_cs); +(*apply f to both sides and simplify using right_inverse + -- could also use etac[subst_context RS box_equals] in this proof *) +by (rtac simp_equals 2); +by (REPEAT (eresolve_tac [inj_converse_fun, right_inverse RS sym, ssubst] 1 + ORELSE ares_tac [refl,rangeI] 1)); +val inj_converse_inj = result(); + +goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> converse(f): bij(B,A)"; +by (etac IntE 1); +by (eresolve_tac [(surj_range RS subst)] 1); +by (rtac IntI 1); +by (etac inj_converse_inj 1); +by (etac inj_converse_surj 1); +val bij_converse_bij = result(); + + +(** Composition of two relations **) + +(*The inductive definition package could derive these theorems for R o S*) + +goalw Perm.thy [comp_def] "!!r s. [| :s; :r |] ==> : r O s"; +by (fast_tac ZF_cs 1); +val compI = result(); + +val prems = goalw Perm.thy [comp_def] + "[| xz : r O s; \ +\ !!x y z. [| xz=; :s; :r |] ==> P \ +\ |] ==> P"; +by (cut_facts_tac prems 1); +by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1)); +val compE = result(); + +val compEpair = + rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac) + THEN prune_params_tac) + (read_instantiate [("xz","")] compE); + +val comp_cs = ZF_cs addIs [compI,idI] addSEs [compE,idE]; + +(** Domain and Range -- see Suppes, section 3.1 **) + +(*Boyer et al., Set Theory in First-Order Logic, JAR 2 (1986), 287-327*) +goal Perm.thy "range(r O s) <= range(r)"; +by (fast_tac comp_cs 1); +val range_comp = result(); + +goal Perm.thy "!!r s. domain(r) <= range(s) ==> range(r O s) = range(r)"; +by (rtac (range_comp RS equalityI) 1); +by (fast_tac comp_cs 1); +val range_comp_eq = result(); + +goal Perm.thy "domain(r O s) <= domain(s)"; +by (fast_tac comp_cs 1); +val domain_comp = result(); + +goal Perm.thy "!!r s. range(s) <= domain(r) ==> domain(r O s) = domain(s)"; +by (rtac (domain_comp RS equalityI) 1); +by (fast_tac comp_cs 1); +val domain_comp_eq = result(); + +(** Other results **) + +goal Perm.thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; +by (fast_tac comp_cs 1); +val comp_mono = result(); + +(*composition preserves relations*) +goal Perm.thy "!!r s. [| s<=A*B; r<=B*C |] ==> (r O s) <= A*C"; +by (fast_tac comp_cs 1); +val comp_rel = result(); + +(*associative law for composition*) +goal Perm.thy "(r O s) O t = r O (s O t)"; +by (fast_tac (comp_cs addIs [equalityI]) 1); +val comp_assoc = result(); + +(*left identity of composition; provable inclusions are + id(A) O r <= r + and [| r<=A*B; B<=C |] ==> r <= id(C) O r *) +goal Perm.thy "!!r A B. r<=A*B ==> id(B) O r = r"; +by (fast_tac (comp_cs addIs [equalityI]) 1); +val left_comp_id = result(); + +(*right identity of composition; provable inclusions are + r O id(A) <= r + and [| r<=A*B; A<=C |] ==> r <= r O id(C) *) +goal Perm.thy "!!r A B. r<=A*B ==> r O id(A) = r"; +by (fast_tac (comp_cs addIs [equalityI]) 1); +val right_comp_id = result(); + + +(** Composition preserves functions, injections, and surjections **) + +goal Perm.thy "!!f g. [| g: A->B; f: B->C |] ==> (f O g) : A->C"; +by (REPEAT (ares_tac [PiI,comp_rel,ex1I,compI] 1 + ORELSE eresolve_tac [fun_is_rel,apply_Pair,apply_type] 1)); +by (fast_tac (comp_cs addDs [apply_equality]) 1); +val comp_func = result(); + +goal Perm.thy "!!f g. [| g: A->B; f: B->C; a:A |] ==> (f O g)`a = f`(g`a)"; +by (REPEAT (ares_tac [comp_func,apply_equality,compI, + apply_Pair,apply_type] 1)); +val comp_func_apply = result(); + +goalw Perm.thy [inj_def] + "!!f g. [| g: inj(A,B); f: inj(B,C) |] ==> (f O g) : inj(A,C)"; +by (REPEAT (eresolve_tac [bspec RS bspec RS mp, box_equals] 1 + ORELSE step_tac (ZF_cs addSIs [comp_func,apply_type,comp_func_apply]) 1)); +val comp_inj = result(); + +goalw Perm.thy [surj_def] + "!!f g. [| g: surj(A,B); f: surj(B,C) |] ==> (f O g) : surj(A,C)"; +by (best_tac (ZF_cs addSIs [comp_func,comp_func_apply]) 1); +val comp_surj = result(); + +goalw Perm.thy [bij_def] + "!!f g. [| g: bij(A,B); f: bij(B,C) |] ==> (f O g) : bij(A,C)"; +by (fast_tac (ZF_cs addIs [comp_inj,comp_surj]) 1); +val comp_bij = result(); + + +(** Dual properties of inj and surj -- useful for proofs from + D Pastre. Automatic theorem proving in set theory. + Artificial Intelligence, 10:1--27, 1978. **) + +goalw Perm.thy [inj_def] + "!!f g. [| (f O g): inj(A,C); g: A->B; f: B->C |] ==> g: inj(A,B)"; +by (safe_tac comp_cs); +by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1)); +by (ASM_SIMP_TAC (ZF_ss addrews [comp_func_apply]) 1); +val comp_mem_injD1 = result(); + +goalw Perm.thy [inj_def,surj_def] + "!!f g. [| (f O g): inj(A,C); g: surj(A,B); f: B->C |] ==> f: inj(B,C)"; +by (safe_tac comp_cs); +by (res_inst_tac [("x1", "x")] (bspec RS bexE) 1); +by (eres_inst_tac [("x1", "w")] (bspec RS bexE) 3); +by (REPEAT (assume_tac 1)); +by (safe_tac (comp_cs addSIs ZF_congs)); +by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1)); +by (ASM_SIMP_TAC (ZF_ss addrews [comp_func_apply]) 1); +val comp_mem_injD2 = result(); + +goalw Perm.thy [surj_def] + "!!f g. [| (f O g): surj(A,C); g: A->B; f: B->C |] ==> f: surj(B,C)"; +by (fast_tac (comp_cs addSIs [comp_func_apply RS sym, apply_type]) 1); +val comp_mem_surjD1 = result(); + +goal Perm.thy + "!!f g. [| (f O g)`a = c; g: A->B; f: B->C; a:A |] ==> f`(g`a) = c"; +by (REPEAT (ares_tac [comp_func_apply RS sym RS trans] 1)); +val comp_func_applyD = result(); + +goalw Perm.thy [inj_def,surj_def] + "!!f g. [| (f O g): surj(A,C); g: A->B; f: inj(B,C) |] ==> g: surj(A,B)"; +by (safe_tac comp_cs); +by (eres_inst_tac [("x1", "f`y")] (bspec RS bexE) 1); +by (REPEAT (ares_tac [apply_type] 1 ORELSE dtac comp_func_applyD 1)); +by (best_tac (comp_cs addSIs [apply_type]) 1); +val comp_mem_surjD2 = result(); + + +(** inverses of composition **) + +(*left inverse of composition; one inclusion is + f: A->B ==> id(A) <= converse(f) O f *) +val [prem] = goal Perm.thy + "f: inj(A,B) ==> converse(f) O f = id(A)"; +val injfD = prem RSN (3,inj_equality); +by (cut_facts_tac [prem RS inj_is_fun] 1); +by (fast_tac (comp_cs addIs [equalityI,apply_Pair] + addEs [domain_type, make_elim injfD]) 1); +val left_comp_inverse = result(); + +(*right inverse of composition; one inclusion is + f: A->B ==> f O converse(f) <= id(B) *) +val [prem] = goalw Perm.thy [surj_def] + "f: surj(A,B) ==> f O converse(f) = id(B)"; +val appfD = (prem RS CollectD1) RSN (3,apply_equality2); +by (cut_facts_tac [prem] 1); +by (rtac equalityI 1); +by (best_tac (comp_cs addEs [domain_type, range_type, make_elim appfD]) 1); +by (best_tac (comp_cs addIs [apply_Pair]) 1); +val right_comp_inverse = result(); + +(*Injective case applies converse(f) to both sides then simplifies + using left_inverse_lemma*) +goalw Perm.thy [bij_def,inj_def,surj_def] + "!!f A B. [| converse(f): B->A; f: A->B |] ==> f : bij(A,B)"; +val cf_cong = read_instantiate_sg (sign_of Perm.thy) + [("t","%x.?f`x")] subst_context; +by (fast_tac (ZF_cs addIs [left_inverse_lemma, right_inverse_lemma, apply_type] + addEs [cf_cong RS box_equals]) 1); +val invertible_imp_bijective = result(); + +(** Unions of functions -- cf similar theorems on func.ML **) + +goal Perm.thy "converse(r Un s) = converse(r) Un converse(s)"; +by (rtac equalityI 1); +by (DEPTH_SOLVE_1 + (resolve_tac [Un_least,converse_mono, Un_upper1,Un_upper2] 2)); +by (fast_tac ZF_cs 1); +val converse_of_Un = result(); + +goalw Perm.thy [surj_def] + "!!f g. [| f: surj(A,B); g: surj(C,D); A Int C = 0 |] ==> \ +\ (f Un g) : surj(A Un C, B Un D)"; +by (DEPTH_SOLVE_1 (eresolve_tac [fun_disjoint_apply1, fun_disjoint_apply2] 1 + ORELSE ball_tac 1 + ORELSE (rtac trans 1 THEN atac 2) + ORELSE step_tac (ZF_cs addIs [fun_disjoint_Un]) 1)); +val surj_disjoint_Un = result(); + +(*A simple, high-level proof; the version for injections follows from it, + using f:inj(A,B)<->f:bij(A,range(f)) *) +goal Perm.thy + "!!f g. [| f: bij(A,B); g: bij(C,D); A Int C = 0; B Int D = 0 |] ==> \ +\ (f Un g) : bij(A Un C, B Un D)"; +by (rtac invertible_imp_bijective 1); +by (rtac (converse_of_Un RS ssubst) 1); +by (REPEAT (ares_tac [fun_disjoint_Un, bij_is_fun, bij_converse_bij] 1)); +val bij_disjoint_Un = result(); + + +(** Restrictions as surjections and bijections *) + +val prems = goalw Perm.thy [surj_def] + "f: Pi(A,B) ==> f: surj(A, f``A)"; +val rls = apply_equality :: (prems RL [apply_Pair,Pi_type]); +by (fast_tac (ZF_cs addIs rls) 1); +val surj_image = result(); + +goal Perm.thy + "!!f. [| f: Pi(C,B); A<=C |] ==> restrict(f,A)``A = f``A"; +by (rtac equalityI 1); +by (SELECT_GOAL (rewtac restrict_def) 2); +by (REPEAT (eresolve_tac [imageE, apply_equality RS subst] 2 + ORELSE ares_tac [subsetI,lamI,imageI] 2)); +by (REPEAT (ares_tac [image_mono,restrict_subset,subset_refl] 1)); +val restrict_image = result(); + +goalw Perm.thy [inj_def] + "!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): inj(C,B)"; +by (safe_tac (ZF_cs addSEs [restrict_type2])); +by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp, subsetD, + box_equals, restrict] 1)); +val restrict_inj = result(); + +val prems = goal Perm.thy + "[| f: Pi(A,B); C<=A |] ==> restrict(f,C): surj(C, f``C)"; +by (rtac (restrict_image RS subst) 1); +by (rtac (restrict_type2 RS surj_image) 3); +by (REPEAT (resolve_tac prems 1)); +val restrict_surj = result(); + +goalw Perm.thy [inj_def,bij_def] + "!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)"; +by (safe_tac ZF_cs); +by (REPEAT (eresolve_tac [bspec RS bspec RS mp, subsetD, + box_equals, restrict] 1 + ORELSE ares_tac [surj_is_fun,restrict_surj] 1)); +val restrict_bij = result(); + + +(*** Lemmas for Ramsey's Theorem ***) + +goalw Perm.thy [inj_def] "!!f. [| f: inj(A,B); B<=D |] ==> f: inj(A,D)"; +by (fast_tac (ZF_cs addSEs [fun_weaken_type]) 1); +val inj_weaken_type = result(); + +val [major] = goal Perm.thy + "[| f: inj(succ(m), A) |] ==> restrict(f,m) : inj(m, A-{f`m})"; +by (rtac (major RS restrict_bij RS bij_is_inj RS inj_weaken_type) 1); +by (fast_tac ZF_cs 1); +by (cut_facts_tac [major] 1); +by (rewtac inj_def); +by (safe_tac ZF_cs); +by (etac range_type 1); +by (assume_tac 1); +by (dtac apply_equality 1); +by (assume_tac 1); +by (res_inst_tac [("a","m")] mem_anti_refl 1); +by (fast_tac ZF_cs 1); +val inj_succ_restrict = result(); + +goalw Perm.thy [inj_def] + "!!f. [| f: inj(A,B); ~ a:A; ~ b:B |] ==> \ +\ cons(,f) : inj(cons(a,A), cons(b,B))"; +(*cannot use safe_tac: must preserve the implication*) +by (etac CollectE 1); +by (rtac CollectI 1); +by (etac fun_extend 1); +by (REPEAT (ares_tac [ballI] 1)); +by (REPEAT_FIRST (eresolve_tac [consE,ssubst])); +(*Assumption ALL w:A. ALL x:A. f`w = f`x --> w=x makes ASM_SIMP_TAC loop!*) +by (ALLGOALS (SIMP_TAC (ZF_ss addrews [fun_extend_apply2,fun_extend_apply1]))); +by (ALLGOALS (fast_tac (ZF_cs addIs [apply_type]))); +val inj_extend = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/Perm.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Perm.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,36 @@ +(* Title: ZF/perm + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +The theory underlying permutation groups + -- Composition of relations, the identity relation + -- Injections, surjections, bijections + -- Lemmas for the Schroeder-Bernstein Theorem +*) + +Perm = ZF + +consts + O :: "[i,i]=>i" (infixr 60) + id :: "i=>i" + inj,surj,bij:: "[i,i]=>i" + +rules + + (*composition of relations and functions; NOT Suppes's relative product*) + comp_def "r O s == {xz : domain(s)*range(r) . \ +\ EX x y z. xz= & :s & :r}" + + (*the identity function for A*) + id_def "id(A) == (lam x:A. x)" + + (*one-to-one functions from A to B*) + inj_def "inj(A,B) == { f: A->B. ALL w:A. ALL x:A. f`w=f`x --> w=x}" + + (*onto functions from A to B*) + surj_def "surj(A,B) == { f: A->B . ALL y:B. EX x:A. f`x=y}" + + (*one-to-one and onto functions*) + bij_def "bij(A,B) == inj(A,B) Int surj(A,B)" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/QPair.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/QPair.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,299 @@ +(* Title: ZF/qpair.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For qpair.thy. + +Quine-inspired ordered pairs and disjoint sums, for non-well-founded data +structures in ZF. Does not precisely follow Quine's construction. Thanks +to Thomas Forster for suggesting this approach! + +W. V. Quine, On Ordered Pairs and Relations, in Selected Logic Papers, +1966. + +Many proofs are borrowed from pair.ML and sum.ML + +Do we EVER have rank(a) < rank() ? Perhaps if the latter rank + is not a limit ordinal? +*) + + +open QPair; + +(**** Quine ordered pairing ****) + +(** Lemmas for showing that uniquely determines a and b **) + +val QPair_iff = prove_goalw QPair.thy [QPair_def] + " = <-> a=c & b=d" + (fn _=> [rtac sum_equal_iff 1]); + +val QPair_inject = standard (QPair_iff RS iffD1 RS conjE); + +val QPair_inject1 = prove_goal QPair.thy " = ==> a=c" + (fn [major]=> + [ (rtac (major RS QPair_inject) 1), (assume_tac 1) ]); + +val QPair_inject2 = prove_goal QPair.thy " = ==> b=d" + (fn [major]=> + [ (rtac (major RS QPair_inject) 1), (assume_tac 1) ]); + + +(*** QSigma: Disjoint union of a family of sets + Generalizes Cartesian product ***) + +val QSigmaI = prove_goalw QPair.thy [QSigma_def] + "[| a:A; b:B(a) |] ==> : QSigma(A,B)" + (fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]); + +(*The general elimination rule*) +val QSigmaE = prove_goalw QPair.thy [QSigma_def] + "[| c: QSigma(A,B); \ +\ !!x y.[| x:A; y:B(x); c= |] ==> P \ +\ |] ==> P" + (fn major::prems=> + [ (cut_facts_tac [major] 1), + (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]); + +(** Elimination rules for :A*B -- introducing no eigenvariables **) + +val QSigmaE2 = + rule_by_tactic (REPEAT_FIRST (etac QPair_inject ORELSE' bound_hyp_subst_tac) + THEN prune_params_tac) + (read_instantiate [("c","")] QSigmaE); + +val QSigmaD1 = prove_goal QPair.thy " : QSigma(A,B) ==> a : A" + (fn [major]=> + [ (rtac (major RS QSigmaE2) 1), (assume_tac 1) ]); + +val QSigmaD2 = prove_goal QPair.thy " : QSigma(A,B) ==> b : B(a)" + (fn [major]=> + [ (rtac (major RS QSigmaE2) 1), (assume_tac 1) ]); + +val QSigma_cong = prove_goalw QPair.thy [QSigma_def] + "[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> \ +\ QSigma(A,B) = QSigma(A',B')" + (fn prems=> [ (prove_cong_tac (prems@[RepFun_cong]) 1) ]); + +val QSigma_empty1 = prove_goal QPair.thy "QSigma(0,B) = 0" + (fn _ => [ (fast_tac (ZF_cs addIs [equalityI] addSEs [QSigmaE]) 1) ]); + +val QSigma_empty2 = prove_goal QPair.thy "A <*> 0 = 0" + (fn _ => [ (fast_tac (ZF_cs addIs [equalityI] addSEs [QSigmaE]) 1) ]); + + +(*** Eliminator - qsplit ***) + +val qsplit = prove_goalw QPair.thy [qsplit_def] + "qsplit(%x y.c(x,y), ) = c(a,b)" + (fn _ => [ (fast_tac (ZF_cs addIs [the_equality] addEs [QPair_inject]) 1) ]); + +val qsplit_type = prove_goal QPair.thy + "[| p:QSigma(A,B); \ +\ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C() \ +\ |] ==> qsplit(%x y.c(x,y), p) : C(p)" + (fn major::prems=> + [ (rtac (major RS QSigmaE) 1), + (etac ssubst 1), + (REPEAT (ares_tac (prems @ [qsplit RS ssubst]) 1)) ]); + +(*This congruence rule uses NO typing information...*) +val qsplit_cong = prove_goalw QPair.thy [qsplit_def] + "[| p=p'; !!x y.c(x,y) = c'(x,y) |] ==> \ +\ qsplit(%x y.c(x,y), p) = qsplit(%x y.c'(x,y), p')" + (fn prems=> [ (prove_cong_tac (prems@[the_cong]) 1) ]); + + +val qpair_cs = ZF_cs addSIs [QSigmaI] addSEs [QSigmaE2, QSigmaE, QPair_inject]; + +(*** qconverse ***) + +val qconverseI = prove_goalw QPair.thy [qconverse_def] + "!!a b r. :r ==> :qconverse(r)" + (fn _ => [ (fast_tac qpair_cs 1) ]); + +val qconverseD = prove_goalw QPair.thy [qconverse_def] + "!!a b r. : qconverse(r) ==> : r" + (fn _ => [ (fast_tac qpair_cs 1) ]); + +val qconverseE = prove_goalw QPair.thy [qconverse_def] + "[| yx : qconverse(r); \ +\ !!x y. [| yx=; :r |] ==> P \ +\ |] ==> P" + (fn [major,minor]=> + [ (rtac (major RS ReplaceE) 1), + (REPEAT (eresolve_tac [exE, conjE, minor] 1)), + (hyp_subst_tac 1), + (assume_tac 1) ]); + +val qconverse_cs = qpair_cs addSIs [qconverseI] + addSEs [qconverseD,qconverseE]; + +val qconverse_of_qconverse = prove_goal QPair.thy + "!!A B r. r<=QSigma(A,B) ==> qconverse(qconverse(r)) = r" + (fn _ => [ (fast_tac (qconverse_cs addSIs [equalityI]) 1) ]); + +val qconverse_type = prove_goal QPair.thy + "!!A B r. r <= A <*> B ==> qconverse(r) <= B <*> A" + (fn _ => [ (fast_tac qconverse_cs 1) ]); + +val qconverse_of_prod = prove_goal QPair.thy "qconverse(A <*> B) = B <*> A" + (fn _ => [ (fast_tac (qconverse_cs addSIs [equalityI]) 1) ]); + +val qconverse_empty = prove_goal QPair.thy "qconverse(0) = 0" + (fn _ => [ (fast_tac (qconverse_cs addSIs [equalityI]) 1) ]); + + +(*** qsplit for predicates: result type o ***) + +goalw QPair.thy [qfsplit_def] "!!R a b. R(a,b) ==> qfsplit(R, )"; +by (REPEAT (ares_tac [refl,exI,conjI] 1)); +val qfsplitI = result(); + +val major::prems = goalw QPair.thy [qfsplit_def] + "[| qfsplit(R,z); !!x y. [| z = ; R(x,y) |] ==> P |] ==> P"; +by (cut_facts_tac [major] 1); +by (REPEAT (eresolve_tac (prems@[asm_rl,exE,conjE]) 1)); +val qfsplitE = result(); + +goal QPair.thy "!!R a b. qfsplit(R,) ==> R(a,b)"; +by (REPEAT (eresolve_tac [asm_rl,qfsplitE,QPair_inject,ssubst] 1)); +val qfsplitD = result(); + + +(**** The Quine-inspired notion of disjoint sum ****) + +val qsum_defs = [qsum_def,QInl_def,QInr_def,qcase_def]; + +(** Introduction rules for the injections **) + +goalw QPair.thy qsum_defs "!!a A B. a : A ==> QInl(a) : A <+> B"; +by (REPEAT (ares_tac [UnI1,QSigmaI,singletonI] 1)); +val QInlI = result(); + +goalw QPair.thy qsum_defs "!!b A B. b : B ==> QInr(b) : A <+> B"; +by (REPEAT (ares_tac [UnI2,QSigmaI,singletonI] 1)); +val QInrI = result(); + +(** Elimination rules **) + +val major::prems = goalw QPair.thy qsum_defs + "[| u: A <+> B; \ +\ !!x. [| x:A; u=QInl(x) |] ==> P; \ +\ !!y. [| y:B; u=QInr(y) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS UnE) 1); +by (REPEAT (rtac refl 1 + ORELSE eresolve_tac (prems@[QSigmaE,singletonE,ssubst]) 1)); +val qsumE = result(); + +(** QInjection and freeness rules **) + +val [major] = goalw QPair.thy qsum_defs "QInl(a)=QInl(b) ==> a=b"; +by (EVERY1 [rtac (major RS QPair_inject), assume_tac]); +val QInl_inject = result(); + +val [major] = goalw QPair.thy qsum_defs "QInr(a)=QInr(b) ==> a=b"; +by (EVERY1 [rtac (major RS QPair_inject), assume_tac]); +val QInr_inject = result(); + +val [major] = goalw QPair.thy qsum_defs "QInl(a)=QInr(b) ==> P"; +by (rtac (major RS QPair_inject) 1); +by (etac (sym RS one_neq_0) 1); +val QInl_neq_QInr = result(); + +val QInr_neq_QInl = sym RS QInl_neq_QInr; + +(** Injection and freeness equivalences, for rewriting **) + +goal QPair.thy "QInl(a)=QInl(b) <-> a=b"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [QInl_inject,subst_context] 1)); +val QInl_iff = result(); + +goal QPair.thy "QInr(a)=QInr(b) <-> a=b"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [QInr_inject,subst_context] 1)); +val QInr_iff = result(); + +goal QPair.thy "QInl(a)=QInr(b) <-> False"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [QInl_neq_QInr,FalseE] 1)); +val QInl_QInr_iff = result(); + +goal QPair.thy "QInr(b)=QInl(a) <-> False"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [QInr_neq_QInl,FalseE] 1)); +val QInr_QInl_iff = result(); + +val qsum_cs = + ZF_cs addIs [QInlI,QInrI] addSEs [qsumE,QInl_neq_QInr,QInr_neq_QInl] + addSDs [QInl_inject,QInr_inject]; + +(** <+> is itself injective... who cares?? **) + +goal QPair.thy + "u: A <+> B <-> (EX x. x:A & u=QInl(x)) | (EX y. y:B & u=QInr(y))"; +by (fast_tac qsum_cs 1); +val qsum_iff = result(); + +goal QPair.thy "A <+> B <= C <+> D <-> A<=C & B<=D"; +by (fast_tac qsum_cs 1); +val qsum_subset_iff = result(); + +goal QPair.thy "A <+> B = C <+> D <-> A=C & B=D"; +by (SIMP_TAC (ZF_ss addrews [extension,qsum_subset_iff]) 1); +by (fast_tac ZF_cs 1); +val qsum_equal_iff = result(); + +(*** Eliminator -- qcase ***) + +goalw QPair.thy qsum_defs "qcase(c, d, QInl(a)) = c(a)"; +by (rtac (qsplit RS trans) 1); +by (rtac cond_0 1); +val qcase_QInl = result(); + +goalw QPair.thy qsum_defs "qcase(c, d, QInr(b)) = d(b)"; +by (rtac (qsplit RS trans) 1); +by (rtac cond_1 1); +val qcase_QInr = result(); + +val prems = goalw QPair.thy [qcase_def] + "[| u=u'; !!x. c(x)=c'(x); !!y. d(y)=d'(y) |] ==> \ +\ qcase(c,d,u)=qcase(c',d',u')"; +by (REPEAT (resolve_tac ([refl,qsplit_cong,cond_cong] @ prems) 1)); +val qcase_cong = result(); + +val major::prems = goal QPair.thy + "[| u: A <+> B; \ +\ !!x. x: A ==> c(x): C(QInl(x)); \ +\ !!y. y: B ==> d(y): C(QInr(y)) \ +\ |] ==> qcase(c,d,u) : C(u)"; +by (rtac (major RS qsumE) 1); +by (ALLGOALS (etac ssubst)); +by (ALLGOALS (ASM_SIMP_TAC (ZF_ss addrews + (prems@[qcase_QInl,qcase_QInr])))); +val qcase_type = result(); + +(** Rules for the Part primitive **) + +goal QPair.thy "Part(A <+> B,QInl) = {QInl(x). x: A}"; +by (fast_tac (qsum_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_QInl = result(); + +goal QPair.thy "Part(A <+> B,QInr) = {QInr(y). y: B}"; +by (fast_tac (qsum_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_QInr = result(); + +goal QPair.thy "Part(A <+> B, %x.QInr(h(x))) = {QInr(y). y: Part(B,h)}"; +by (fast_tac (qsum_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_QInr2 = result(); + +goal QPair.thy "!!A B C. C <= A <+> B ==> Part(C,QInl) Un Part(C,QInr) = C"; +by (rtac equalityI 1); +by (rtac Un_least 1); +by (rtac Part_subset 1); +by (rtac Part_subset 1); +by (fast_tac (ZF_cs addIs [PartI] addSEs [qsumE]) 1); +val Part_qsum_equality = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/QPair.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/QPair.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,52 @@ +(* Title: ZF/qpair.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Quine-inspired ordered pairs and disjoint sums, for non-well-founded data +structures in ZF. Does not precisely follow Quine's construction. Thanks +to Thomas Forster for suggesting this approach! + +W. V. Quine, On Ordered Pairs and Relations, in Selected Logic Papers, +1966. +*) + +QPair = Sum + +consts + QPair :: "[i, i] => i" ("<(_;/ _)>") + qsplit :: "[[i,i] => i, i] => i" + qfsplit :: "[[i,i] => o, i] => o" + qconverse :: "i => i" + "@QSUM" :: "[idt, i, i] => i" ("(3QSUM _:_./ _)" 10) + " <*>" :: "[i, i] => i" ("(_ <*>/ _)" [81, 80] 80) + QSigma :: "[i, i => i] => i" + + "<+>" :: "[i,i]=>i" (infixr 65) + QInl,QInr :: "i=>i" + qcase :: "[i=>i, i=>i, i]=>i" + +translations + "QSUM x:A. B" => "QSigma(A, %x. B)" + +rules + QPair_def " == a+b" + qsplit_def "qsplit(c,p) == THE y. EX a b. p= & y=c(a,b)" + qfsplit_def "qfsplit(R,z) == EX x y. z= & R(x,y)" + qconverse_def "qconverse(r) == {z. w:r, EX x y. w= & z=}" + QSigma_def "QSigma(A,B) == UN x:A. UN y:B(x). {}" + + qsum_def "A <+> B == QSigma({0}, %x.A) Un QSigma({1}, %x.B)" + QInl_def "QInl(a) == <0;a>" + QInr_def "QInr(b) == <1;b>" + qcase_def "qcase(c,d) == qsplit(%y z. cond(y, d(z), c(z)))" +end + +ML + +(* 'Dependent' type operators *) + +val parse_translation = + [(" <*>", ndependent_tr "QSigma")]; + +val print_translation = + [("QSigma", dependent_tr' ("@QSUM", " <*>"))]; diff -r 000000000000 -r a5a9c433f639 src/ZF/QUniv.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/QUniv.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,334 @@ +(* Title: ZF/quniv + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For quniv.thy. A small universe for lazy recursive types +*) + +open QUniv; + +(** Introduction and elimination rules avoid tiresome folding/unfolding **) + +goalw QUniv.thy [quniv_def] + "!!X A. X <= univ(eclose(A)) ==> X : quniv(A)"; +be PowI 1; +val qunivI = result(); + +goalw QUniv.thy [quniv_def] + "!!X A. X : quniv(A) ==> X <= univ(eclose(A))"; +be PowD 1; +val qunivD = result(); + +goalw QUniv.thy [quniv_def] "!!A B. A<=B ==> quniv(A) <= quniv(B)"; +by (etac (eclose_mono RS univ_mono RS Pow_mono) 1); +val quniv_mono = result(); + +(*** Closure properties ***) + +goalw QUniv.thy [quniv_def] "univ(eclose(A)) <= quniv(A)"; +by (rtac (Transset_iff_Pow RS iffD1) 1); +by (rtac (Transset_eclose RS Transset_univ) 1); +val univ_eclose_subset_quniv = result(); + +goal QUniv.thy "univ(A) <= quniv(A)"; +by (rtac (arg_subset_eclose RS univ_mono RS subset_trans) 1); +by (rtac univ_eclose_subset_quniv 1); +val univ_subset_quniv = result(); + +val univ_into_quniv = standard (univ_subset_quniv RS subsetD); + +goalw QUniv.thy [quniv_def] "Pow(univ(A)) <= quniv(A)"; +by (rtac (arg_subset_eclose RS univ_mono RS Pow_mono) 1); +val Pow_univ_subset_quniv = result(); + +val univ_subset_into_quniv = standard + (PowI RS (Pow_univ_subset_quniv RS subsetD)); + +val zero_in_quniv = standard (zero_in_univ RS univ_into_quniv); +val one_in_quniv = standard (one_in_univ RS univ_into_quniv); +val two_in_quniv = standard (two_in_univ RS univ_into_quniv); + +val A_subset_quniv = standard + ([A_subset_univ, univ_subset_quniv] MRS subset_trans); + +val A_into_quniv = A_subset_quniv RS subsetD; + +(*** univ(A) closure for Quine-inspired pairs and injections ***) + +(*Quine ordered pairs*) +goalw QUniv.thy [QPair_def] + "!!A a. [| a <= univ(A); b <= univ(A) |] ==> <= univ(A)"; +by (REPEAT (ares_tac [sum_subset_univ] 1)); +val QPair_subset_univ = result(); + +(** Quine disjoint sum **) + +goalw QUniv.thy [QInl_def] "!!A a. a <= univ(A) ==> QInl(a) <= univ(A)"; +by (etac (empty_subsetI RS QPair_subset_univ) 1); +val QInl_subset_univ = result(); + +val naturals_subset_nat = + rewrite_rule [Transset_def] (Ord_nat RS Ord_is_Transset) + RS bspec; + +val naturals_subset_univ = + [naturals_subset_nat, nat_subset_univ] MRS subset_trans; + +goalw QUniv.thy [QInr_def] "!!A a. a <= univ(A) ==> QInr(a) <= univ(A)"; +by (etac (nat_1I RS naturals_subset_univ RS QPair_subset_univ) 1); +val QInr_subset_univ = result(); + +(*** Closure for Quine-inspired products and sums ***) + +(*Quine ordered pairs*) +goalw QUniv.thy [quniv_def,QPair_def] + "!!A a. [| a: quniv(A); b: quniv(A) |] ==> : quniv(A)"; +by (REPEAT (dtac PowD 1)); +by (REPEAT (ares_tac [PowI, sum_subset_univ] 1)); +val QPair_in_quniv = result(); + +goal QUniv.thy "quniv(A) <*> quniv(A) <= quniv(A)"; +by (REPEAT (ares_tac [subsetI, QPair_in_quniv] 1 + ORELSE eresolve_tac [QSigmaE, ssubst] 1)); +val QSigma_quniv = result(); + +val QSigma_subset_quniv = standard + (QSigma_mono RS (QSigma_quniv RSN (2,subset_trans))); + +(*The opposite inclusion*) +goalw QUniv.thy [quniv_def,QPair_def] + "!!A a b. : quniv(A) ==> a: quniv(A) & b: quniv(A)"; +be ([Transset_eclose RS Transset_univ, PowD] MRS + Transset_includes_summands RS conjE) 1; +by (REPEAT (ares_tac [conjI,PowI] 1)); +val quniv_QPair_D = result(); + +val quniv_QPair_E = standard (quniv_QPair_D RS conjE); + +goal QUniv.thy " : quniv(A) <-> a: quniv(A) & b: quniv(A)"; +by (REPEAT (ares_tac [iffI, QPair_in_quniv, quniv_QPair_D] 1 + ORELSE etac conjE 1)); +val quniv_QPair_iff = result(); + + +(** Quine disjoint sum **) + +goalw QUniv.thy [QInl_def] "!!A a. a: quniv(A) ==> QInl(a) : quniv(A)"; +by (REPEAT (ares_tac [zero_in_quniv,QPair_in_quniv] 1)); +val QInl_in_quniv = result(); + +goalw QUniv.thy [QInr_def] "!!A b. b: quniv(A) ==> QInr(b) : quniv(A)"; +by (REPEAT (ares_tac [one_in_quniv, QPair_in_quniv] 1)); +val QInr_in_quniv = result(); + +goal QUniv.thy "quniv(C) <+> quniv(C) <= quniv(C)"; +by (REPEAT (ares_tac [subsetI, QInl_in_quniv, QInr_in_quniv] 1 + ORELSE eresolve_tac [qsumE, ssubst] 1)); +val qsum_quniv = result(); + +val qsum_subset_quniv = standard + (qsum_mono RS (qsum_quniv RSN (2,subset_trans))); + +(*** The natural numbers ***) + +val nat_subset_quniv = standard + ([nat_subset_univ, univ_subset_quniv] MRS subset_trans); + +(* n:nat ==> n:quniv(A) *) +val nat_into_quniv = standard (nat_subset_quniv RS subsetD); + +val bool_subset_quniv = standard + ([bool_subset_univ, univ_subset_quniv] MRS subset_trans); + +val bool_into_quniv = standard (bool_subset_quniv RS subsetD); + + +(**** Properties of Vfrom analogous to the "take-lemma" ****) + +(*** Intersecting a*b with Vfrom... ***) + +(*This version says a, b exist one level down, in the smaller set Vfrom(X,i)*) +goal Univ.thy + "!!X. [| {a,b} : Vfrom(X,succ(i)); Transset(X) |] ==> \ +\ a: Vfrom(X,i) & b: Vfrom(X,i)"; +bd (Transset_Vfrom_succ RS equalityD1 RS subsetD RS PowD) 1; +ba 1; +by (fast_tac ZF_cs 1); +val doubleton_in_Vfrom_D = result(); + +(*This weaker version says a, b exist at the same level*) +val Vfrom_doubleton_D = standard (Transset_Vfrom RS Transset_doubleton_D); + +(** Using only the weaker theorem would prove : Vfrom(X,i) + implies a, b : Vfrom(X,i), which is useless for induction. + Using only the stronger theorem would prove : Vfrom(X,succ(succ(i))) + implies a, b : Vfrom(X,i), leaving the succ(i) case untreated. + The combination gives a reduction by precisely one level, which is + most convenient for proofs. +**) + +goalw Univ.thy [Pair_def] + "!!X. [| : Vfrom(X,succ(i)); Transset(X) |] ==> \ +\ a: Vfrom(X,i) & b: Vfrom(X,i)"; +by (fast_tac (ZF_cs addSDs [doubleton_in_Vfrom_D, Vfrom_doubleton_D]) 1); +val Pair_in_Vfrom_D = result(); + +goal Univ.thy + "!!X. Transset(X) ==> \ +\ (a*b) Int Vfrom(X, succ(i)) <= (a Int Vfrom(X,i)) * (b Int Vfrom(X,i))"; +by (fast_tac (ZF_cs addSDs [Pair_in_Vfrom_D]) 1); +val product_Int_Vfrom_subset = result(); + +(*** Intersecting with Vfrom... ***) + +goalw QUniv.thy [QPair_def,sum_def] + "!!X. Transset(X) ==> \ +\ Int Vfrom(X, succ(i)) <= "; +br (Int_Un_distrib RS ssubst) 1; +br Un_mono 1; +by (REPEAT (ares_tac [product_Int_Vfrom_subset RS subset_trans, + [Int_lower1, subset_refl] MRS Sigma_mono] 1)); +val QPair_Int_Vfrom_succ_subset = result(); + +(** Pairs in quniv -- for handling the base case **) + +goal QUniv.thy "!!X. : quniv(X) ==> : univ(eclose(X))"; +be ([qunivD, Transset_eclose] MRS Transset_Pair_subset_univ) 1; +val Pair_in_quniv_D = result(); + +goal QUniv.thy "a*b Int quniv(A) = a*b Int univ(eclose(A))"; +br equalityI 1; +br ([subset_refl, univ_eclose_subset_quniv] MRS Int_mono) 2; +by (fast_tac (ZF_cs addSEs [Pair_in_quniv_D]) 1); +val product_Int_quniv_eq = result(); + +goalw QUniv.thy [QPair_def,sum_def] + " Int quniv(A) = Int univ(eclose(A))"; +by (SIMP_TAC (ZF_ss addrews [Int_Un_distrib, product_Int_quniv_eq]) 1); +val QPair_Int_quniv_eq = result(); + +(**** "Take-lemma" rules for proving c: quniv(A) ****) + +goalw QUniv.thy [quniv_def] "Transset(quniv(A))"; +br (Transset_eclose RS Transset_univ RS Transset_Pow) 1; +val Transset_quniv = result(); + +val [aprem, iprem] = goal QUniv.thy + "[| a: quniv(quniv(X)); \ +\ !!i. i:nat ==> a Int Vfrom(quniv(X),i) : quniv(A) \ +\ |] ==> a : quniv(A)"; +br (univ_Int_Vfrom_subset RS qunivI) 1; +br (aprem RS qunivD) 1; +by (rtac (Transset_quniv RS Transset_eclose_eq_arg RS ssubst) 1); +be (iprem RS qunivD) 1; +val quniv_Int_Vfrom = result(); + +(** Rules for level 0 **) + +goal QUniv.thy " Int quniv(A) : quniv(A)"; +br (QPair_Int_quniv_eq RS ssubst) 1; +br (Int_lower2 RS qunivI) 1; +val QPair_Int_quniv_in_quniv = result(); + +(*Unused; kept as an example. QInr rule is similar*) +goalw QUniv.thy [QInl_def] "QInl(a) Int quniv(A) : quniv(A)"; +br QPair_Int_quniv_in_quniv 1; +val QInl_Int_quniv_in_quniv = result(); + +goal QUniv.thy "!!a A X. a : quniv(A) ==> a Int X : quniv(A)"; +be ([Int_lower1, qunivD] MRS subset_trans RS qunivI) 1; +val Int_quniv_in_quniv = result(); + +goal QUniv.thy + "!!X. a Int X : quniv(A) ==> a Int Vfrom(X, 0) : quniv(A)"; +by (etac (Vfrom_0 RS ssubst) 1); +val Int_Vfrom_0_in_quniv = result(); + +(** Rules for level succ(i), decreasing it to i **) + +goal QUniv.thy + "!!X. [| a Int Vfrom(X,i) : quniv(A); \ +\ b Int Vfrom(X,i) : quniv(A); \ +\ Transset(X) \ +\ |] ==> Int Vfrom(X, succ(i)) : quniv(A)"; +br (QPair_Int_Vfrom_succ_subset RS subset_trans RS qunivI) 1; +br (QPair_in_quniv RS qunivD) 2; +by (REPEAT (assume_tac 1)); +val QPair_Int_Vfrom_succ_in_quniv = result(); + +val zero_Int_in_quniv = standard + ([Int_lower1,empty_subsetI] MRS subset_trans RS qunivI); + +val one_Int_in_quniv = standard + ([Int_lower1, one_in_quniv RS qunivD] MRS subset_trans RS qunivI); + +(*Unused; kept as an example. QInr rule is similar*) +goalw QUniv.thy [QInl_def] + "!!X. [| a Int Vfrom(X,i) : quniv(A); Transset(X) \ +\ |] ==> QInl(a) Int Vfrom(X, succ(i)) : quniv(A)"; +br QPair_Int_Vfrom_succ_in_quniv 1; +by (REPEAT (ares_tac [zero_Int_in_quniv] 1)); +val QInl_Int_Vfrom_succ_in_quniv = result(); + +(** Rules for level i -- preserving the level, not decreasing it **) + +goalw QUniv.thy [QPair_def] + "!!X. Transset(X) ==> \ +\ Int Vfrom(X,i) <= "; +be (Transset_Vfrom RS Transset_sum_Int_subset) 1; +val QPair_Int_Vfrom_subset = result(); + +goal QUniv.thy + "!!X. [| a Int Vfrom(X,i) : quniv(A); \ +\ b Int Vfrom(X,i) : quniv(A); \ +\ Transset(X) \ +\ |] ==> Int Vfrom(X,i) : quniv(A)"; +br (QPair_Int_Vfrom_subset RS subset_trans RS qunivI) 1; +br (QPair_in_quniv RS qunivD) 2; +by (REPEAT (assume_tac 1)); +val QPair_Int_Vfrom_in_quniv = result(); + + +(**** "Take-lemma" rules for proving a=b by co-induction ****) + +(** Unfortunately, the technique used above does not apply here, since + the base case appears impossible to prove: it involves an intersection + with eclose(X) for arbitrary X. So a=b is proved by transfinite induction + over ALL ordinals, using Vset(i) instead of Vfrom(X,i). +**) + +(*Rule for level 0*) +goal QUniv.thy "a Int Vset(0) <= b"; +by (rtac (Vfrom_0 RS ssubst) 1); +by (fast_tac ZF_cs 1); +val Int_Vset_0_subset = result(); + +(*Rule for level succ(i), decreasing it to i*) +goal QUniv.thy + "!!i. [| a Int Vset(i) <= c; \ +\ b Int Vset(i) <= d \ +\ |] ==> Int Vset(succ(i)) <= "; +br ([Transset_0 RS QPair_Int_Vfrom_succ_subset, QPair_mono] + MRS subset_trans) 1; +by (REPEAT (assume_tac 1)); +val QPair_Int_Vset_succ_subset_trans = result(); + +(*Unused; kept as an example. QInr rule is similar*) +goalw QUniv.thy [QInl_def] + "!!i. a Int Vset(i) <= b ==> QInl(a) Int Vset(succ(i)) <= QInl(b)"; +be (Int_lower1 RS QPair_Int_Vset_succ_subset_trans) 1; +val QInl_Int_Vset_succ_subset_trans = result(); + +(*Rule for level i -- preserving the level, not decreasing it*) +goal QUniv.thy + "!!i. [| a Int Vset(i) <= c; \ +\ b Int Vset(i) <= d \ +\ |] ==> Int Vset(i) <= "; +br ([Transset_0 RS QPair_Int_Vfrom_subset, QPair_mono] + MRS subset_trans) 1; +by (REPEAT (assume_tac 1)); +val QPair_Int_Vset_subset_trans = result(); + + + diff -r 000000000000 -r a5a9c433f639 src/ZF/QUniv.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/QUniv.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,16 @@ +(* Title: ZF/univ.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +A small universe for lazy recursive types +*) + +QUniv = Univ + QPair + +consts + quniv :: "i=>i" + +rules + quniv_def "quniv(A) == Pow(univ(eclose(A)))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/README Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,22 @@ + ZF: Zermelo-Fraenkel Set Theory + +This directory contains the Standard ML sources of the Isabelle system for +ZF Set Theory. It is loaded upon FOL, not Pure Isabelle. Important files +include + +ROOT.ML -- loads all source files. Enter an ML image containing FOL and +type: use "ROOT.ML"; + +Makefile -- compiles the files under Poly/ML or SML of New Jersey + +ex -- subdirectory containing examples. To execute them, enter an ML image +containing ZF and type: use "ex/ROOT.ML"; + +Useful references on ZF set theory: + + Paul R. Halmos, Naive Set Theory (Van Nostrand, 1960) + + Patrick Suppes, Axiomatic Set Theory (Dover, 1972) + + Keith J. Devlin, + Fundamentals of Contemporary Set Theory (Springer, 1979) diff -r 000000000000 -r a5a9c433f639 src/ZF/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,74 @@ +(* Title: ZF/ROOT + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Adds Zermelo-Fraenkel Set Theory to a database containing First-Order Logic. + +This theory is the work of Martin Coen, Philippe Noel and Lawrence Paulson. +*) + +val banner = "ZF Set Theory (in FOL)"; +writeln banner; + +(*For Pure/drule?? Multiple resolution infixes*) +infix 0 MRS MRL; + +(*Resolve a list of rules against bottom_rl from right to left*) +fun rls MRS bottom_rl = + let fun rs_aux i [] = bottom_rl + | rs_aux i (rl::rls) = rl RSN (i, rs_aux (i+1) rls) + in rs_aux 1 rls end; + +fun rlss MRL bottom_rls = + let fun rs_aux i [] = bottom_rls + | rs_aux i (rls::rlss) = rls RLN (i, rs_aux (i+1) rlss) + in rs_aux 1 rlss end; + +print_depth 1; +use_thy "zf"; + +use "upair.ML"; +use "subset.ML"; +use "pair.ML"; +use "domrange.ML"; +use "func.ML"; +use "equalities.ML"; +use "simpdata.ML"; + +(*further development*) +use_thy "bool"; +use_thy "sum"; +use_thy "qpair"; +use "mono.ML"; +use_thy "fixedpt"; + +(*Inductive/co-inductive definitions*) +use "ind-syntax.ML"; +use "intr-elim.ML"; +use "indrule.ML"; +use "inductive.ML"; +use "co-inductive.ML"; + +use_thy "perm"; +use_thy "trancl"; +use_thy "wf"; +use_thy "ordinal"; +use_thy "nat"; +use_thy "epsilon"; +use_thy "arith"; + +(*Datatype/co-datatype definitions*) +use_thy "univ"; +use_thy "quniv"; +use "constructor.ML"; +use "datatype.ML"; + +use "fin.ML"; +use "list.ML"; +use_thy "list-fn"; + +(*printing functions are inherited from FOL*) +print_depth 8; + +val ZF_build_completed = (); (*indicate successful build*) diff -r 000000000000 -r a5a9c433f639 src/ZF/Sum.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Sum.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,170 @@ +(* Title: ZF/sum + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Disjoint sums in Zermelo-Fraenkel Set Theory +*) + +open Sum; + +val sum_defs = [sum_def,Inl_def,Inr_def,case_def]; + +(** Introduction rules for the injections **) + +goalw Sum.thy sum_defs "!!a A B. a : A ==> Inl(a) : A+B"; +by (REPEAT (ares_tac [UnI1,SigmaI,singletonI] 1)); +val InlI = result(); + +goalw Sum.thy sum_defs "!!b A B. b : B ==> Inr(b) : A+B"; +by (REPEAT (ares_tac [UnI2,SigmaI,singletonI] 1)); +val InrI = result(); + +(** Elimination rules **) + +val major::prems = goalw Sum.thy sum_defs + "[| u: A+B; \ +\ !!x. [| x:A; u=Inl(x) |] ==> P; \ +\ !!y. [| y:B; u=Inr(y) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS UnE) 1); +by (REPEAT (rtac refl 1 + ORELSE eresolve_tac (prems@[SigmaE,singletonE,ssubst]) 1)); +val sumE = result(); + +(** Injection and freeness rules **) + +val [major] = goalw Sum.thy sum_defs "Inl(a)=Inl(b) ==> a=b"; +by (EVERY1 [rtac (major RS Pair_inject), assume_tac]); +val Inl_inject = result(); + +val [major] = goalw Sum.thy sum_defs "Inr(a)=Inr(b) ==> a=b"; +by (EVERY1 [rtac (major RS Pair_inject), assume_tac]); +val Inr_inject = result(); + +val [major] = goalw Sum.thy sum_defs "Inl(a)=Inr(b) ==> P"; +by (rtac (major RS Pair_inject) 1); +by (etac (sym RS one_neq_0) 1); +val Inl_neq_Inr = result(); + +val Inr_neq_Inl = sym RS Inl_neq_Inr; + +(** Injection and freeness equivalences, for rewriting **) + +goal Sum.thy "Inl(a)=Inl(b) <-> a=b"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [Inl_inject,subst_context] 1)); +val Inl_iff = result(); + +goal Sum.thy "Inr(a)=Inr(b) <-> a=b"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [Inr_inject,subst_context] 1)); +val Inr_iff = result(); + +goal Sum.thy "Inl(a)=Inr(b) <-> False"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [Inl_neq_Inr,FalseE] 1)); +val Inl_Inr_iff = result(); + +goal Sum.thy "Inr(b)=Inl(a) <-> False"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [Inr_neq_Inl,FalseE] 1)); +val Inr_Inl_iff = result(); + +val sum_cs = ZF_cs addIs [InlI,InrI] addSEs [sumE,Inl_neq_Inr,Inr_neq_Inl] + addSDs [Inl_inject,Inr_inject]; + +goal Sum.thy "u: A+B <-> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))"; +by (fast_tac sum_cs 1); +val sum_iff = result(); + +goal Sum.thy "A+B <= C+D <-> A<=C & B<=D"; +by (fast_tac sum_cs 1); +val sum_subset_iff = result(); + +goal Sum.thy "A+B = C+D <-> A=C & B=D"; +by (SIMP_TAC (ZF_ss addrews [extension,sum_subset_iff]) 1); +by (fast_tac ZF_cs 1); +val sum_equal_iff = result(); + +(*** Eliminator -- case ***) + +goalw Sum.thy sum_defs "case(c, d, Inl(a)) = c(a)"; +by (rtac (split RS trans) 1); +by (rtac cond_0 1); +val case_Inl = result(); + +goalw Sum.thy sum_defs "case(c, d, Inr(b)) = d(b)"; +by (rtac (split RS trans) 1); +by (rtac cond_1 1); +val case_Inr = result(); + +val prems = goalw Sum.thy [case_def] + "[| u=u'; !!x. c(x)=c'(x); !!y. d(y)=d'(y) |] ==> \ +\ case(c,d,u)=case(c',d',u')"; +by (REPEAT (resolve_tac ([refl,split_cong,cond_cong] @ prems) 1)); +val case_cong = result(); + +val major::prems = goal Sum.thy + "[| u: A+B; \ +\ !!x. x: A ==> c(x): C(Inl(x)); \ +\ !!y. y: B ==> d(y): C(Inr(y)) \ +\ |] ==> case(c,d,u) : C(u)"; +by (rtac (major RS sumE) 1); +by (ALLGOALS (etac ssubst)); +by (ALLGOALS (ASM_SIMP_TAC (ZF_ss addrews + (prems@[case_Inl,case_Inr])))); +val case_type = result(); + +(** Rules for the Part primitive **) + +goalw Sum.thy [Part_def] + "!!a b A h. [| a : A; a=h(b) |] ==> a : Part(A,h)"; +by (REPEAT (ares_tac [exI,CollectI] 1)); +val Part_eqI = result(); + +val PartI = refl RSN (2,Part_eqI); + +val major::prems = goalw Sum.thy [Part_def] + "[| a : Part(A,h); !!z. [| a : A; a=h(z) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS CollectE) 1); +by (etac exE 1); +by (REPEAT (ares_tac prems 1)); +val PartE = result(); + +goalw Sum.thy [Part_def] "Part(A,h) <= A"; +by (rtac Collect_subset 1); +val Part_subset = result(); + +goal Sum.thy "!!A B h. A<=B ==> Part(A,h)<=Part(B,h)"; +by (fast_tac (ZF_cs addIs [PartI] addSEs [PartE]) 1); +val Part_mono = result(); + +goal Sum.thy "Part(A+B,Inl) = {Inl(x). x: A}"; +by (fast_tac (sum_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_Inl = result(); + +goal Sum.thy "Part(A+B,Inr) = {Inr(y). y: B}"; +by (fast_tac (sum_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_Inr = result(); + +goalw Sum.thy [Part_def] "!!a. a : Part(A,h) ==> a : A"; +by (etac CollectD1 1); +val PartD1 = result(); + +goal Sum.thy "Part(A,%x.x) = A"; +by (fast_tac (ZF_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_id = result(); + +goal Sum.thy "Part(A+B, %x.Inr(h(x))) = {Inr(y). y: Part(B,h)}"; +by (fast_tac (sum_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_Inr2 = result(); + +goal Sum.thy "!!A B C. C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C"; +by (rtac equalityI 1); +by (rtac Un_least 1); +by (rtac Part_subset 1); +by (rtac Part_subset 1); +by (fast_tac (ZF_cs addIs [PartI] addSEs [sumE]) 1); +val Part_sum_equality = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/Sum.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Sum.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,25 @@ +(* Title: ZF/sum.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Disjoint sums in Zermelo-Fraenkel Set Theory +"Part" primitive for simultaneous recursive type definitions +*) + +Sum = Bool + +consts + "+" :: "[i,i]=>i" (infixr 65) + Inl,Inr :: "i=>i" + case :: "[i=>i, i=>i, i]=>i" + Part :: "[i,i=>i] => i" + +rules + sum_def "A+B == {0}*A Un {1}*B" + Inl_def "Inl(a) == <0,a>" + Inr_def "Inr(b) == <1,b>" + case_def "case(c,d) == split(%y z. cond(y, d(z), c(z)))" + + (*operator for selecting out the various summands*) + Part_def "Part(A,h) == {x: A. EX z. x = h(z)}" +end diff -r 000000000000 -r a5a9c433f639 src/ZF/Trancl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Trancl.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,193 @@ +(* Title: ZF/trancl.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For trancl.thy. Transitive closure of a relation +*) + +open Trancl; + +val major::prems = goalw Trancl.thy [trans_def] + "[| trans(r); :r; :r |] ==> :r"; +by (rtac (major RS spec RS spec RS spec RS mp RS mp) 1); +by (REPEAT (resolve_tac prems 1)); +val transD = result(); + +goal Trancl.thy "bnd_mono(field(r)*field(r), %s. id(field(r)) Un (r O s))"; +by (rtac bnd_monoI 1); +by (REPEAT (ares_tac [subset_refl, Un_mono, comp_mono] 2)); +by (fast_tac comp_cs 1); +val rtrancl_bnd_mono = result(); + +val [prem] = goalw Trancl.thy [rtrancl_def] "r<=s ==> r^* <= s^*"; +by (rtac lfp_mono 1); +by (REPEAT (resolve_tac [rtrancl_bnd_mono, prem, subset_refl, id_mono, + comp_mono, Un_mono, field_mono, Sigma_mono] 1)); +val rtrancl_mono = result(); + +(* r^* = id(field(r)) Un ( r O r^* ) *) +val rtrancl_unfold = rtrancl_bnd_mono RS (rtrancl_def RS def_lfp_Tarski); + +(** The relation rtrancl **) + +val rtrancl_type = standard (rtrancl_def RS def_lfp_subset); + +(*Reflexivity of rtrancl*) +val [prem] = goal Trancl.thy "[| a: field(r) |] ==> : r^*"; +by (resolve_tac [rtrancl_unfold RS ssubst] 1); +by (rtac (prem RS idI RS UnI1) 1); +val rtrancl_refl = result(); + +(*Closure under composition with r *) +val prems = goal Trancl.thy + "[| : r^*; : r |] ==> : r^*"; +by (resolve_tac [rtrancl_unfold RS ssubst] 1); +by (rtac (compI RS UnI2) 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +val rtrancl_into_rtrancl = result(); + +(*rtrancl of r contains all pairs in r *) +val prems = goal Trancl.thy " : r ==> : r^*"; +by (resolve_tac [rtrancl_refl RS rtrancl_into_rtrancl] 1); +by (REPEAT (resolve_tac (prems@[fieldI1]) 1)); +val r_into_rtrancl = result(); + +(*The premise ensures that r consists entirely of pairs*) +val prems = goal Trancl.thy "r <= Sigma(A,B) ==> r <= r^*"; +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addIs [r_into_rtrancl]) 1); +val r_subset_rtrancl = result(); + +goal Trancl.thy "field(r^*) = field(r)"; +by (fast_tac (eq_cs addIs [r_into_rtrancl] + addSDs [rtrancl_type RS subsetD]) 1); +val rtrancl_field = result(); + + +(** standard induction rule **) + +val major::prems = goal Trancl.thy + "[| : r^*; \ +\ !!x. x: field(r) ==> P(); \ +\ !!x y z.[| P(); : r^*; : r |] ==> P() |] \ +\ ==> P()"; +by (rtac ([rtrancl_def, rtrancl_bnd_mono, major] MRS def_induct) 1); +by (fast_tac (ZF_cs addIs prems addSEs [idE,compE]) 1); +val rtrancl_full_induct = result(); + +(*nice induction rule. + Tried adding the typing hypotheses y,z:field(r), but these + caused expensive case splits!*) +val major::prems = goal Trancl.thy + "[| : r^*; \ +\ P(a); \ +\ !!y z.[| : r^*; : r; P(y) |] ==> P(z) \ +\ |] ==> P(b)"; +(*by induction on this formula*) +by (subgoal_tac "ALL y. = --> P(y)" 1); +(*now solve first subgoal: this formula is sufficient*) +by (EVERY1 [etac (spec RS mp), rtac refl]); +(*now do the induction*) +by (resolve_tac [major RS rtrancl_full_induct] 1); +by (ALLGOALS (fast_tac (ZF_cs addIs prems))); +val rtrancl_induct = result(); + +(*transitivity of transitive closure!! -- by induction.*) +goalw Trancl.thy [trans_def] "trans(r^*)"; +by (REPEAT (resolve_tac [allI,impI] 1)); +by (eres_inst_tac [("b","z")] rtrancl_induct 1); +by (DEPTH_SOLVE (eresolve_tac [asm_rl, rtrancl_into_rtrancl] 1)); +val trans_rtrancl = result(); + +(*elimination of rtrancl -- by induction on a special formula*) +val major::prems = goal Trancl.thy + "[| : r^*; (a=b) ==> P; \ +\ !!y.[| : r^*; : r |] ==> P |] \ +\ ==> P"; +by (subgoal_tac "a = b | (EX y. : r^* & : r)" 1); +(*see HOL/trancl*) +by (rtac (major RS rtrancl_induct) 2); +by (ALLGOALS (fast_tac (ZF_cs addSEs prems))); +val rtranclE = result(); + + +(**** The relation trancl ****) + +(*Transitivity of r^+ is proved by transitivity of r^* *) +goalw Trancl.thy [trans_def,trancl_def] "trans(r^+)"; +by (safe_tac comp_cs); +by (rtac (rtrancl_into_rtrancl RS (trans_rtrancl RS transD RS compI)) 1); +by (REPEAT (assume_tac 1)); +val trans_trancl = result(); + +(** Conversions between trancl and rtrancl **) + +val [major] = goalw Trancl.thy [trancl_def] " : r^+ ==> : r^*"; +by (resolve_tac [major RS compEpair] 1); +by (REPEAT (ares_tac [rtrancl_into_rtrancl] 1)); +val trancl_into_rtrancl = result(); + +(*r^+ contains all pairs in r *) +val [prem] = goalw Trancl.thy [trancl_def] " : r ==> : r^+"; +by (REPEAT (ares_tac [prem,compI,rtrancl_refl,fieldI1] 1)); +val r_into_trancl = result(); + +(*The premise ensures that r consists entirely of pairs*) +val prems = goal Trancl.thy "r <= Sigma(A,B) ==> r <= r^+"; +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addIs [r_into_trancl]) 1); +val r_subset_trancl = result(); + +(*intro rule by definition: from r^* and r *) +val prems = goalw Trancl.thy [trancl_def] + "[| : r^*; : r |] ==> : r^+"; +by (REPEAT (resolve_tac ([compI]@prems) 1)); +val rtrancl_into_trancl1 = result(); + +(*intro rule from r and r^* *) +val prems = goal Trancl.thy + "[| : r; : r^* |] ==> : r^+"; +by (resolve_tac (prems RL [rtrancl_induct]) 1); +by (resolve_tac (prems RL [r_into_trancl]) 1); +by (etac (trans_trancl RS transD) 1); +by (etac r_into_trancl 1); +val rtrancl_into_trancl2 = result(); + +(*Nice induction rule for trancl*) +val major::prems = goal Trancl.thy + "[| : r^+; \ +\ !!y. [| : r |] ==> P(y); \ +\ !!y z.[| : r^+; : r; P(y) |] ==> P(z) \ +\ |] ==> P(b)"; +by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); +(*by induction on this formula*) +by (subgoal_tac "ALL z. : r --> P(z)" 1); +(*now solve first subgoal: this formula is sufficient*) +by (fast_tac ZF_cs 1); +by (etac rtrancl_induct 1); +by (ALLGOALS (fast_tac (ZF_cs addIs (rtrancl_into_trancl1::prems)))); +val trancl_induct = result(); + +(*elimination of r^+ -- NOT an induction rule*) +val major::prems = goal Trancl.thy + "[| : r^+; \ +\ : r ==> P; \ +\ !!y.[| : r^+; : r |] ==> P \ +\ |] ==> P"; +by (subgoal_tac " : r | (EX y. : r^+ & : r)" 1); +by (fast_tac (ZF_cs addIs prems) 1); +by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); +by (etac rtranclE 1); +by (ALLGOALS (fast_tac (ZF_cs addIs [rtrancl_into_trancl1]))); +val tranclE = result(); + +goalw Trancl.thy [trancl_def] "r^+ <= field(r)*field(r)"; +by (fast_tac (ZF_cs addEs [compE, rtrancl_type RS subsetD RS SigmaE2]) 1); +val trancl_type = result(); + +val [prem] = goalw Trancl.thy [trancl_def] "r<=s ==> r^+ <= s^+"; +by (REPEAT (resolve_tac [prem, comp_mono, rtrancl_mono] 1)); +val trancl_mono = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/Trancl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Trancl.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,21 @@ +(* Title: ZF/trancl.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Transitive closure of a relation +*) + +Trancl = Fixedpt + Perm + +consts + "rtrancl" :: "i=>i" ("(_^*)" [100] 100) (*refl/transitive closure*) + "trancl" :: "i=>i" ("(_^+)" [100] 100) (*transitive closure*) + "trans" :: "i=>o" (*transitivity predicate*) + +rules + trans_def "trans(r) == ALL x y z. : r --> : r --> : r" + + rtrancl_def "r^* == lfp(field(r)*field(r), %s. id(field(r)) Un (r O s))" + + trancl_def "r^+ == r O r^*" +end diff -r 000000000000 -r a5a9c433f639 src/ZF/Univ.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Univ.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,642 @@ +(* Title: ZF/univ + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +The cumulative hierarchy and a small universe for recursive types +*) + +open Univ; + +(*NOT SUITABLE FOR REWRITING -- RECURSIVE!*) +goal Univ.thy "Vfrom(A,i) = A Un (UN j:i. Pow(Vfrom(A,j)))"; +by (rtac (Vfrom_def RS def_transrec RS ssubst) 1); +by (SIMP_TAC ZF_ss 1); +val Vfrom = result(); + +(** Monotonicity **) + +goal Univ.thy "!!A B. A<=B ==> ALL j. i<=j --> Vfrom(A,i) <= Vfrom(B,j)"; +by (eps_ind_tac "i" 1); +by (rtac (impI RS allI) 1); +by (rtac (Vfrom RS ssubst) 1); +by (rtac (Vfrom RS ssubst) 1); +by (etac Un_mono 1); +by (rtac UN_mono 1); +by (assume_tac 1); +by (rtac Pow_mono 1); +by (etac (bspec RS spec RS mp) 1); +by (assume_tac 1); +by (rtac subset_refl 1); +val Vfrom_mono_lemma = result(); + +(* [| A<=B; i<=x |] ==> Vfrom(A,i) <= Vfrom(B,x) *) +val Vfrom_mono = standard (Vfrom_mono_lemma RS spec RS mp); + + +(** A fundamental equality: Vfrom does not require ordinals! **) + +goal Univ.thy "Vfrom(A,x) <= Vfrom(A,rank(x))"; +by (eps_ind_tac "x" 1); +by (rtac (Vfrom RS ssubst) 1); +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac (ZF_cs addSIs [rank_lt]) 1); +val Vfrom_rank_subset1 = result(); + +goal Univ.thy "Vfrom(A,rank(x)) <= Vfrom(A,x)"; +by (eps_ind_tac "x" 1); +by (rtac (Vfrom RS ssubst) 1); +by (rtac (Vfrom RS ssubst) 1); +br (subset_refl RS Un_mono) 1; +br UN_least 1; +by (etac (rank_implies_mem RS bexE) 1); +br subset_trans 1; +be UN_upper 2; +by (etac (subset_refl RS Vfrom_mono RS subset_trans RS Pow_mono) 1); +by (etac bspec 1); +by (assume_tac 1); +val Vfrom_rank_subset2 = result(); + +goal Univ.thy "Vfrom(A,rank(x)) = Vfrom(A,x)"; +by (rtac equalityI 1); +by (rtac Vfrom_rank_subset2 1); +by (rtac Vfrom_rank_subset1 1); +val Vfrom_rank_eq = result(); + + +(*** Basic closure properties ***) + +goal Univ.thy "!!x y. y:x ==> 0 : Vfrom(A,x)"; +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac ZF_cs 1); +val zero_in_Vfrom = result(); + +goal Univ.thy "i <= Vfrom(A,i)"; +by (eps_ind_tac "i" 1); +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac ZF_cs 1); +val i_subset_Vfrom = result(); + +goal Univ.thy "A <= Vfrom(A,i)"; +by (rtac (Vfrom RS ssubst) 1); +by (rtac Un_upper1 1); +val A_subset_Vfrom = result(); + +goal Univ.thy "!!A a i. a <= Vfrom(A,i) ==> a: Vfrom(A,succ(i))"; +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac ZF_cs 1); +val subset_mem_Vfrom = result(); + +(** Finite sets and ordered pairs **) + +goal Univ.thy "!!a. a: Vfrom(A,i) ==> {a} : Vfrom(A,succ(i))"; +by (rtac subset_mem_Vfrom 1); +by (safe_tac ZF_cs); +val singleton_in_Vfrom = result(); + +goal Univ.thy + "!!A. [| a: Vfrom(A,i); b: Vfrom(A,i) |] ==> {a,b} : Vfrom(A,succ(i))"; +by (rtac subset_mem_Vfrom 1); +by (safe_tac ZF_cs); +val doubleton_in_Vfrom = result(); + +goalw Univ.thy [Pair_def] + "!!A. [| a: Vfrom(A,i); b: Vfrom(A,i) |] ==> \ +\ : Vfrom(A,succ(succ(i)))"; +by (REPEAT (ares_tac [doubleton_in_Vfrom] 1)); +val Pair_in_Vfrom = result(); + +val [prem] = goal Univ.thy + "a<=Vfrom(A,i) ==> succ(a) : Vfrom(A,succ(succ(i)))"; +by (REPEAT (resolve_tac [subset_mem_Vfrom, succ_subsetI] 1)); +by (rtac (Vfrom_mono RSN (2,subset_trans)) 2); +by (REPEAT (resolve_tac [prem, subset_refl, subset_succI] 1)); +val succ_in_Vfrom = result(); + +(*** 0, successor and limit equations fof Vfrom ***) + +goal Univ.thy "Vfrom(A,0) = A"; +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac eq_cs 1); +val Vfrom_0 = result(); + +goal Univ.thy "!!i. Ord(i) ==> Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"; +by (rtac (Vfrom RS trans) 1); +brs ([refl] RL ZF_congs) 1; +by (rtac equalityI 1); +by (rtac (succI1 RS RepFunI RS Union_upper) 2); +by (rtac UN_least 1); +by (rtac (subset_refl RS Vfrom_mono RS Pow_mono) 1); +by (etac member_succD 1); +by (assume_tac 1); +val Vfrom_succ_lemma = result(); + +goal Univ.thy "Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"; +by (res_inst_tac [("x1", "succ(i)")] (Vfrom_rank_eq RS subst) 1); +by (res_inst_tac [("x1", "i")] (Vfrom_rank_eq RS subst) 1); +by (rtac (rank_succ RS ssubst) 1); +by (rtac (Ord_rank RS Vfrom_succ_lemma) 1); +val Vfrom_succ = result(); + +(*The premise distinguishes this from Vfrom(A,0); allowing X=0 forces + the conclusion to be Vfrom(A,Union(X)) = A Un (UN y:X. Vfrom(A,y)) *) +val [prem] = goal Univ.thy "y:X ==> Vfrom(A,Union(X)) = (UN y:X. Vfrom(A,y))"; +by (rtac (Vfrom RS ssubst) 1); +by (rtac equalityI 1); +(*first inclusion*) +by (rtac Un_least 1); +br (A_subset_Vfrom RS subset_trans) 1; +br (prem RS UN_upper) 1; +br UN_least 1; +be UnionE 1; +br subset_trans 1; +be UN_upper 2; +by (rtac (Vfrom RS ssubst) 1); +be ([UN_upper, Un_upper2] MRS subset_trans) 1; +(*opposite inclusion*) +br UN_least 1; +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac ZF_cs 1); +val Vfrom_Union = result(); + +(*** Limit ordinals -- general properties ***) + +goalw Univ.thy [Limit_def] "!!i. Limit(i) ==> Union(i) = i"; +by (fast_tac (eq_cs addEs [Ord_trans]) 1); +val Limit_Union_eq = result(); + +goalw Univ.thy [Limit_def] "!!i. Limit(i) ==> Ord(i)"; +by (etac conjunct1 1); +val Limit_is_Ord = result(); + +goalw Univ.thy [Limit_def] "!!i. Limit(i) ==> 0 : i"; +by (fast_tac ZF_cs 1); +val Limit_has_0 = result(); + +goalw Univ.thy [Limit_def] "!!i. [| Limit(i); j:i |] ==> succ(j) : i"; +by (fast_tac ZF_cs 1); +val Limit_has_succ = result(); + +goalw Univ.thy [Limit_def] "Limit(nat)"; +by (REPEAT (ares_tac [conjI, ballI, nat_0I, nat_succI, Ord_nat] 1)); +val Limit_nat = result(); + +goalw Univ.thy [Limit_def] + "!!i. [| Ord(i); 0:i; ALL y. ~ succ(y)=i |] ==> Limit(i)"; +by (safe_tac subset_cs); +br Ord_member 1; +by (REPEAT_FIRST (eresolve_tac [asm_rl, Ord_in_Ord RS Ord_succ] + ORELSE' dresolve_tac [Ord_succ_subsetI])); +by (fast_tac (subset_cs addSIs [equalityI]) 1); +val non_succ_LimitI = result(); + +goal Univ.thy "!!i. Ord(i) ==> i=0 | (EX j. i=succ(j)) | Limit(i)"; +by (fast_tac (ZF_cs addSIs [non_succ_LimitI, Ord_0_member_iff RS iffD2]) 1); +val Ord_cases_lemma = result(); + +val major::prems = goal Univ.thy + "[| Ord(i); \ +\ i=0 ==> P; \ +\ !!j. i=succ(j) ==> P; \ +\ Limit(i) ==> P \ +\ |] ==> P"; +by (cut_facts_tac [major RS Ord_cases_lemma] 1); +by (REPEAT (eresolve_tac (prems@[disjE, exE]) 1)); +val Ord_cases = result(); + + +(*** Vfrom applied to Limit ordinals ***) + +(*NB. limit ordinals are non-empty; + Vfrom(A,0) = A = A Un (UN y:0. Vfrom(A,y)) *) +val [limiti] = goal Univ.thy + "Limit(i) ==> Vfrom(A,i) = (UN y:i. Vfrom(A,y))"; +by (rtac (limiti RS Limit_has_0 RS Vfrom_Union RS subst) 1); +by (rtac (limiti RS Limit_Union_eq RS ssubst) 1); +by (rtac refl 1); +val Limit_Vfrom_eq = result(); + +val Limit_VfromE = standard (Limit_Vfrom_eq RS equalityD1 RS subsetD RS UN_E); + +val [major,limiti] = goal Univ.thy + "[| a: Vfrom(A,i); Limit(i) |] ==> {a} : Vfrom(A,i)"; +by (rtac (limiti RS Limit_VfromE) 1); +by (rtac major 1); +by (rtac (limiti RS Limit_Vfrom_eq RS ssubst) 1); +by (rtac UN_I 1); +by (etac singleton_in_Vfrom 2); +by (etac (limiti RS Limit_has_succ) 1); +val singleton_in_Vfrom_limit = result(); + +val Vfrom_UnI1 = Un_upper1 RS (subset_refl RS Vfrom_mono RS subsetD) +and Vfrom_UnI2 = Un_upper2 RS (subset_refl RS Vfrom_mono RS subsetD); + +(*Hard work is finding a single j:i such that {a,b}<=Vfrom(A,j)*) +val [aprem,bprem,limiti] = goal Univ.thy + "[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i) |] ==> \ +\ {a,b} : Vfrom(A,i)"; +by (rtac (aprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (bprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (limiti RS Limit_Vfrom_eq RS ssubst) 1); +by (rtac UN_I 1); +by (rtac doubleton_in_Vfrom 2); +by (etac Vfrom_UnI1 2); +by (etac Vfrom_UnI2 2); +by (REPEAT (ares_tac[limiti, Limit_has_succ, Ord_member_UnI, Limit_is_Ord] 1)); +val doubleton_in_Vfrom_limit = result(); + +val [aprem,bprem,limiti] = goal Univ.thy + "[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i) |] ==> \ +\ : Vfrom(A,i)"; +(*Infer that a, b occur at ordinals x,xa < i.*) +by (rtac (aprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (bprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (limiti RS Limit_Vfrom_eq RS ssubst) 1); +by (rtac UN_I 1); +by (rtac Pair_in_Vfrom 2); +(*Infer that succ(succ(x Un xa)) < i *) +by (etac Vfrom_UnI1 2); +by (etac Vfrom_UnI2 2); +by (REPEAT (ares_tac[limiti, Limit_has_succ, Ord_member_UnI, Limit_is_Ord] 1)); +val Pair_in_Vfrom_limit = result(); + + +(*** Properties assuming Transset(A) ***) + +goal Univ.thy "!!i A. Transset(A) ==> Transset(Vfrom(A,i))"; +by (eps_ind_tac "i" 1); +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac (ZF_cs addSIs [Transset_Union_family, Transset_Un, + Transset_Pow]) 1); +val Transset_Vfrom = result(); + +goal Univ.thy "!!A i. Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))"; +by (rtac (Vfrom_succ RS trans) 1); +br (Un_upper2 RSN (2,equalityI)) 1;; +br (subset_refl RSN (2,Un_least)) 1;; +br (A_subset_Vfrom RS subset_trans) 1; +be (Transset_Vfrom RS (Transset_iff_Pow RS iffD1)) 1; +val Transset_Vfrom_succ = result(); + +goalw Ord.thy [Pair_def,Transset_def] + "!!C. [| <= C; Transset(C) |] ==> a: C & b: C"; +by (fast_tac ZF_cs 1); +val Transset_Pair_subset = result(); + +goal Univ.thy + "!!a b.[| <= Vfrom(A,i); Transset(A); Limit(i) |] ==> \ +\ : Vfrom(A,i)"; +be (Transset_Pair_subset RS conjE) 1; +be Transset_Vfrom 1; +by (REPEAT (ares_tac [Pair_in_Vfrom_limit] 1)); +val Transset_Pair_subset_Vfrom_limit = result(); + + +(*** Closure under product/sum applied to elements -- thus Vfrom(A,i) + is a model of simple type theory provided A is a transitive set + and i is a limit ordinal +***) + +(*There are three nearly identical proofs below -- needs a general theorem + for proving ...a...b : Vfrom(A,i) where i is a limit ordinal*) + +(** products **) + +goal Univ.thy + "!!A. [| a: Vfrom(A,i); b: Vfrom(A,i); Transset(A) |] ==> \ +\ a*b : Vfrom(A, succ(succ(succ(i))))"; +by (dtac Transset_Vfrom 1); +by (rtac subset_mem_Vfrom 1); +by (rewtac Transset_def); +by (fast_tac (ZF_cs addIs [Pair_in_Vfrom]) 1); +val prod_in_Vfrom = result(); + +val [aprem,bprem,limiti,transset] = goal Univ.thy + "[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); Transset(A) |] ==> \ +\ a*b : Vfrom(A,i)"; +(*Infer that a, b occur at ordinals x,xa < i.*) +by (rtac (aprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (bprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (limiti RS Limit_Vfrom_eq RS ssubst) 1); +by (rtac UN_I 1); +by (rtac prod_in_Vfrom 2); +(*Infer that succ(succ(succ(x Un xa))) < i *) +by (etac (Vfrom_UnI1 RS Vfrom_UnI2) 2); +by (etac (Vfrom_UnI2 RS Vfrom_UnI2) 2); +by (REPEAT (ares_tac [limiti RS Limit_has_succ, + Ord_member_UnI, limiti RS Limit_is_Ord, transset] 1)); +val prod_in_Vfrom_limit = result(); + +(** Disjoint sums, aka Quine ordered pairs **) + +goalw Univ.thy [sum_def] + "!!A. [| a: Vfrom(A,i); b: Vfrom(A,i); Transset(A); 1:i |] ==> \ +\ a+b : Vfrom(A, succ(succ(succ(i))))"; +by (dtac Transset_Vfrom 1); +by (rtac subset_mem_Vfrom 1); +by (rewtac Transset_def); +by (fast_tac (ZF_cs addIs [zero_in_Vfrom, Pair_in_Vfrom, + i_subset_Vfrom RS subsetD]) 1); +val sum_in_Vfrom = result(); + +val [aprem,bprem,limiti,transset] = goal Univ.thy + "[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); Transset(A) |] ==> \ +\ a+b : Vfrom(A,i)"; +(*Infer that a, b occur at ordinals x,xa < i.*) +by (rtac (aprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (bprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (limiti RS Limit_Vfrom_eq RS ssubst) 1); +by (rtac UN_I 1); +by (rtac (rewrite_rule [one_def] sum_in_Vfrom) 2); +by (rtac (succI1 RS UnI1) 5); +(*Infer that succ(succ(succ(x Un xa))) < i *) +by (etac (Vfrom_UnI1 RS Vfrom_UnI2) 2); +by (etac (Vfrom_UnI2 RS Vfrom_UnI2) 2); +by (REPEAT (ares_tac [limiti RS Limit_has_0, + limiti RS Limit_has_succ, + Ord_member_UnI, limiti RS Limit_is_Ord, transset] 1)); +val sum_in_Vfrom_limit = result(); + +(** function space! **) + +goalw Univ.thy [Pi_def] + "!!A. [| a: Vfrom(A,i); b: Vfrom(A,i); Transset(A) |] ==> \ +\ a->b : Vfrom(A, succ(succ(succ(succ(i)))))"; +by (dtac Transset_Vfrom 1); +by (rtac subset_mem_Vfrom 1); +by (rtac (Collect_subset RS subset_trans) 1); +by (rtac (Vfrom RS ssubst) 1); +by (rtac (subset_trans RS subset_trans) 1); +by (rtac Un_upper2 3); +by (rtac (succI1 RS UN_upper) 2); +by (rtac Pow_mono 1); +by (rewtac Transset_def); +by (fast_tac (ZF_cs addIs [Pair_in_Vfrom]) 1); +val fun_in_Vfrom = result(); + +val [aprem,bprem,limiti,transset] = goal Univ.thy + "[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); Transset(A) |] ==> \ +\ a->b : Vfrom(A,i)"; +(*Infer that a, b occur at ordinals x,xa < i.*) +by (rtac (aprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (bprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (limiti RS Limit_Vfrom_eq RS ssubst) 1); +by (rtac UN_I 1); +by (rtac fun_in_Vfrom 2); +(*Infer that succ(succ(succ(x Un xa))) < i *) +by (etac (Vfrom_UnI1 RS Vfrom_UnI2) 2); +by (etac (Vfrom_UnI2 RS Vfrom_UnI2) 2); +by (REPEAT (ares_tac [limiti RS Limit_has_succ, + Ord_member_UnI, limiti RS Limit_is_Ord, transset] 1)); +val fun_in_Vfrom_limit = result(); + + +(*** The set Vset(i) ***) + +goal Univ.thy "Vset(i) = (UN j:i. Pow(Vset(j)))"; +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac eq_cs 1); +val Vset = result(); + +val Vset_succ = Transset_0 RS Transset_Vfrom_succ; + +val Transset_Vset = Transset_0 RS Transset_Vfrom; + +(** Characterisation of the elements of Vset(i) **) + +val [ordi] = goal Univ.thy "Ord(i) ==> ALL b. b : Vset(i) --> rank(b) : i"; +by (rtac (ordi RS trans_induct) 1); +by (rtac (Vset RS ssubst) 1); +by (safe_tac ZF_cs); +by (rtac (rank RS ssubst) 1); +by (rtac sup_least2 1); +by (assume_tac 1); +by (assume_tac 1); +by (fast_tac ZF_cs 1); +val Vset_rank_imp1 = result(); + +(* [| Ord(i); x : Vset(i) |] ==> rank(x) : i *) +val Vset_D = standard (Vset_rank_imp1 RS spec RS mp); + +val [ordi] = goal Univ.thy "Ord(i) ==> ALL b. rank(b) : i --> b : Vset(i)"; +by (rtac (ordi RS trans_induct) 1); +by (rtac allI 1); +by (rtac (Vset RS ssubst) 1); +by (fast_tac (ZF_cs addSIs [rank_lt]) 1); +val Vset_rank_imp2 = result(); + +(* [| Ord(i); rank(x) : i |] ==> x : Vset(i) *) +val VsetI = standard (Vset_rank_imp2 RS spec RS mp); + +val [ordi] = goal Univ.thy "Ord(i) ==> b : Vset(i) <-> rank(b) : i"; +by (rtac iffI 1); +by (etac (ordi RS Vset_D) 1); +by (etac (ordi RS VsetI) 1); +val Vset_Ord_rank_iff = result(); + +goal Univ.thy "b : Vset(a) <-> rank(b) : rank(a)"; +by (rtac (Vfrom_rank_eq RS subst) 1); +by (rtac (Ord_rank RS Vset_Ord_rank_iff) 1); +val Vset_rank_iff = result(); + +goal Univ.thy "!!i. Ord(i) ==> rank(Vset(i)) = i"; +by (rtac (rank RS ssubst) 1); +by (rtac equalityI 1); +by (safe_tac ZF_cs); +by (EVERY' [wtac UN_I, + etac (i_subset_Vfrom RS subsetD), + etac (Ord_in_Ord RS rank_of_Ord RS ssubst), + assume_tac, + rtac succI1] 3); +by (REPEAT (eresolve_tac [asm_rl,Vset_D,Ord_trans] 1)); +val rank_Vset = result(); + +(** Lemmas for reasoning about sets in terms of their elements' ranks **) + +(* rank(x) : rank(a) ==> x : Vset(rank(a)) *) +val Vset_rankI = Ord_rank RS VsetI; + +goal Univ.thy "a <= Vset(rank(a))"; +br subsetI 1; +be (rank_lt RS Vset_rankI) 1; +val arg_subset_Vset_rank = result(); + +val [iprem] = goal Univ.thy + "[| !!i. Ord(i) ==> a Int Vset(i) <= b |] ==> a <= b"; +br ([subset_refl, arg_subset_Vset_rank] MRS Int_greatest RS subset_trans) 1; +br (Ord_rank RS iprem) 1; +val Int_Vset_subset = result(); + +(** Set up an environment for simplification **) + +val rank_rls = [rank_Inl, rank_Inr, rank_pair1, rank_pair2]; +val rank_trans_rls = rank_rls @ (rank_rls RLN (2, [rank_trans])); + +val rank_ss = ZF_ss + addrews [split, case_Inl, case_Inr, Vset_rankI] + addrews rank_trans_rls; + +(** Recursion over Vset levels! **) + +(*NOT SUITABLE FOR REWRITING: recursive!*) +goalw Univ.thy [Vrec_def] "Vrec(a,H) = H(a, lam x:Vset(rank(a)). Vrec(x,H))"; +by (rtac (transrec RS ssubst) 1); +by (SIMP_TAC (wf_ss addrews [Ord_rank, Ord_succ, Vset_D RS beta, + VsetI RS beta]) 1); +val Vrec = result(); + +(*This form avoids giant explosions in proofs. NOTE USE OF == *) +val rew::prems = goal Univ.thy + "[| !!x. h(x)==Vrec(x,H) |] ==> \ +\ h(a) = H(a, lam x: Vset(rank(a)). h(x))"; +by (rewtac rew); +by (rtac Vrec 1); +val def_Vrec = result(); + +val prems = goalw Univ.thy [Vrec_def] + "[| a=a'; !!x u. H(x,u)=H'(x,u) |] ==> Vrec(a,H)=Vrec(a',H')"; +val Vrec_ss = ZF_ss addcongs ([transrec_cong] @ mk_congs Univ.thy ["rank"]) + addrews (prems RL [sym]); +by (SIMP_TAC Vrec_ss 1); +val Vrec_cong = result(); + + +(*** univ(A) ***) + +goalw Univ.thy [univ_def] "!!A B. A<=B ==> univ(A) <= univ(B)"; +by (etac Vfrom_mono 1); +by (rtac subset_refl 1); +val univ_mono = result(); + +goalw Univ.thy [univ_def] "!!A. Transset(A) ==> Transset(univ(A))"; +by (etac Transset_Vfrom 1); +val Transset_univ = result(); + +(** univ(A) as a limit **) + +goalw Univ.thy [univ_def] "univ(A) = (UN i:nat. Vfrom(A,i))"; +br (Limit_nat RS Limit_Vfrom_eq) 1; +val univ_eq_UN = result(); + +goal Univ.thy "!!c. c <= univ(A) ==> c = (UN i:nat. c Int Vfrom(A,i))"; +br (subset_UN_iff_eq RS iffD1) 1; +be (univ_eq_UN RS subst) 1; +val subset_univ_eq_Int = result(); + +val [aprem, iprem] = goal Univ.thy + "[| a <= univ(X); \ +\ !!i. i:nat ==> a Int Vfrom(X,i) <= b \ +\ |] ==> a <= b"; +br (aprem RS subset_univ_eq_Int RS ssubst) 1; +br UN_least 1; +be iprem 1; +val univ_Int_Vfrom_subset = result(); + +val prems = goal Univ.thy + "[| a <= univ(X); b <= univ(X); \ +\ !!i. i:nat ==> a Int Vfrom(X,i) = b Int Vfrom(X,i) \ +\ |] ==> a = b"; +br equalityI 1; +by (ALLGOALS + (resolve_tac (prems RL [univ_Int_Vfrom_subset]) THEN' + eresolve_tac (prems RL [equalityD1,equalityD2] RL [subset_trans]) THEN' + rtac Int_lower1)); +val univ_Int_Vfrom_eq = result(); + +(** Closure properties **) + +goalw Univ.thy [univ_def] "0 : univ(A)"; +by (rtac (nat_0I RS zero_in_Vfrom) 1); +val zero_in_univ = result(); + +goalw Univ.thy [univ_def] "A <= univ(A)"; +by (rtac A_subset_Vfrom 1); +val A_subset_univ = result(); + +val A_into_univ = A_subset_univ RS subsetD; + +(** Closure under unordered and ordered pairs **) + +goalw Univ.thy [univ_def] "!!A a. a: univ(A) ==> {a} : univ(A)"; +by (rtac singleton_in_Vfrom_limit 1); +by (REPEAT (ares_tac [Ord_nat,Limit_nat] 1)); +val singleton_in_univ = result(); + +goalw Univ.thy [univ_def] + "!!A a. [| a: univ(A); b: univ(A) |] ==> {a,b} : univ(A)"; +by (rtac doubleton_in_Vfrom_limit 1); +by (REPEAT (ares_tac [Ord_nat,Limit_nat] 1)); +val doubleton_in_univ = result(); + +goalw Univ.thy [univ_def] + "!!A a. [| a: univ(A); b: univ(A) |] ==> : univ(A)"; +by (rtac Pair_in_Vfrom_limit 1); +by (REPEAT (ares_tac [Ord_nat,Limit_nat] 1)); +val Pair_in_univ = result(); + +goal Univ.thy "univ(A)*univ(A) <= univ(A)"; +by (REPEAT (ares_tac [subsetI,Pair_in_univ] 1 + ORELSE eresolve_tac [SigmaE, ssubst] 1)); +val product_univ = result(); + +val Sigma_subset_univ = standard + (Sigma_mono RS (product_univ RSN (2,subset_trans))); + +goalw Univ.thy [univ_def] + "!!a b.[| <= univ(A); Transset(A) |] ==> : univ(A)"; +be Transset_Pair_subset_Vfrom_limit 1; +by (REPEAT (ares_tac [Ord_nat,Limit_nat] 1)); +val Transset_Pair_subset_univ = result(); + + +(** The natural numbers **) + +goalw Univ.thy [univ_def] "nat <= univ(A)"; +by (rtac i_subset_Vfrom 1); +val nat_subset_univ = result(); + +(* n:nat ==> n:univ(A) *) +val nat_into_univ = standard (nat_subset_univ RS subsetD); + +(** instances for 1 and 2 **) + +goalw Univ.thy [one_def] "1 : univ(A)"; +by (REPEAT (ares_tac [nat_into_univ, nat_0I, nat_succI] 1)); +val one_in_univ = result(); + +(*unused!*) +goal Univ.thy "succ(succ(0)) : univ(A)"; +by (REPEAT (ares_tac [nat_into_univ, nat_0I, nat_succI] 1)); +val two_in_univ = result(); + +goalw Univ.thy [bool_def] "bool <= univ(A)"; +by (fast_tac (ZF_cs addSIs [zero_in_univ,one_in_univ]) 1); +val bool_subset_univ = result(); + +val bool_into_univ = standard (bool_subset_univ RS subsetD); + + +(** Closure under disjoint union **) + +goalw Univ.thy [Inl_def] "!!A a. a: univ(A) ==> Inl(a) : univ(A)"; +by (REPEAT (ares_tac [zero_in_univ,Pair_in_univ] 1)); +val Inl_in_univ = result(); + +goalw Univ.thy [Inr_def] "!!A b. b: univ(A) ==> Inr(b) : univ(A)"; +by (REPEAT (ares_tac [one_in_univ, Pair_in_univ] 1)); +val Inr_in_univ = result(); + +goal Univ.thy "univ(C)+univ(C) <= univ(C)"; +by (REPEAT (ares_tac [subsetI,Inl_in_univ,Inr_in_univ] 1 + ORELSE eresolve_tac [sumE, ssubst] 1)); +val sum_univ = result(); + +val sum_subset_univ = standard + (sum_mono RS (sum_univ RSN (2,subset_trans))); + + +(** Closure under binary union -- use Un_least **) +(** Closure under Collect -- use (Collect_subset RS subset_trans) **) +(** Closure under RepFun -- use RepFun_subset **) + + diff -r 000000000000 -r a5a9c433f639 src/ZF/Univ.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Univ.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,34 @@ +(* Title: ZF/univ.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +The cumulative hierarchy and a small universe for recursive types + +Standard notation for Vset(i) is V(i), but users might want V for a variable +*) + +Univ = Arith + Sum + +consts + Limit :: "i=>o" + Vfrom :: "[i,i]=>i" + Vset :: "i=>i" + Vrec :: "[i, [i,i]=>i] =>i" + univ :: "i=>i" + +translations + (*Apparently a bug prevents using "Vset" == "Vfrom(0)" *) + "Vset(x)" == "Vfrom(0,x)" + +rules + Limit_def "Limit(i) == Ord(i) & 0:i & (ALL y:i. succ(y): i)" + + Vfrom_def "Vfrom(A,i) == transrec(i, %x f. A Un (UN y:x. Pow(f`y)))" + + Vrec_def + "Vrec(a,H) == transrec(rank(a), %x g. lam z: Vset(succ(x)). \ +\ H(z, lam w:Vset(x). g`rank(w)`w)) ` a" + + univ_def "univ(A) == Vfrom(A,nat)" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/WF.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/WF.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,262 @@ +(* Title: ZF/wf.ML + ID: $Id$ + Author: Tobias Nipkow and Lawrence C Paulson + Copyright 1992 University of Cambridge + +For wf.thy. Well-founded Recursion + +Derived first for transitive relations, and finally for arbitrary WF relations +via wf_trancl and trans_trancl. + +It is difficult to derive this general case directly, using r^+ instead of +r. In is_recfun, the two occurrences of the relation must have the same +form. Inserting r^+ in the_recfun or wftrec yields a recursion rule with +r^+ -`` {a} instead of r-``{a}. This recursion rule is stronger in +principle, but harder to use, especially to prove wfrec_eclose_eq in +epsilon.ML. Expanding out the definition of wftrec in wfrec would yield +a mess. +*) + +open WF; + +val [H_cong] = mk_typed_congs WF.thy[("H","[i,i]=>i")]; + +val wf_ss = ZF_ss addcongs [H_cong]; + + +(*** Well-founded relations ***) + +(*Are these two theorems at all useful??*) + +(*If every subset of field(r) possesses an r-minimal element then wf(r). + Seems impossible to prove this for domain(r) or range(r) instead... + Consider in particular finite wf relations!*) +val [prem1,prem2] = goalw WF.thy [wf_def] + "[| field(r)<=A; \ +\ !!Z u. [| Z<=A; u:Z; ALL x:Z. EX y:Z. :r |] ==> False |] \ +\ ==> wf(r)"; +by (rtac (equals0I RS disjCI RS allI) 1); +by (rtac prem2 1); +by (res_inst_tac [ ("B1", "Z") ] (prem1 RS (Int_lower1 RS subset_trans)) 1); +by (fast_tac ZF_cs 1); +by (fast_tac ZF_cs 1); +val wfI = result(); + +(*If r allows well-founded induction then wf(r)*) +val [prem1,prem2] = goal WF.thy + "[| field(r)<=A; \ +\ !!B. ALL x:A. (ALL y. : r --> y:B) --> x:B ==> A<=B |] \ +\ ==> wf(r)"; +by (rtac (prem1 RS wfI) 1); +by (res_inst_tac [ ("B", "A-Z") ] (prem2 RS subsetCE) 1); +by (fast_tac ZF_cs 3); +by (fast_tac ZF_cs 2); +by (fast_tac ZF_cs 1); +val wfI2 = result(); + + +(** Well-founded Induction **) + +(*Consider the least z in domain(r) Un {a} such that P(z) does not hold...*) +val major::prems = goalw WF.thy [wf_def] + "[| wf(r); \ +\ !!x.[| ALL y. : r --> P(y) |] ==> P(x) \ +\ |] ==> P(a)"; +by (res_inst_tac [ ("x", "{z:domain(r) Un {a}. ~P(z)}") ] (major RS allE) 1); +by (etac disjE 1); +by (rtac classical 1); +by (etac equals0D 1); +by (etac (singletonI RS UnI2 RS CollectI) 1); +by (etac bexE 1); +by (etac CollectE 1); +by (etac swap 1); +by (resolve_tac prems 1); +by (fast_tac ZF_cs 1); +val wf_induct = result(); + +(*Perform induction on i, then prove the wf(r) subgoal using prems. *) +fun wf_ind_tac a prems i = + EVERY [res_inst_tac [("a",a)] wf_induct i, + rename_last_tac a ["1"] (i+1), + ares_tac prems i]; + +(*The form of this rule is designed to match wfI2*) +val wfr::amem::prems = goal WF.thy + "[| wf(r); a:A; field(r)<=A; \ +\ !!x.[| x: A; ALL y. : r --> P(y) |] ==> P(x) \ +\ |] ==> P(a)"; +by (rtac (amem RS rev_mp) 1); +by (wf_ind_tac "a" [wfr] 1); +by (rtac impI 1); +by (eresolve_tac prems 1); +by (fast_tac (ZF_cs addIs (prems RL [subsetD])) 1); +val wf_induct2 = result(); + +val prems = goal WF.thy "[| wf(r); :r; :r |] ==> False"; +by (subgoal_tac "ALL x. :r --> :r --> False" 1); +by (wf_ind_tac "a" prems 2); +by (fast_tac ZF_cs 2); +by (fast_tac (FOL_cs addIs prems) 1); +val wf_anti_sym = result(); + +(*transitive closure of a WF relation is WF!*) +val [prem] = goal WF.thy "wf(r) ==> wf(r^+)"; +by (rtac (trancl_type RS field_rel_subset RS wfI2) 1); +by (rtac subsetI 1); +(*must retain the universal formula for later use!*) +by (rtac (bspec RS mp) 1 THEN assume_tac 1 THEN assume_tac 1); +by (eres_inst_tac [("a","x")] (prem RS wf_induct2) 1); +by (rtac subset_refl 1); +by (rtac (impI RS allI) 1); +by (etac tranclE 1); +by (etac (bspec RS mp) 1); +by (etac fieldI1 1); +by (fast_tac ZF_cs 1); +by (fast_tac ZF_cs 1); +val wf_trancl = result(); + +(** r-``{a} is the set of everything under a in r **) + +val underI = standard (vimage_singleton_iff RS iffD2); +val underD = standard (vimage_singleton_iff RS iffD1); + +(** is_recfun **) + +val [major] = goalw WF.thy [is_recfun_def] + "is_recfun(r,a,H,f) ==> f: r-``{a} -> range(f)"; +by (rtac (major RS ssubst) 1); +by (rtac (lamI RS rangeI RS lam_type) 1); +by (assume_tac 1); +val is_recfun_type = result(); + +val [isrec,rel] = goalw WF.thy [is_recfun_def] + "[| is_recfun(r,a,H,f); :r |] ==> f`x = H(x, restrict(f,r-``{x}))"; +by (res_inst_tac [("P", "%x.?t(x) = ?u::i")] (isrec RS ssubst) 1); +by (rtac (rel RS underI RS beta) 1); +val apply_recfun = result(); + +(*eresolve_tac transD solves :r using transitivity AT MOST ONCE + spec RS mp instantiates induction hypotheses*) +fun indhyp_tac hyps = + ares_tac (TrueI::hyps) ORELSE' + (cut_facts_tac hyps THEN' + DEPTH_SOLVE_1 o (ares_tac [TrueI, ballI] ORELSE' + eresolve_tac [underD, transD, spec RS mp])); + +(*** NOTE! some simplifications need a different auto_tac!! ***) +val wf_super_ss = wf_ss setauto indhyp_tac; + +val prems = goalw WF.thy [is_recfun_def] + "[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g) |] ==> \ +\ :r --> :r --> f`x=g`x"; +by (cut_facts_tac prems 1); +by (wf_ind_tac "x" prems 1); +by (REPEAT (rtac impI 1 ORELSE etac ssubst 1)); +by (rewtac restrict_def); +by (ASM_SIMP_TAC (wf_super_ss addrews [vimage_singleton_iff]) 1); +val is_recfun_equal_lemma = result(); +val is_recfun_equal = standard (is_recfun_equal_lemma RS mp RS mp); + +val prems as [wfr,transr,recf,recg,_] = goal WF.thy + "[| wf(r); trans(r); \ +\ is_recfun(r,a,H,f); is_recfun(r,b,H,g); :r |] ==> \ +\ restrict(f, r-``{b}) = g"; +by (cut_facts_tac prems 1); +by (rtac (consI1 RS restrict_type RS fun_extension) 1); +by (etac is_recfun_type 1); +by (ALLGOALS + (ASM_SIMP_TAC (wf_super_ss addrews + [ [wfr,transr,recf,recg] MRS is_recfun_equal ]))); +val is_recfun_cut = result(); + +(*** Main Existence Lemma ***) + +val prems = goal WF.thy + "[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,a,H,g) |] ==> f=g"; +by (cut_facts_tac prems 1); +by (rtac fun_extension 1); +by (REPEAT (ares_tac [is_recfun_equal] 1 + ORELSE eresolve_tac [is_recfun_type,underD] 1)); +val is_recfun_functional = result(); + +(*If some f satisfies is_recfun(r,a,H,-) then so does the_recfun(r,a,H) *) +val prems = goalw WF.thy [the_recfun_def] + "[| is_recfun(r,a,H,f); wf(r); trans(r) |] \ +\ ==> is_recfun(r, a, H, the_recfun(r,a,H))"; +by (rtac (ex1I RS theI) 1); +by (REPEAT (ares_tac (prems@[is_recfun_functional]) 1)); +val is_the_recfun = result(); + +val prems = goal WF.thy + "[| wf(r); trans(r) |] ==> is_recfun(r, a, H, the_recfun(r,a,H))"; +by (cut_facts_tac prems 1); +by (wf_ind_tac "a" prems 1); +by (res_inst_tac [("f", "lam y: r-``{a1}. wftrec(r,y,H)")] is_the_recfun 1); +by (REPEAT (assume_tac 2)); +by (rewrite_goals_tac [is_recfun_def, wftrec_def]); +(*Applying the substitution: must keep the quantified assumption!!*) +by (REPEAT (dtac underD 1 ORELSE resolve_tac [refl, lam_cong, H_cong] 1)); +by (fold_tac [is_recfun_def]); +by (rtac (consI1 RS restrict_type RSN (2,fun_extension)) 1); +by (rtac is_recfun_type 1); +by (ALLGOALS + (ASM_SIMP_TAC + (wf_super_ss addrews [underI RS beta, apply_recfun, is_recfun_cut]))); +val unfold_the_recfun = result(); + + +(*** Unfolding wftrec ***) + +val prems = goal WF.thy + "[| wf(r); trans(r); :r |] ==> \ +\ restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)"; +by (REPEAT (ares_tac (prems @ [is_recfun_cut, unfold_the_recfun]) 1)); +val the_recfun_cut = result(); + +(*NOT SUITABLE FOR REWRITING since it is recursive!*) +val prems = goalw WF.thy [wftrec_def] + "[| wf(r); trans(r) |] ==> \ +\ wftrec(r,a,H) = H(a, lam x: r-``{a}. wftrec(r,x,H))"; +by (rtac (rewrite_rule [is_recfun_def] unfold_the_recfun RS ssubst) 1); +by (ALLGOALS (ASM_SIMP_TAC + (wf_ss addrews (prems@[vimage_singleton_iff RS iff_sym, + the_recfun_cut])))); +val wftrec = result(); + +(** Removal of the premise trans(r) **) + +(*NOT SUITABLE FOR REWRITING since it is recursive!*) +val [wfr] = goalw WF.thy [wfrec_def] + "wf(r) ==> wfrec(r,a,H) = H(a, lam x:r-``{a}. wfrec(r,x,H))"; +by (rtac (wfr RS wf_trancl RS wftrec RS ssubst) 1); +by (rtac trans_trancl 1); +by (rtac (refl RS H_cong) 1); +by (rtac (vimage_pair_mono RS restrict_lam_eq) 1); +by (etac r_into_trancl 1); +by (rtac subset_refl 1); +val wfrec = result(); + +(*This form avoids giant explosions in proofs. NOTE USE OF == *) +val rew::prems = goal WF.thy + "[| !!x. h(x)==wfrec(r,x,H); wf(r) |] ==> \ +\ h(a) = H(a, lam x: r-``{a}. h(x))"; +by (rewtac rew); +by (REPEAT (resolve_tac (prems@[wfrec]) 1)); +val def_wfrec = result(); + +val prems = goal WF.thy + "[| wf(r); a:A; field(r)<=A; \ +\ !!x u. [| x: A; u: Pi(r-``{x}, B) |] ==> H(x,u) : B(x) \ +\ |] ==> wfrec(r,a,H) : B(a)"; +by (res_inst_tac [("a","a")] wf_induct2 1); +by (rtac (wfrec RS ssubst) 4); +by (REPEAT (ares_tac (prems@[lam_type]) 1 + ORELSE eresolve_tac [spec RS mp, underD] 1)); +val wfrec_type = result(); + +val prems = goalw WF.thy [wfrec_def,wftrec_def,the_recfun_def,is_recfun_def] + "[| r=r'; !!x u. H(x,u)=H'(x,u); a=a' |] \ +\ ==> wfrec(r,a,H)=wfrec(r',a',H')"; +by (EVERY1 (map rtac (prems RL [subst]))); +by (SIMP_TAC (wf_ss addrews (prems RL [sym])) 1); +val wfrec_cong = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/WF.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/WF.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,30 @@ +(* Title: ZF/wf.thy + ID: $Id$ + Author: Tobias Nipkow and Lawrence C Paulson + Copyright 1992 University of Cambridge + +Well-founded Recursion +*) + +WF = Trancl + +consts + wf :: "i=>o" + wftrec,wfrec :: "[i, i, [i,i]=>i] =>i" + is_recfun :: "[i, i, [i,i]=>i, i] =>o" + the_recfun :: "[i, i, [i,i]=>i] =>i" + +rules + (*r is a well-founded relation*) + wf_def "wf(r) == ALL Z. Z=0 | (EX x:Z. ALL y. :r --> ~ y:Z)" + + is_recfun_def "is_recfun(r,a,H,f) == \ +\ (f = (lam x: r-``{a}. H(x, restrict(f, r-``{x}))))" + + the_recfun_def "the_recfun(r,a,H) == (THE f.is_recfun(r,a,H,f))" + + wftrec_def "wftrec(r,a,H) == H(a, the_recfun(r,a,H))" + + (*public version. Does not require r to be transitive*) + wfrec_def "wfrec(r,a,H) == wftrec(r^+, a, %x f. H(x, restrict(f,r-``{x})))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ZF.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ZF.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,455 @@ +(* Title: ZF/zf.ML + ID: $Id$ + Author: Lawrence C Paulson and Martin D Coen, CU Computer Laboratory + Copyright 1992 University of Cambridge + +Basic introduction and elimination rules for Zermelo-Fraenkel Set Theory +*) + +open ZF; + +signature ZF_LEMMAS = + sig + val ballE : thm + val ballI : thm + val ball_cong : thm + val ball_rew : thm + val ball_tac : int -> tactic + val basic_ZF_congs : thm list + val bexCI : thm + val bexE : thm + val bexI : thm + val bex_cong : thm + val bspec : thm + val CollectD1 : thm + val CollectD2 : thm + val CollectE : thm + val CollectI : thm + val Collect_cong : thm + val emptyE : thm + val empty_subsetI : thm + val equalityCE : thm + val equalityD1 : thm + val equalityD2 : thm + val equalityE : thm + val equalityI : thm + val equality_iffI : thm + val equals0D : thm + val equals0I : thm + val ex1_functional : thm + val InterD : thm + val InterE : thm + val InterI : thm + val INT_E : thm + val INT_I : thm + val lemmas_cs : claset + val PowD : thm + val PowI : thm + val prove_cong_tac : thm list -> int -> tactic + val RepFunE : thm + val RepFunI : thm + val RepFun_eqI : thm + val RepFun_cong : thm + val ReplaceE : thm + val ReplaceI : thm + val Replace_iff : thm + val Replace_cong : thm + val rev_ballE : thm + val rev_bspec : thm + val rev_subsetD : thm + val separation : thm + val setup_induction : thm + val set_mp_tac : int -> tactic + val subsetCE : thm + val subsetD : thm + val subsetI : thm + val subset_refl : thm + val subset_trans : thm + val UnionE : thm + val UnionI : thm + val UN_E : thm + val UN_I : thm + end; + + +structure ZF_Lemmas : ZF_LEMMAS = +struct + +val basic_ZF_congs = mk_congs ZF.thy + ["op `", "op ``", "op Int", "op Un", "op -", "op <=", "op :", + "Pow", "Union", "Inter", "fst", "snd", "succ", "Pair", "Upair", "cons", + "domain", "range", "restrict"]; + +fun prove_cong_tac prems i = + REPEAT (ares_tac (prems@[refl]@FOL_congs@basic_ZF_congs) i); + +(*** Bounded universal quantifier ***) + +val ballI = prove_goalw ZF.thy [Ball_def] + "[| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)" + (fn prems=> [ (REPEAT (ares_tac (prems @ [allI,impI]) 1)) ]); + +val bspec = prove_goalw ZF.thy [Ball_def] + "[| ALL x:A. P(x); x: A |] ==> P(x)" + (fn major::prems=> + [ (rtac (major RS spec RS mp) 1), + (resolve_tac prems 1) ]); + +val ballE = prove_goalw ZF.thy [Ball_def] + "[| ALL x:A. P(x); P(x) ==> Q; ~ x:A ==> Q |] ==> Q" + (fn major::prems=> + [ (rtac (major RS allE) 1), + (REPEAT (eresolve_tac (prems@[asm_rl,impCE]) 1)) ]); + +(*Used in the datatype package*) +val rev_bspec = prove_goal ZF.thy + "!!x A P. [| x: A; ALL x:A. P(x) |] ==> P(x)" + (fn _ => + [ REPEAT (ares_tac [bspec] 1) ]); + +(*Instantiates x first: better for automatic theorem proving?*) +val rev_ballE = prove_goal ZF.thy + "[| ALL x:A. P(x); ~ x:A ==> Q; P(x) ==> Q |] ==> Q" + (fn major::prems=> + [ (rtac (major RS ballE) 1), + (REPEAT (eresolve_tac prems 1)) ]); + +(*Takes assumptions ALL x:A.P(x) and a:A; creates assumption P(a)*) +val ball_tac = dtac bspec THEN' assume_tac; + +(*Trival rewrite rule; (ALL x:A.P)<->P holds only if A is nonempty!*) +val ball_rew = prove_goal ZF.thy "(ALL x:A. True) <-> True" + (fn prems=> [ (REPEAT (ares_tac [TrueI,ballI,iffI] 1)) ]); + +(*Congruence rule for rewriting*) +val ball_cong = prove_goalw ZF.thy [Ball_def] + "[| A=A'; !!x. x:A' ==> P(x) <-> P'(x) \ +\ |] ==> (ALL x:A. P(x)) <-> (ALL x:A'. P'(x))" + (fn prems=> [ (prove_cong_tac prems 1) ]); + +(*** Bounded existential quantifier ***) + +val bexI = prove_goalw ZF.thy [Bex_def] + "[| P(x); x: A |] ==> EX x:A. P(x)" + (fn prems=> [ (REPEAT (ares_tac (prems @ [exI,conjI]) 1)) ]); + +(*Not of the general form for such rules; ~EX has become ALL~ *) +val bexCI = prove_goal ZF.thy + "[| ALL x:A. ~P(x) ==> P(a); a: A |] ==> EX x:A.P(x)" + (fn prems=> + [ (rtac classical 1), + (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ]); + +val bexE = prove_goalw ZF.thy [Bex_def] + "[| EX x:A. P(x); !!x. [| x:A; P(x) |] ==> Q \ +\ |] ==> Q" + (fn major::prems=> + [ (rtac (major RS exE) 1), + (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)) ]); + +(*We do not even have (EX x:A. True) <-> True unless A is nonempty!!*) + +val bex_cong = prove_goalw ZF.thy [Bex_def] + "[| A=A'; !!x. x:A' ==> P(x) <-> P'(x) \ +\ |] ==> (EX x:A. P(x)) <-> (EX x:A'. P'(x))" + (fn prems=> [ (prove_cong_tac prems 1) ]); + +(*** Rules for subsets ***) + +val subsetI = prove_goalw ZF.thy [subset_def] + "(!!x.x:A ==> x:B) ==> A <= B" + (fn prems=> [ (REPEAT (ares_tac (prems @ [ballI]) 1)) ]); + +(*Rule in Modus Ponens style [was called subsetE] *) +val subsetD = prove_goalw ZF.thy [subset_def] "[| A <= B; c:A |] ==> c:B" + (fn major::prems=> + [ (rtac (major RS bspec) 1), + (resolve_tac prems 1) ]); + +(*Classical elimination rule*) +val subsetCE = prove_goalw ZF.thy [subset_def] + "[| A <= B; ~(c:A) ==> P; c:B ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS ballE) 1), + (REPEAT (eresolve_tac prems 1)) ]); + +(*Takes assumptions A<=B; c:A and creates the assumption c:B *) +val set_mp_tac = dtac subsetD THEN' assume_tac; + +(*Sometimes useful with premises in this order*) +val rev_subsetD = prove_goal ZF.thy "!!A B c. [| c:A; A<=B |] ==> c:B" + (fn _=> [REPEAT (ares_tac [subsetD] 1)]); + +val subset_refl = prove_goal ZF.thy "A <= A" + (fn _=> [ (rtac subsetI 1), atac 1 ]); + +val subset_trans = prove_goal ZF.thy "[| A<=B; B<=C |] ==> A<=C" + (fn prems=> [ (REPEAT (ares_tac ([subsetI]@(prems RL [subsetD])) 1)) ]); + + +(*** Rules for equality ***) + +(*Anti-symmetry of the subset relation*) +val equalityI = prove_goal ZF.thy "[| A <= B; B <= A |] ==> A = B" + (fn prems=> [ (REPEAT (resolve_tac (prems@[conjI, extension RS iffD2]) 1)) ]); + +val equality_iffI = prove_goal ZF.thy "(!!x. x:A <-> x:B) ==> A = B" + (fn [prem] => + [ (rtac equalityI 1), + (REPEAT (ares_tac [subsetI, prem RS iffD1, prem RS iffD2] 1)) ]); + +val equalityD1 = prove_goal ZF.thy "A = B ==> A<=B" + (fn prems=> + [ (rtac (extension RS iffD1 RS conjunct1) 1), + (resolve_tac prems 1) ]); + +val equalityD2 = prove_goal ZF.thy "A = B ==> B<=A" + (fn prems=> + [ (rtac (extension RS iffD1 RS conjunct2) 1), + (resolve_tac prems 1) ]); + +val equalityE = prove_goal ZF.thy + "[| A = B; [| A<=B; B<=A |] ==> P |] ==> P" + (fn prems=> + [ (DEPTH_SOLVE (resolve_tac (prems@[equalityD1,equalityD2]) 1)) ]); + +val equalityCE = prove_goal ZF.thy + "[| A = B; [| c:A; c:B |] ==> P; [| ~ c:A; ~ c:B |] ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS equalityE) 1), + (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1)) ]); + +(*Lemma for creating induction formulae -- for "pattern matching" on p + To make the induction hypotheses usable, apply "spec" or "bspec" to + put universal quantifiers over the free variables in p. + Would it be better to do subgoal_tac "ALL z. p = f(z) --> R(z)" ??*) +val setup_induction = prove_goal ZF.thy + "[| p: A; !!z. z: A ==> p=z --> R |] ==> R" + (fn prems=> + [ (rtac mp 1), + (REPEAT (resolve_tac (refl::prems) 1)) ]); + + +(*** Rules for Replace -- the derived form of replacement ***) + +val ex1_functional = prove_goal ZF.thy + "[| EX! z. P(a,z); P(a,b); P(a,c) |] ==> b = c" + (fn prems=> + [ (cut_facts_tac prems 1), + (best_tac FOL_dup_cs 1) ]); + +val Replace_iff = prove_goalw ZF.thy [Replace_def] + "b : {y. x:A, P(x,y)} <-> (EX x:A. P(x,b) & (ALL y. P(x,y) --> y=b))" + (fn _=> + [ (rtac (replacement RS iff_trans) 1), + (REPEAT (ares_tac [refl,bex_cong,iffI,ballI,allI,conjI,impI,ex1I] 1 + ORELSE eresolve_tac [conjE, spec RS mp, ex1_functional] 1)) ]); + +(*Introduction; there must be a unique y such that P(x,y), namely y=b. *) +val ReplaceI = prove_goal ZF.thy + "[| x: A; P(x,b); !!y. P(x,y) ==> y=b |] ==> \ +\ b : {y. x:A, P(x,y)}" + (fn prems=> + [ (rtac (Replace_iff RS iffD2) 1), + (REPEAT (ares_tac (prems@[bexI,conjI,allI,impI]) 1)) ]); + +(*Elimination; may asssume there is a unique y such that P(x,y), namely y=b. *) +val ReplaceE = prove_goal ZF.thy + "[| b : {y. x:A, P(x,y)}; \ +\ !!x. [| x: A; P(x,b); ALL y. P(x,y)-->y=b |] ==> R \ +\ |] ==> R" + (fn prems=> + [ (rtac (Replace_iff RS iffD1 RS bexE) 1), + (etac conjE 2), + (REPEAT (ares_tac prems 1)) ]); + +val Replace_cong = prove_goal ZF.thy + "[| A=B; !!x y. x:B ==> P(x,y) <-> Q(x,y) |] ==> \ +\ {y. x:A, P(x,y)} = {y. x:B, Q(x,y)}" + (fn prems=> + let val substprems = prems RL [subst, ssubst] + and iffprems = prems RL [iffD1,iffD2] + in [ (rtac equalityI 1), + (REPEAT (eresolve_tac (substprems@[asm_rl, ReplaceE, spec RS mp]) 1 + ORELSE resolve_tac [subsetI, ReplaceI] 1 + ORELSE (resolve_tac iffprems 1 THEN assume_tac 2))) ] + end); + + +(*** Rules for RepFun ***) + +val RepFunI = prove_goalw ZF.thy [RepFun_def] + "!!a A. a : A ==> f(a) : {f(x). x:A}" + (fn _ => [ (REPEAT (ares_tac [ReplaceI,refl] 1)) ]); + +(*Useful for co-induction proofs*) +val RepFun_eqI = prove_goal ZF.thy + "!!b a f. [| b=f(a); a : A |] ==> b : {f(x). x:A}" + (fn _ => [ etac ssubst 1, etac RepFunI 1 ]); + +val RepFunE = prove_goalw ZF.thy [RepFun_def] + "[| b : {f(x). x:A}; \ +\ !!x.[| x:A; b=f(x) |] ==> P |] ==> \ +\ P" + (fn major::prems=> + [ (rtac (major RS ReplaceE) 1), + (REPEAT (ares_tac prems 1)) ]); + +val RepFun_cong = prove_goalw ZF.thy [RepFun_def] + "[| A=B; !!x. x:B ==> f(x)=g(x) |] ==> \ +\ {f(x). x:A} = {g(x). x:B}" + (fn prems=> [ (prove_cong_tac (prems@[Replace_cong]) 1) ]); + + +(*** Rules for Collect -- forming a subset by separation ***) + +(*Separation is derivable from Replacement*) +val separation = prove_goalw ZF.thy [Collect_def] + "a : {x:A. P(x)} <-> a:A & P(a)" + (fn _=> [ (fast_tac (FOL_cs addIs [bexI,ReplaceI] + addSEs [bexE,ReplaceE]) 1) ]); + +val CollectI = prove_goal ZF.thy + "[| a:A; P(a) |] ==> a : {x:A. P(x)}" + (fn prems=> + [ (rtac (separation RS iffD2) 1), + (REPEAT (resolve_tac (prems@[conjI]) 1)) ]); + +val CollectE = prove_goal ZF.thy + "[| a : {x:A. P(x)}; [| a:A; P(a) |] ==> R |] ==> R" + (fn prems=> + [ (rtac (separation RS iffD1 RS conjE) 1), + (REPEAT (ares_tac prems 1)) ]); + +val CollectD1 = prove_goal ZF.thy "a : {x:A. P(x)} ==> a:A" + (fn [major]=> + [ (rtac (major RS CollectE) 1), + (assume_tac 1) ]); + +val CollectD2 = prove_goal ZF.thy "a : {x:A. P(x)} ==> P(a)" + (fn [major]=> + [ (rtac (major RS CollectE) 1), + (assume_tac 1) ]); + +val Collect_cong = prove_goalw ZF.thy [Collect_def] + "[| A=B; !!x. x:B ==> P(x) <-> Q(x) |] ==> \ +\ {x:A. P(x)} = {x:B. Q(x)}" + (fn prems=> [ (prove_cong_tac (prems@[Replace_cong]) 1) ]); + +(*** Rules for Unions ***) + +(*The order of the premises presupposes that C is rigid; A may be flexible*) +val UnionI = prove_goal ZF.thy "[| B: C; A: B |] ==> A: Union(C)" + (fn prems=> + [ (resolve_tac [union_iff RS iffD2] 1), + (REPEAT (resolve_tac (prems @ [bexI]) 1)) ]); + +val UnionE = prove_goal ZF.thy + "[| A : Union(C); !!B.[| A: B; B: C |] ==> R |] ==> R" + (fn prems=> + [ (resolve_tac [union_iff RS iffD1 RS bexE] 1), + (REPEAT (ares_tac prems 1)) ]); + +(*** Rules for Inter ***) + +(*Not obviously useful towards proving InterI, InterD, InterE*) +val Inter_iff = prove_goalw ZF.thy [Inter_def,Ball_def] + "A : Inter(C) <-> (ALL x:C. A: x) & (EX x. x:C)" + (fn _=> [ (rtac (separation RS iff_trans) 1), + (fast_tac (FOL_cs addIs [UnionI] addSEs [UnionE]) 1) ]); + +(* Intersection is well-behaved only if the family is non-empty! *) +val InterI = prove_goalw ZF.thy [Inter_def] + "[| !!x. x: C ==> A: x; c:C |] ==> A : Inter(C)" + (fn prems=> + [ (DEPTH_SOLVE (ares_tac ([CollectI,UnionI,ballI] @ prems) 1)) ]); + +(*A "destruct" rule -- every B in C contains A as an element, but + A:B can hold when B:C does not! This rule is analogous to "spec". *) +val InterD = prove_goalw ZF.thy [Inter_def] + "[| A : Inter(C); B : C |] ==> A : B" + (fn [major,minor]=> + [ (rtac (major RS CollectD2 RS bspec) 1), + (rtac minor 1) ]); + +(*"Classical" elimination rule -- does not require exhibiting B:C *) +val InterE = prove_goalw ZF.thy [Inter_def] + "[| A : Inter(C); A:B ==> R; ~ B:C ==> R |] ==> R" + (fn major::prems=> + [ (rtac (major RS CollectD2 RS ballE) 1), + (REPEAT (eresolve_tac prems 1)) ]); + +(*** Rules for Unions of families ***) +(* UN x:A. B(x) abbreviates Union({B(x). x:A}) *) + +(*The order of the premises presupposes that A is rigid; b may be flexible*) +val UN_I = prove_goal ZF.thy "[| a: A; b: B(a) |] ==> b: (UN x:A. B(x))" + (fn prems=> + [ (REPEAT (resolve_tac (prems@[UnionI,RepFunI]) 1)) ]); + +val UN_E = prove_goal ZF.thy + "[| b : (UN x:A. B(x)); !!x.[| x: A; b: B(x) |] ==> R |] ==> R" + (fn major::prems=> + [ (rtac (major RS UnionE) 1), + (REPEAT (eresolve_tac (prems@[asm_rl, RepFunE, subst]) 1)) ]); + + +(*** Rules for Intersections of families ***) +(* INT x:A. B(x) abbreviates Inter({B(x). x:A}) *) + +val INT_I = prove_goal ZF.thy + "[| !!x. x: A ==> b: B(x); a: A |] ==> b: (INT x:A. B(x))" + (fn prems=> + [ (REPEAT (ares_tac (prems@[InterI,RepFunI]) 1 + ORELSE eresolve_tac [RepFunE,ssubst] 1)) ]); + +val INT_E = prove_goal ZF.thy + "[| b : (INT x:A. B(x)); a: A |] ==> b : B(a)" + (fn [major,minor]=> + [ (rtac (major RS InterD) 1), + (rtac (minor RS RepFunI) 1) ]); + + +(*** Rules for Powersets ***) + +val PowI = prove_goal ZF.thy "A <= B ==> A : Pow(B)" + (fn [prem]=> [ (rtac (prem RS (power_set RS iffD2)) 1) ]); + +val PowD = prove_goal ZF.thy "A : Pow(B) ==> A<=B" + (fn [major]=> [ (rtac (major RS (power_set RS iffD1)) 1) ]); + + +(*** Rules for the empty set ***) + +(*The set {x:0.False} is empty; by foundation it equals 0 + See Suppes, page 21.*) +val emptyE = prove_goal ZF.thy "a:0 ==> P" + (fn [major]=> + [ (rtac (foundation RS disjE) 1), + (etac (equalityD2 RS subsetD RS CollectD2 RS FalseE) 1), + (rtac major 1), + (etac bexE 1), + (etac (CollectD2 RS FalseE) 1) ]); + +val empty_subsetI = prove_goal ZF.thy "0 <= A" + (fn _ => [ (REPEAT (ares_tac [equalityI,subsetI,emptyE] 1)) ]); + +val equals0I = prove_goal ZF.thy "[| !!y. y:A ==> False |] ==> A=0" + (fn prems=> + [ (REPEAT (ares_tac (prems@[empty_subsetI,subsetI,equalityI]) 1 + ORELSE eresolve_tac (prems RL [FalseE]) 1)) ]); + +val equals0D = prove_goal ZF.thy "[| A=0; a:A |] ==> P" + (fn [major,minor]=> + [ (rtac (minor RS (major RS equalityD1 RS subsetD RS emptyE)) 1) ]); + +val lemmas_cs = FOL_cs + addSIs [ballI, InterI, CollectI, PowI, subsetI] + addIs [bexI, UnionI, ReplaceI, RepFunI] + addSEs [bexE, make_elim PowD, UnionE, ReplaceE, RepFunE, + CollectE, emptyE] + addEs [rev_ballE, InterD, make_elim InterD, subsetD, subsetCE]; + +end; + +open ZF_Lemmas; diff -r 000000000000 -r a5a9c433f639 src/ZF/ZF.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ZF.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,227 @@ +(* Title: ZF/zf.thy + ID: $Id$ + Author: Lawrence C Paulson and Martin D Coen, CU Computer Laboratory + Copyright 1993 University of Cambridge + +Zermelo-Fraenkel Set Theory +*) + +ZF = FOL + + +types + i, is, syntax 0 + +arities + i :: term + + +consts + + "0" :: "i" ("0") (*the empty set*) + Pow :: "i => i" (*power sets*) + Inf :: "i" (*infinite set*) + + (* Bounded Quantifiers *) + + "@Ball" :: "[idt, i, o] => o" ("(3ALL _:_./ _)" 10) + "@Bex" :: "[idt, i, o] => o" ("(3EX _:_./ _)" 10) + Ball :: "[i, i => o] => o" + Bex :: "[i, i => o] => o" + + (* General Union and Intersection *) + + "@INTER" :: "[idt, i, i] => i" ("(3INT _:_./ _)" 10) + "@UNION" :: "[idt, i, i] => i" ("(3UN _:_./ _)" 10) + Union, Inter :: "i => i" + + (* Variations on Replacement *) + + "@Replace" :: "[idt, idt, i, o] => i" ("(1{_ ./ _: _, _})") + "@RepFun" :: "[i, idt, i] => i" ("(1{_ ./ _: _})") + "@Collect" :: "[idt, i, o] => i" ("(1{_: _ ./ _})") + PrimReplace :: "[i, [i, i] => o] => i" + Replace :: "[i, [i, i] => o] => i" + RepFun :: "[i, i => i] => i" + Collect :: "[i, i => o] => i" + + (* Descriptions *) + + "@THE" :: "[idt, o] => i" ("(3THE _./ _)" 10) + The :: "[i => o] => i" + if :: "[o, i, i] => i" + + (* Enumerations of type i *) + + "" :: "i => is" ("_") + "@Enum" :: "[i, is] => is" ("_,/ _") + + (* Finite Sets *) + + "@Finset" :: "is => i" ("{(_)}") + Upair, cons :: "[i, i] => i" + succ :: "i => i" + + (* Ordered Pairing and n-Tuples *) + + "@Tuple" :: "[i, is] => i" ("<(_,/ _)>") + PAIR :: "syntax" + Pair :: "[i, i] => i" + fst, snd :: "i => i" + split :: "[[i,i] => i, i] => i" + fsplit :: "[[i,i] => o, i] => o" + + (* Sigma and Pi Operators *) + + "@PROD" :: "[idt, i, i] => i" ("(3PROD _:_./ _)" 10) + "@SUM" :: "[idt, i, i] => i" ("(3SUM _:_./ _)" 10) + "@lam" :: "[idt, i, i] => i" ("(3lam _:_./ _)" 10) + Pi, Sigma :: "[i, i => i] => i" + + (* Relations and Functions *) + + domain :: "i => i" + range :: "i => i" + field :: "i => i" + converse :: "i => i" + Lambda :: "[i, i => i] => i" + restrict :: "[i, i] => i" + + (* Infixes in order of decreasing precedence *) + + "``" :: "[i, i] => i" (infixl 90) (*image*) + "-``" :: "[i, i] => i" (infixl 90) (*inverse image*) + "`" :: "[i, i] => i" (infixl 90) (*function application*) + + (*Except for their translations, * and -> are right-associating infixes*) + " *" :: "[i, i] => i" ("(_ */ _)" [81, 80] 80) (*Cartesian product*) + "Int" :: "[i, i] => i" (infixl 70) (*binary intersection*) + "Un" :: "[i, i] => i" (infixl 65) (*binary union*) + "-" :: "[i, i] => i" (infixl 65) (*set difference*) + " ->" :: "[i, i] => i" ("(_ ->/ _)" [61, 60] 60) (*function space*) + "<=" :: "[i, i] => o" (infixl 50) (*subset relation*) + ":" :: "[i, i] => o" (infixl 50) (*membership relation*) + + +translations + "{x, xs}" == "cons(x, {xs})" + "{x}" == "cons(x, 0)" + + "PAIR(x, Pair(y, z))" <= "Pair(x, Pair(y, z))" + "PAIR(x, PAIR(y, z))" <= "Pair(x, PAIR(y, z))" + "" <= "PAIR(x, )" + "" == "Pair(x, )" + "" == "Pair(x, y)" + + "{x:A. P}" == "Collect(A, %x. P)" + "{y. x:A, Q}" == "Replace(A, %x y. Q)" + "{f. x:A}" == "RepFun(A, %x. f)" + "INT x:A. B" == "Inter({B. x:A})" + "UN x:A. B" == "Union({B. x:A})" + "PROD x:A. B" => "Pi(A, %x. B)" + "SUM x:A. B" => "Sigma(A, %x. B)" + "THE x. P" == "The(%x. P)" + "lam x:A. f" == "Lambda(A, %x. f)" + "ALL x:A. P" == "Ball(A, %x. P)" + "EX x:A. P" == "Bex(A, %x. P)" + + +rules + + (* Bounded Quantifiers *) +Ball_def "Ball(A,P) == ALL x. x:A --> P(x)" +Bex_def "Bex(A,P) == EX x. x:A & P(x)" +subset_def "A <= B == ALL x:A. x:B" + + (* ZF axioms -- see Suppes p.238 + Axioms for Union, Pow and Replace state existence only, + uniqueness is derivable using extensionality. *) + +extension "A = B <-> A <= B & B <= A" +union_iff "A : Union(C) <-> (EX B:C. A:B)" +power_set "A : Pow(B) <-> A <= B" +succ_def "succ(i) == cons(i,i)" + + (*We may name this set, though it is not uniquely defined. *) +infinity "0:Inf & (ALL y:Inf. succ(y): Inf)" + + (*This formulation facilitates case analysis on A. *) +foundation "A=0 | (EX x:A. ALL y:x. ~ y:A)" + + (* Schema axiom since predicate P is a higher-order variable *) +replacement "(ALL x:A. ALL y z. P(x,y) & P(x,z) --> y=z) ==> \ +\ b : PrimReplace(A,P) <-> (EX x:A. P(x,b))" + + (* Derived form of replacement, restricting P to its functional part. + The resulting set (for functional P) is the same as with + PrimReplace, but the rules are simpler. *) +Replace_def "Replace(A,P) == PrimReplace(A, %x y. (EX!z.P(x,z)) & P(x,y))" + + (* Functional form of replacement -- analgous to ML's map functional *) +RepFun_def "RepFun(A,f) == {y . x:A, y=f(x)}" + + (* Separation and Pairing can be derived from the Replacement + and Powerset Axioms using the following definitions. *) + +Collect_def "Collect(A,P) == {y . x:A, x=y & P(x)}" + + (*Unordered pairs (Upair) express binary union/intersection and cons; + set enumerations translate as {a,...,z} = cons(a,...,cons(z,0)...) *) +Upair_def "Upair(a,b) == {y. x:Pow(Pow(0)), (x=0 & y=a) | (x=Pow(0) & y=b)}" +cons_def "cons(a,A) == Upair(a,a) Un A" + + (* Difference, general intersection, binary union and small intersection *) + +Diff_def "A - B == { x:A . ~(x:B) }" +Inter_def "Inter(A) == { x:Union(A) . ALL y:A. x:y}" +Un_def "A Un B == Union(Upair(A,B))" +Int_def "A Int B == Inter(Upair(A,B))" + + (* Definite descriptions -- via Replace over the set "1" *) + +the_def "The(P) == Union({y . x:{0}, P(y)})" +if_def "if(P,a,b) == THE z. P & z=a | ~P & z=b" + + (* Ordered pairs and disjoint union of a family of sets *) + + (* this "symmetric" definition works better than {{a}, {a,b}} *) +Pair_def " == {{a,a}, {a,b}}" +fst_def "fst == split(%x y.x)" +snd_def "snd == split(%x y.y)" +split_def "split(c,p) == THE y. EX a b. p= & y=c(a,b)" +fsplit_def "fsplit(R,z) == EX x y. z= & R(x,y)" +Sigma_def "Sigma(A,B) == UN x:A. UN y:B(x). {}" + + (* Operations on relations *) + +(*converse of relation r, inverse of function*) +converse_def "converse(r) == {z. w:r, EX x y. w= & z=}" + +domain_def "domain(r) == {x. w:r, EX y. w=}" +range_def "range(r) == domain(converse(r))" +field_def "field(r) == domain(r) Un range(r)" +image_def "r `` A == {y : range(r) . EX x:A. : r}" +vimage_def "r -`` A == converse(r)``A" + + (* Abstraction, application and Cartesian product of a family of sets *) + +lam_def "Lambda(A,b) == { . x:A}" +apply_def "f`a == THE y. : f" +Pi_def "Pi(A,B) == {f: Pow(Sigma(A,B)). ALL x:A. EX! y. : f}" + + (* Restrict the function f to the domain A *) +restrict_def "restrict(f,A) == lam x:A.f`x" + +end + + +ML + +(* 'Dependent' type operators *) + +val parse_translation = + [(" ->", ndependent_tr "Pi"), + (" *", ndependent_tr "Sigma")]; + +val print_translation = + [("Pi", dependent_tr' ("@PROD", " ->")), + ("Sigma", dependent_tr' ("@SUM", " *"))]; diff -r 000000000000 -r a5a9c433f639 src/ZF/arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/arith.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,356 @@ +(* Title: ZF/arith.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For arith.thy. Arithmetic operators and their definitions + +Proofs about elementary arithmetic: addition, multiplication, etc. + +Could prove def_rec_0, def_rec_succ... +*) + +open Arith; + +(*"Difference" is subtraction of natural numbers. + There are no negative numbers; we have + m #- n = 0 iff m<=n and m #- n = succ(k) iff m>n. + Also, rec(m, 0, %z w.z) is pred(m). +*) + +(** rec -- better than nat_rec; the succ case has no type requirement! **) + +val rec_trans = rec_def RS def_transrec RS trans; + +goal Arith.thy "rec(0,a,b) = a"; +by (rtac rec_trans 1); +by (rtac nat_case_0 1); +val rec_0 = result(); + +goal Arith.thy "rec(succ(m),a,b) = b(m, rec(m,a,b))"; +val rec_ss = ZF_ss + addcongs (mk_typed_congs Arith.thy [("b", "[i,i]=>i")]) + addrews [nat_case_succ, nat_succI]; +by (rtac rec_trans 1); +by (SIMP_TAC rec_ss 1); +val rec_succ = result(); + +val major::prems = goal Arith.thy + "[| n: nat; \ +\ a: C(0); \ +\ !!m z. [| m: nat; z: C(m) |] ==> b(m,z): C(succ(m)) \ +\ |] ==> rec(n,a,b) : C(n)"; +by (rtac (major RS nat_induct) 1); +by (ALLGOALS + (ASM_SIMP_TAC (ZF_ss addrews (prems@[rec_0,rec_succ])))); +val rec_type = result(); + +val prems = goalw Arith.thy [rec_def] + "[| n=n'; a=a'; !!m z. b(m,z)=b'(m,z) \ +\ |] ==> rec(n,a,b)=rec(n',a',b')"; +by (SIMP_TAC (ZF_ss addcongs [transrec_cong,nat_case_cong] + addrews (prems RL [sym])) 1); +val rec_cong = result(); + +val nat_typechecks = [rec_type,nat_0I,nat_1I,nat_succI,Ord_nat]; + +val nat_ss = ZF_ss addcongs [nat_case_cong,rec_cong] + addrews ([rec_0,rec_succ] @ nat_typechecks); + + +(** Addition **) + +val add_type = prove_goalw Arith.thy [add_def] + "[| m:nat; n:nat |] ==> m #+ n : nat" + (fn prems=> [ (typechk_tac (prems@nat_typechecks@ZF_typechecks)) ]); + +val add_0 = prove_goalw Arith.thy [add_def] + "0 #+ n = n" + (fn _ => [ (rtac rec_0 1) ]); + +val add_succ = prove_goalw Arith.thy [add_def] + "succ(m) #+ n = succ(m #+ n)" + (fn _=> [ (rtac rec_succ 1) ]); + +(** Multiplication **) + +val mult_type = prove_goalw Arith.thy [mult_def] + "[| m:nat; n:nat |] ==> m #* n : nat" + (fn prems=> + [ (typechk_tac (prems@[add_type]@nat_typechecks@ZF_typechecks)) ]); + +val mult_0 = prove_goalw Arith.thy [mult_def] + "0 #* n = 0" + (fn _ => [ (rtac rec_0 1) ]); + +val mult_succ = prove_goalw Arith.thy [mult_def] + "succ(m) #* n = n #+ (m #* n)" + (fn _ => [ (rtac rec_succ 1) ]); + +(** Difference **) + +val diff_type = prove_goalw Arith.thy [diff_def] + "[| m:nat; n:nat |] ==> m #- n : nat" + (fn prems=> [ (typechk_tac (prems@nat_typechecks@ZF_typechecks)) ]); + +val diff_0 = prove_goalw Arith.thy [diff_def] + "m #- 0 = m" + (fn _ => [ (rtac rec_0 1) ]); + +val diff_0_eq_0 = prove_goalw Arith.thy [diff_def] + "n:nat ==> 0 #- n = 0" + (fn [prem]=> + [ (rtac (prem RS nat_induct) 1), + (ALLGOALS (ASM_SIMP_TAC nat_ss)) ]); + +(*Must simplify BEFORE the induction!! (Else we get a critical pair) + succ(m) #- succ(n) rewrites to pred(succ(m) #- n) *) +val diff_succ_succ = prove_goalw Arith.thy [diff_def] + "[| m:nat; n:nat |] ==> succ(m) #- succ(n) = m #- n" + (fn prems=> + [ (ASM_SIMP_TAC (nat_ss addrews prems) 1), + (nat_ind_tac "n" prems 1), + (ALLGOALS (ASM_SIMP_TAC (nat_ss addrews prems))) ]); + +val prems = goal Arith.thy + "[| m:nat; n:nat |] ==> m #- n : succ(m)"; +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +by (etac succE 3); +by (ALLGOALS + (ASM_SIMP_TAC + (nat_ss addrews (prems@[diff_0,diff_0_eq_0,diff_succ_succ])))); +val diff_leq = result(); + +(*** Simplification over add, mult, diff ***) + +val arith_typechecks = [add_type, mult_type, diff_type]; +val arith_rews = [add_0, add_succ, + mult_0, mult_succ, + diff_0, diff_0_eq_0, diff_succ_succ]; + +val arith_congs = mk_congs Arith.thy ["op #+", "op #-", "op #*"]; + +val arith_ss = nat_ss addcongs arith_congs + addrews (arith_rews@arith_typechecks); + +(*** Addition ***) + +(*Associative law for addition*) +val add_assoc = prove_goal Arith.thy + "m:nat ==> (m #+ n) #+ k = m #+ (n #+ k)" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews prems))) ]); + +(*The following two lemmas are used for add_commute and sometimes + elsewhere, since they are safe for rewriting.*) +val add_0_right = prove_goal Arith.thy + "m:nat ==> m #+ 0 = m" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews prems))) ]); + +val add_succ_right = prove_goal Arith.thy + "m:nat ==> m #+ succ(n) = succ(m #+ n)" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews prems))) ]); + +(*Commutative law for addition*) +val add_commute = prove_goal Arith.thy + "[| m:nat; n:nat |] ==> m #+ n = n #+ m" + (fn prems=> + [ (nat_ind_tac "n" prems 1), + (ALLGOALS + (ASM_SIMP_TAC + (arith_ss addrews (prems@[add_0_right, add_succ_right])))) ]); + +(*Cancellation law on the left*) +val [knat,eqn] = goal Arith.thy + "[| k:nat; k #+ m = k #+ n |] ==> m=n"; +by (rtac (eqn RS rev_mp) 1); +by (nat_ind_tac "k" [knat] 1); +by (ALLGOALS (SIMP_TAC arith_ss)); +by (fast_tac ZF_cs 1); +val add_left_cancel = result(); + +(*** Multiplication ***) + +(*right annihilation in product*) +val mult_0_right = prove_goal Arith.thy + "m:nat ==> m #* 0 = 0" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews prems))) ]); + +(*right successor law for multiplication*) +val mult_succ_right = prove_goal Arith.thy + "[| m:nat; n:nat |] ==> m #* succ(n) = m #+ (m #* n)" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews ([add_assoc RS sym]@prems)))), + (*The final goal requires the commutative law for addition*) + (REPEAT (ares_tac (prems@[refl,add_commute]@ZF_congs@arith_congs) 1)) ]); + +(*Commutative law for multiplication*) +val mult_commute = prove_goal Arith.thy + "[| m:nat; n:nat |] ==> m #* n = n #* m" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC + (arith_ss addrews (prems@[mult_0_right, mult_succ_right])))) ]); + +(*addition distributes over multiplication*) +val add_mult_distrib = prove_goal Arith.thy + "[| m:nat; k:nat |] ==> (m #+ n) #* k = (m #* k) #+ (n #* k)" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews ([add_assoc RS sym]@prems)))) ]); + +(*Distributive law on the left; requires an extra typing premise*) +val add_mult_distrib_left = prove_goal Arith.thy + "[| m:nat; n:nat; k:nat |] ==> k #* (m #+ n) = (k #* m) #+ (k #* n)" + (fn prems=> + let val mult_commute' = read_instantiate [("m","k")] mult_commute + val ss = arith_ss addrews ([mult_commute',add_mult_distrib]@prems) + in [ (SIMP_TAC ss 1) ] + end); + +(*Associative law for multiplication*) +val mult_assoc = prove_goal Arith.thy + "[| m:nat; n:nat; k:nat |] ==> (m #* n) #* k = m #* (n #* k)" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews (prems@[add_mult_distrib])))) ]); + + +(*** Difference ***) + +val diff_self_eq_0 = prove_goal Arith.thy + "m:nat ==> m #- m = 0" + (fn prems=> + [ (nat_ind_tac "m" prems 1), + (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews prems))) ]); + +(*Addition is the inverse of subtraction: if n<=m then n+(m-n) = m. *) +val notless::prems = goal Arith.thy + "[| ~m:n; m:nat; n:nat |] ==> n #+ (m#-n) = m"; +by (rtac (notless RS rev_mp) 1); +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +by (ALLGOALS (ASM_SIMP_TAC + (arith_ss addrews (prems@[succ_mem_succ_iff, Ord_0_mem_succ, + naturals_are_ordinals])))); +val add_diff_inverse = result(); + + +(*Subtraction is the inverse of addition. *) +val [mnat,nnat] = goal Arith.thy + "[| m:nat; n:nat |] ==> (n#+m) #-n = m"; +by (rtac (nnat RS nat_induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews [mnat]))); +val diff_add_inverse = result(); + +val [mnat,nnat] = goal Arith.thy + "[| m:nat; n:nat |] ==> n #- (n#+m) = 0"; +by (rtac (nnat RS nat_induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews [mnat]))); +val diff_add_0 = result(); + + +(*** Remainder ***) + +(*In ordinary notation: if 0 m #- n : m"; +by (cut_facts_tac prems 1); +by (etac rev_mp 1); +by (etac rev_mp 1); +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +by (ALLGOALS (ASM_SIMP_TAC + (nat_ss addrews (prems@[diff_leq,diff_succ_succ])))); +val div_termination = result(); + +val div_rls = + [Ord_transrec_type, apply_type, div_termination, if_type] @ + nat_typechecks; + +(*Type checking depends upon termination!*) +val prems = goalw Arith.thy [mod_def] + "[| 0:n; m:nat; n:nat |] ==> m mod n : nat"; +by (REPEAT (ares_tac (prems @ div_rls) 1 ORELSE etac Ord_trans 1)); +val mod_type = result(); + +val div_ss = ZF_ss addrews [naturals_are_ordinals,div_termination]; + +val prems = goal Arith.thy "[| 0:n; m:n; m:nat; n:nat |] ==> m mod n = m"; +by (rtac (mod_def RS def_transrec RS trans) 1); +by (SIMP_TAC (div_ss addrews prems) 1); +val mod_less = result(); + +val prems = goal Arith.thy + "[| 0:n; ~m:n; m:nat; n:nat |] ==> m mod n = (m#-n) mod n"; +by (rtac (mod_def RS def_transrec RS trans) 1); +by (SIMP_TAC (div_ss addrews prems) 1); +val mod_geq = result(); + +(*** Quotient ***) + +(*Type checking depends upon termination!*) +val prems = goalw Arith.thy [div_def] + "[| 0:n; m:nat; n:nat |] ==> m div n : nat"; +by (REPEAT (ares_tac (prems @ div_rls) 1 ORELSE etac Ord_trans 1)); +val div_type = result(); + +val prems = goal Arith.thy + "[| 0:n; m:n; m:nat; n:nat |] ==> m div n = 0"; +by (rtac (div_def RS def_transrec RS trans) 1); +by (SIMP_TAC (div_ss addrews prems) 1); +val div_less = result(); + +val prems = goal Arith.thy + "[| 0:n; ~m:n; m:nat; n:nat |] ==> m div n = succ((m#-n) div n)"; +by (rtac (div_def RS def_transrec RS trans) 1); +by (SIMP_TAC (div_ss addrews prems) 1); +val div_geq = result(); + +(*Main Result.*) +val prems = goal Arith.thy + "[| 0:n; m:nat; n:nat |] ==> (m div n)#*n #+ m mod n = m"; +by (res_inst_tac [("i","m")] complete_induct 1); +by (resolve_tac prems 1); +by (res_inst_tac [("Q","x:n")] (excluded_middle RS disjE) 1); +by (ALLGOALS + (ASM_SIMP_TAC + (arith_ss addrews ([mod_type,div_type] @ prems @ + [mod_less,mod_geq, div_less, div_geq, + add_assoc, add_diff_inverse, div_termination])))); +val mod_div_equality = result(); + + +(**** Additional theorems about "less than" ****) + +val [mnat,nnat] = goal Arith.thy + "[| m:nat; n:nat |] ==> ~ (m #+ n) : n"; +by (rtac (mnat RS nat_induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (arith_ss addrews [mem_not_refl]))); +by (rtac notI 1); +by (etac notE 1); +by (etac (succI1 RS Ord_trans) 1); +by (rtac (nnat RS naturals_are_ordinals) 1); +val add_not_less_self = result(); + +val [mnat,nnat] = goal Arith.thy + "[| m:nat; n:nat |] ==> m : succ(m #+ n)"; +by (rtac (mnat RS nat_induct) 1); +(*May not simplify even with ZF_ss because it would expand m:succ(...) *) +by (rtac (add_0 RS ssubst) 1); +by (rtac (add_succ RS ssubst) 2); +by (REPEAT (ares_tac [nnat, Ord_0_mem_succ, succ_mem_succI, + naturals_are_ordinals, nat_succI, add_type] 1)); +val add_less_succ_self = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/arith.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/arith.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,26 @@ +(* Title: ZF/arith.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Arithmetic operators and their definitions +*) + +Arith = Epsilon + +consts + rec :: "[i, i, [i,i]=>i]=>i" + "#*" :: "[i,i]=>i" (infixl 70) + div :: "[i,i]=>i" (infixl 70) + mod :: "[i,i]=>i" (infixl 70) + "#+" :: "[i,i]=>i" (infixl 65) + "#-" :: "[i,i]=>i" (infixl 65) + +rules + rec_def "rec(k,a,b) == transrec(k, %n f. nat_case(n, a, %m. b(m, f`m)))" + + add_def "m#+n == rec(m, n, %u v.succ(v))" + diff_def "m#-n == rec(n, m, %u v. rec(v, 0, %x y.x))" + mult_def "m#*n == rec(m, 0, %u v. n #+ v)" + mod_def "m mod n == transrec(m, %j f. if(j:n, j, f`(j#-n)))" + div_def "m div n == transrec(m, %j f. if(j:n, 0, succ(f`(j#-n))))" +end diff -r 000000000000 -r a5a9c433f639 src/ZF/bool.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/bool.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,99 @@ +(* Title: ZF/bool + ID: $Id$ + Author: Martin D Coen, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For ZF/bool.thy. Booleans in Zermelo-Fraenkel Set Theory +*) + +open Bool; + +val bool_defs = [bool_def,one_def,cond_def]; + +(* Introduction rules *) + +goalw Bool.thy bool_defs "1 : bool"; +by (rtac (consI1 RS consI2) 1); +val bool_1I = result(); + +goalw Bool.thy bool_defs "0 : bool"; +by (rtac consI1 1); +val bool_0I = result(); + +goalw Bool.thy bool_defs "~ 1=0"; +by (rtac succ_not_0 1); +val one_not_0 = result(); + +(** 1=0 ==> R **) +val one_neq_0 = one_not_0 RS notE; + +val prems = goalw Bool.thy bool_defs "[| c: bool; P(1); P(0) |] ==> P(c)"; +by (cut_facts_tac prems 1); +by (fast_tac ZF_cs 1); +val boolE = result(); + +(** cond **) + +(*1 means true*) +goalw Bool.thy bool_defs "cond(1,c,d) = c"; +by (rtac (refl RS if_P) 1); +val cond_1 = result(); + +(*0 means false*) +goalw Bool.thy bool_defs "cond(0,c,d) = d"; +by (rtac (succ_not_0 RS not_sym RS if_not_P) 1); +val cond_0 = result(); + +val major::prems = goal Bool.thy + "[| b: bool; c: A(1); d: A(0) |] ==> cond(b,c,d): A(b)"; +by (rtac (major RS boolE) 1); +by (rtac (cond_0 RS ssubst) 2); +by (resolve_tac prems 2); +by (rtac (cond_1 RS ssubst) 1); +by (resolve_tac prems 1); +val cond_type = result(); + +val [cond_cong] = mk_congs Bool.thy ["cond"]; +val bool_congs = mk_congs Bool.thy ["cond","not","op and","op or","op xor"]; + +val [rew] = goal Bool.thy "[| !!b. j(b)==cond(b,c,d) |] ==> j(1) = c"; +by (rewtac rew); +by (rtac cond_1 1); +val def_cond_1 = result(); + +val [rew] = goal Bool.thy "[| !!b. j(b)==cond(b,c,d) |] ==> j(0) = d"; +by (rewtac rew); +by (rtac cond_0 1); +val def_cond_0 = result(); + +fun conds def = [standard (def RS def_cond_1), standard (def RS def_cond_0)]; + +val [not_1,not_0] = conds not_def; + +val [and_1,and_0] = conds and_def; + +val [or_1,or_0] = conds or_def; + +val [xor_1,xor_0] = conds xor_def; + +val not_type = prove_goalw Bool.thy [not_def] + "a:bool ==> not(a) : bool" + (fn prems=> [ (typechk_tac (prems@[bool_1I, bool_0I, cond_type])) ]); + +val and_type = prove_goalw Bool.thy [and_def] + "[| a:bool; b:bool |] ==> a and b : bool" + (fn prems=> [ (typechk_tac (prems@[bool_1I, bool_0I, cond_type])) ]); + +val or_type = prove_goalw Bool.thy [or_def] + "[| a:bool; b:bool |] ==> a or b : bool" + (fn prems=> [ (typechk_tac (prems@[bool_1I, bool_0I, cond_type])) ]); + +val xor_type = prove_goalw Bool.thy [xor_def] + "[| a:bool; b:bool |] ==> a xor b : bool" + (fn prems=> [ (typechk_tac(prems@[bool_1I, bool_0I, cond_type, not_type])) ]); + +val bool_typechecks = [bool_1I, bool_0I, cond_type, not_type, and_type, + or_type, xor_type] + +val bool_rews = [cond_1,cond_0,not_1,not_0,and_1,and_0,or_1,or_0,xor_1,xor_0]; + diff -r 000000000000 -r a5a9c433f639 src/ZF/bool.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/bool.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,27 @@ +(* Title: ZF/bool.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Booleans in Zermelo-Fraenkel Set Theory +*) + +Bool = ZF + +consts + "1" :: "i" ("1") + bool :: "i" + cond :: "[i,i,i]=>i" + not :: "i=>i" + and :: "[i,i]=>i" (infixl 70) + or :: "[i,i]=>i" (infixl 65) + xor :: "[i,i]=>i" (infixl 65) + +rules + one_def "1 == succ(0)" + bool_def "bool == {0,1}" + cond_def "cond(b,c,d) == if(b=1,c,d)" + not_def "not(b) == cond(b,0,1)" + and_def "a and b == cond(a,b,0)" + or_def "a or b == cond(a,1,b)" + xor_def "a xor b == cond(a,not(b),b)" +end diff -r 000000000000 -r a5a9c433f639 src/ZF/co-inductive.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/co-inductive.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,67 @@ +(* Title: ZF/co-inductive.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Co-inductive Definitions for Zermelo-Fraenkel Set Theory + +Uses greatest fixedpoints with Quine-inspired products and sums + +Sums are used only for mutual recursion; +Products are used only to derive "streamlined" induction rules for relations +*) + +structure Gfp = + struct + val oper = Const("gfp", [iT,iT-->iT]--->iT) + val bnd_mono = Const("bnd_mono", [iT,iT-->iT]--->oT) + val bnd_monoI = bnd_monoI + val subs = def_gfp_subset + val Tarski = def_gfp_Tarski + val induct = def_Collect_coinduct + end; + +structure Quine_Prod = + struct + val sigma = Const("QSigma", [iT, iT-->iT]--->iT) + val pair = Const("QPair", [iT,iT]--->iT) + val split_const = Const("qsplit", [[iT,iT]--->iT, iT]--->iT) + val fsplit_const = Const("qfsplit", [[iT,iT]--->oT, iT]--->oT) + val pair_iff = QPair_iff + val split_eq = qsplit + val fsplitI = qfsplitI + val fsplitD = qfsplitD + val fsplitE = qfsplitE + end; + +structure Quine_Sum = + struct + val sum = Const("op <+>", [iT,iT]--->iT) + val inl = Const("QInl", iT-->iT) + val inr = Const("QInr", iT-->iT) + val elim = Const("qcase", [iT-->iT, iT-->iT, iT]--->iT) + val case_inl = qcase_QInl + val case_inr = qcase_QInr + val inl_iff = QInl_iff + val inr_iff = QInr_iff + val distinct = QInl_QInr_iff + val distinct' = QInr_QInl_iff + end; + +signature CO_INDRULE = + sig + val co_induct : thm + end; + + +functor Co_Inductive_Fun (Ind: INDUCTIVE) + : sig include INTR_ELIM CO_INDRULE end = +struct +structure Intr_elim = + Intr_elim_Fun(structure Ind=Ind and Fp=Gfp and + Pr=Quine_Prod and Su=Quine_Sum); + +open Intr_elim +val co_induct = raw_induct +end; + diff -r 000000000000 -r a5a9c433f639 src/ZF/coinductive.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/coinductive.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,67 @@ +(* Title: ZF/co-inductive.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Co-inductive Definitions for Zermelo-Fraenkel Set Theory + +Uses greatest fixedpoints with Quine-inspired products and sums + +Sums are used only for mutual recursion; +Products are used only to derive "streamlined" induction rules for relations +*) + +structure Gfp = + struct + val oper = Const("gfp", [iT,iT-->iT]--->iT) + val bnd_mono = Const("bnd_mono", [iT,iT-->iT]--->oT) + val bnd_monoI = bnd_monoI + val subs = def_gfp_subset + val Tarski = def_gfp_Tarski + val induct = def_Collect_coinduct + end; + +structure Quine_Prod = + struct + val sigma = Const("QSigma", [iT, iT-->iT]--->iT) + val pair = Const("QPair", [iT,iT]--->iT) + val split_const = Const("qsplit", [[iT,iT]--->iT, iT]--->iT) + val fsplit_const = Const("qfsplit", [[iT,iT]--->oT, iT]--->oT) + val pair_iff = QPair_iff + val split_eq = qsplit + val fsplitI = qfsplitI + val fsplitD = qfsplitD + val fsplitE = qfsplitE + end; + +structure Quine_Sum = + struct + val sum = Const("op <+>", [iT,iT]--->iT) + val inl = Const("QInl", iT-->iT) + val inr = Const("QInr", iT-->iT) + val elim = Const("qcase", [iT-->iT, iT-->iT, iT]--->iT) + val case_inl = qcase_QInl + val case_inr = qcase_QInr + val inl_iff = QInl_iff + val inr_iff = QInr_iff + val distinct = QInl_QInr_iff + val distinct' = QInr_QInl_iff + end; + +signature CO_INDRULE = + sig + val co_induct : thm + end; + + +functor Co_Inductive_Fun (Ind: INDUCTIVE) + : sig include INTR_ELIM CO_INDRULE end = +struct +structure Intr_elim = + Intr_elim_Fun(structure Ind=Ind and Fp=Gfp and + Pr=Quine_Prod and Su=Quine_Sum); + +open Intr_elim +val co_induct = raw_induct +end; + diff -r 000000000000 -r a5a9c433f639 src/ZF/constructor.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/constructor.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,200 @@ +(* Title: ZF/constructor.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Constructor function module -- for Datatype Definitions + +Defines constructors and a case-style eliminator (no primitive recursion) + +Features: +* least or greatest fixedpoints +* user-specified product and sum constructions +* mutually recursive datatypes +* recursion over arbitrary monotone operators +* flexible: can derive any reasonable set of introduction rules +* automatically constructs a case analysis operator (but no recursion op) +* efficient treatment of large declarations (e.g. 60 constructors) +*) + +(** STILL NEEDS: some treatment of recursion **) + +signature CONSTRUCTOR = + sig + val thy : theory (*parent theory*) + val rec_specs : (string * string * (string list * string)list) list + (*recursion ops, types, domains, constructors*) + val rec_styp : string (*common type of all recursion ops*) + val ext : Syntax.sext option (*syntax extension for new theory*) + val sintrs : string list (*desired introduction rules*) + val monos : thm list (*monotonicity of each M operator*) + val type_intrs : thm list (*type-checking intro rules*) + val type_elims : thm list (*type-checking elim rules*) + end; + +signature CONSTRUCTOR_RESULT = + sig + val con_thy : theory (*theory defining the constructors*) + val con_defs : thm list (*definitions made in con_thy*) + val case_eqns : thm list (*equations for case operator*) + val free_iffs : thm list (*freeness rewrite rules*) + val free_SEs : thm list (*freeness destruct rules*) + val mk_free : string -> thm (*makes freeness theorems*) + val congs : thm list (*congruence rules for simplifier*) + end; + + +functor Constructor_Fun (structure Const: CONSTRUCTOR and + Pr : PR and Su : SU) : CONSTRUCTOR_RESULT = +struct +open Logic Const; + +val dummy = writeln"Defining the constructor functions..."; + +val case_name = "f"; (*name for case variables*) + +(** Extract basic information from arguments **) + +val sign = sign_of thy; +val rdty = Sign.typ_of o Sign.read_ctyp sign; + +val rec_names = map #1 rec_specs; + +val dummy = assert_all Syntax.is_identifier rec_names + (fn a => "Name of recursive set not an identifier: " ^ a); + +(*Expands multiple constant declarations*) +fun pairtypes (cs,st) = map (rpair st) cs; + +(*Constructors with types and arguments*) +fun mk_con_ty_list cons_pairs = + let fun mk_con_ty (a,st) = + let val T = rdty st + val args = mk_frees "xa" (binder_types T) + in (a,T,args) end + in map mk_con_ty (flat (map pairtypes cons_pairs)) end; + +val con_ty_lists = map (mk_con_ty_list o #3) rec_specs; + +(** Define the constructors **) + +(*We identify 0 (the empty set) with the empty tuple*) +fun mk_tuple [] = Const("0",iT) + | mk_tuple args = foldr1 (app Pr.pair) args; + +fun mk_inject n k u = access_bal(ap Su.inl, ap Su.inr, u) n k; + +val npart = length rec_names; (*number of mutually recursive parts*) + +(*Make constructor definition*) +fun mk_con_defs (kpart, con_ty_list) = + let val ncon = length con_ty_list (*number of constructors*) + fun mk_def ((a,T,args), kcon) = (*kcon = index of this constructor*) + mk_defpair sign + (list_comb (Const(a,T), args), + mk_inject npart kpart (mk_inject ncon kcon (mk_tuple args))) + in map mk_def (con_ty_list ~~ (1 upto ncon)) end; + +(** Define the case operator **) + +(*Combine split terms using case; yields the case operator for one part*) +fun call_case case_list = + let fun call_f (free,args) = + ap_split Pr.split_const free (map (#2 o dest_Free) args) + in fold_bal (app Su.elim) (map call_f case_list) end; + +(** Generating function variables for the case definition + Non-identifiers (e.g. infixes) get a name of the form f_op_nnn. **) + +(*Treatment of a single constructor*) +fun add_case ((a,T,args), (opno,cases)) = + if Syntax.is_identifier a + then (opno, + (Free(case_name ^ "_" ^ a, T), args) :: cases) + else (opno+1, + (Free(case_name ^ "_op_" ^ string_of_int opno, T), args) :: cases); + +(*Treatment of a list of constructors, for one part*) +fun add_case_list (con_ty_list, (opno,case_lists)) = + let val (opno',case_list) = foldr add_case (con_ty_list, (opno,[])) + in (opno', case_list :: case_lists) end; + +(*Treatment of all parts*) +val (_, case_lists) = foldr add_case_list (con_ty_lists, (1,[])); + +val big_case_typ = flat (map (map #2) con_ty_lists) ---> (iT-->iT); + +val big_rec_name = space_implode "_" rec_names; + +val big_case_name = big_rec_name ^ "_case"; + +(*The list of all the function variables*) +val big_case_args = flat (map (map #1) case_lists); + +val big_case_tm = + list_comb (Const(big_case_name, big_case_typ), big_case_args); + +val big_case_def = + mk_defpair sign + (big_case_tm, fold_bal (app Su.elim) (map call_case case_lists)); + +(** Build the new theory **) + +val axpairs = + big_case_def :: flat (map mk_con_defs ((1 upto npart) ~~ con_ty_lists)); + +val const_decs = remove_mixfixes ext + (([big_case_name], flatten_typ sign big_case_typ) :: + (big_rec_name ins rec_names, rec_styp) :: + flat (map #3 rec_specs)); + +val con_thy = extend_theory thy (big_rec_name ^ "_Constructors") + ([], [], [], [], const_decs, ext) axpairs; + +(*1st element is the case definition; others are the constructors*) +val con_defs = map (get_axiom con_thy o #1) axpairs; + +(** Prove the case theorem **) + +(*Each equation has the form + rec_case(f_con1,...,f_conn)(coni(args)) = f_coni(args) *) +fun mk_case_equation ((a,T,args), case_free) = + mk_tprop + (eq_const $ (big_case_tm $ (list_comb (Const(a,T), args))) + $ (list_comb (case_free, args))); + +val case_trans = hd con_defs RS def_trans; + +(*proves a single case equation*) +fun case_tacsf con_def _ = + [rewtac con_def, + rtac case_trans 1, + REPEAT (resolve_tac [refl, Pr.split_eq RS trans, + Su.case_inl RS trans, + Su.case_inr RS trans] 1)]; + +fun prove_case_equation (arg,con_def) = + prove_term (sign_of con_thy) [] + (mk_case_equation arg, case_tacsf con_def); + +val free_iffs = + map standard (con_defs RL [def_swap_iff]) @ + [Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff, Pr.pair_iff]; + +val free_SEs = map (gen_make_elim [conjE,FalseE]) (free_iffs RL [iffD1]); + +val free_cs = ZF_cs addSEs free_SEs; + +(*Typical theorems have the form ~con1=con2, con1=con2==>False, + con1(x)=con1(y) ==> x=y, con1(x)=con1(y) <-> x=y, etc. *) +fun mk_free s = + prove_goalw con_thy con_defs s + (fn prems => [cut_facts_tac prems 1, fast_tac free_cs 1]); + +val case_eqns = map prove_case_equation + (flat con_ty_lists ~~ big_case_args ~~ tl con_defs); + +val congs = mk_congs con_thy (flat (map #1 (const_decs @ ext_constants ext))); +end; + + diff -r 000000000000 -r a5a9c433f639 src/ZF/datatype.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/datatype.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,66 @@ +(* Title: ZF/datatype.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +(Co-) Datatype Definitions for Zermelo-Fraenkel Set Theory + +*) + + +(*Datatype definitions use least fixedpoints, standard products and sums*) +functor Datatype_Fun (Const: CONSTRUCTOR) + : sig include CONSTRUCTOR_RESULT INTR_ELIM INDRULE end = +struct +structure Constructor = Constructor_Fun (structure Const=Const and + Pr=Standard_Prod and Su=Standard_Sum); +open Const Constructor; + +structure Inductive = Inductive_Fun + (val thy = con_thy; + val rec_doms = (map #1 rec_specs) ~~ (map #2 rec_specs); + val sintrs = sintrs; + val monos = monos; + val con_defs = con_defs; + val type_intrs = type_intrs; + val type_elims = type_elims); + +open Inductive +end; + + +(*Co-datatype definitions use greatest fixedpoints, Quine products and sums*) +functor Co_Datatype_Fun (Const: CONSTRUCTOR) + : sig include CONSTRUCTOR_RESULT INTR_ELIM CO_INDRULE end = +struct +structure Constructor = Constructor_Fun (structure Const=Const and + Pr=Quine_Prod and Su=Quine_Sum); +open Const Constructor; + +structure Co_Inductive = Co_Inductive_Fun + (val thy = con_thy; + val rec_doms = (map #1 rec_specs) ~~ (map #2 rec_specs); + val sintrs = sintrs; + val monos = monos; + val con_defs = con_defs; + val type_intrs = type_intrs; + val type_elims = type_elims); + +open Co_Inductive +end; + + + +(*For most datatypes involving univ*) +val data_typechecks = + [SigmaI, InlI, InrI, + Pair_in_univ, Inl_in_univ, Inr_in_univ, + zero_in_univ, A_into_univ, nat_into_univ, sum_univ RS subsetD]; + + +(*For most co-datatypes involving quniv*) +val co_data_typechecks = + [QSigmaI, QInlI, QInrI, + QPair_in_quniv, QInl_in_quniv, QInr_in_quniv, + zero_in_quniv, A_into_quniv, nat_into_quniv, qsum_quniv RS subsetD]; + diff -r 000000000000 -r a5a9c433f639 src/ZF/domrange.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/domrange.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,229 @@ +(* Title: ZF/domrange + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Converse, domain, range of a relation or function +*) + +(*** converse ***) + +val converseI = prove_goalw ZF.thy [converse_def] + "!!a b r. :r ==> :converse(r)" + (fn _ => [ (fast_tac pair_cs 1) ]); + +val converseD = prove_goalw ZF.thy [converse_def] + "!!a b r. : converse(r) ==> : r" + (fn _ => [ (fast_tac pair_cs 1) ]); + +val converseE = prove_goalw ZF.thy [converse_def] + "[| yx : converse(r); \ +\ !!x y. [| yx=; :r |] ==> P \ +\ |] ==> P" + (fn [major,minor]=> + [ (rtac (major RS ReplaceE) 1), + (REPEAT (eresolve_tac [exE, conjE, minor] 1)), + (hyp_subst_tac 1), + (assume_tac 1) ]); + +val converse_cs = pair_cs addSIs [converseI] + addSEs [converseD,converseE]; + +val converse_of_converse = prove_goal ZF.thy + "!!A B r. r<=Sigma(A,B) ==> converse(converse(r)) = r" + (fn _ => [ (fast_tac (converse_cs addSIs [equalityI]) 1) ]); + +val converse_type = prove_goal ZF.thy "!!A B r. r<=A*B ==> converse(r)<=B*A" + (fn _ => [ (fast_tac converse_cs 1) ]); + +val converse_of_prod = prove_goal ZF.thy "converse(A*B) = B*A" + (fn _ => [ (fast_tac (converse_cs addSIs [equalityI]) 1) ]); + +val converse_empty = prove_goal ZF.thy "converse(0) = 0" + (fn _ => [ (fast_tac (converse_cs addSIs [equalityI]) 1) ]); + +(*** domain ***) + +val domain_iff = prove_goalw ZF.thy [domain_def] + "a: domain(r) <-> (EX y. : r)" + (fn _=> [ (fast_tac pair_cs 1) ]); + +val domainI = prove_goal ZF.thy "!!a b r. : r ==> a: domain(r)" + (fn _ => [ (etac (exI RS (domain_iff RS iffD2)) 1) ]); + +val domainE = prove_goal ZF.thy + "[| a : domain(r); !!y. : r ==> P |] ==> P" + (fn prems=> + [ (rtac (domain_iff RS iffD1 RS exE) 1), + (REPEAT (ares_tac prems 1)) ]); + +val domain_of_prod = prove_goal ZF.thy "!!A B. b:B ==> domain(A*B) = A" + (fn _ => + [ (REPEAT (eresolve_tac [domainE,SigmaE2] 1 + ORELSE ares_tac [domainI,equalityI,subsetI,SigmaI] 1)) ]); + +val domain_empty = prove_goal ZF.thy "domain(0) = 0" + (fn _ => + [ (REPEAT (eresolve_tac [domainE,emptyE] 1 + ORELSE ares_tac [equalityI,subsetI] 1)) ]); + +val domain_subset = prove_goal ZF.thy "domain(Sigma(A,B)) <= A" + (fn _ => + [ (rtac subsetI 1), + (etac domainE 1), + (etac SigmaD1 1) ]); + + +(*** range ***) + +val rangeI = prove_goalw ZF.thy [range_def] "!!a b r.: r ==> b : range(r)" + (fn _ => [ (etac (converseI RS domainI) 1) ]); + +val rangeE = prove_goalw ZF.thy [range_def] + "[| b : range(r); !!x. : r ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS domainE) 1), + (resolve_tac prems 1), + (etac converseD 1) ]); + +val range_of_prod = prove_goalw ZF.thy [range_def] + "!!a A B. a:A ==> range(A*B) = B" + (fn _ => + [ (rtac (converse_of_prod RS ssubst) 1), + (etac domain_of_prod 1) ]); + +val range_empty = prove_goalw ZF.thy [range_def] "range(0) = 0" + (fn _ => + [ (rtac (converse_empty RS ssubst) 1), + (rtac domain_empty 1) ]); + +val range_subset = prove_goalw ZF.thy [range_def] "range(A*B) <= B" + (fn _ => + [ (rtac (converse_of_prod RS ssubst) 1), + (rtac domain_subset 1) ]); + + +(*** field ***) + +val fieldI1 = prove_goalw ZF.thy [field_def] ": r ==> a : field(r)" + (fn [prem]=> + [ (rtac (prem RS domainI RS UnI1) 1) ]); + +val fieldI2 = prove_goalw ZF.thy [field_def] ": r ==> b : field(r)" + (fn [prem]=> + [ (rtac (prem RS rangeI RS UnI2) 1) ]); + +val fieldCI = prove_goalw ZF.thy [field_def] + "(~ :r ==> : r) ==> a : field(r)" + (fn [prem]=> + [ (rtac (prem RS domainI RS UnCI) 1), + (swap_res_tac [rangeI] 1), + (etac notnotD 1) ]); + +val fieldE = prove_goalw ZF.thy [field_def] + "[| a : field(r); \ +\ !!x. : r ==> P; \ +\ !!x. : r ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS UnE) 1), + (REPEAT (eresolve_tac (prems@[domainE,rangeE]) 1)) ]); + +val field_of_prod = prove_goal ZF.thy "field(A*A) = A" + (fn _ => + [ (fast_tac (pair_cs addIs [fieldCI,equalityI] addSEs [fieldE]) 1) ]); + +val field_subset = prove_goal ZF.thy "field(A*B) <= A Un B" + (fn _ => [ (fast_tac (pair_cs addIs [fieldCI] addSEs [fieldE]) 1) ]); + +val domain_subset_field = prove_goalw ZF.thy [field_def] + "domain(r) <= field(r)" + (fn _ => [ (rtac Un_upper1 1) ]); + +val range_subset_field = prove_goalw ZF.thy [field_def] + "range(r) <= field(r)" + (fn _ => [ (rtac Un_upper2 1) ]); + +val domain_times_range = prove_goal ZF.thy + "!!A B r. r <= Sigma(A,B) ==> r <= domain(r)*range(r)" + (fn _ => [ (fast_tac (pair_cs addIs [domainI,rangeI]) 1) ]); + +val field_times_field = prove_goal ZF.thy + "!!A B r. r <= Sigma(A,B) ==> r <= field(r)*field(r)" + (fn _ => [ (fast_tac (pair_cs addIs [fieldI1,fieldI2]) 1) ]); + + +(*** Image of a set under a function/relation ***) + +val image_iff = prove_goalw ZF.thy [image_def] + "b : r``A <-> (EX x:A. :r)" + (fn _ => [ fast_tac (pair_cs addIs [rangeI]) 1 ]); + +val image_singleton_iff = prove_goal ZF.thy + "b : r``{a} <-> :r" + (fn _ => [ rtac (image_iff RS iff_trans) 1, + fast_tac pair_cs 1 ]); + +val imageI = prove_goalw ZF.thy [image_def] + "!!a b r. [| : r; a:A |] ==> b : r``A" + (fn _ => [ (REPEAT (ares_tac [CollectI,rangeI,bexI] 1)) ]); + +val imageE = prove_goalw ZF.thy [image_def] + "[| b: r``A; !!x.[| : r; x:A |] ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS CollectE) 1), + (REPEAT (etac bexE 1 ORELSE ares_tac prems 1)) ]); + +val image_subset = prove_goal ZF.thy + "!!A B r. [| r <= A*B; C<=A |] ==> r``C <= B" + (fn _ => + [ (rtac subsetI 1), + (REPEAT (eresolve_tac [asm_rl, imageE, subsetD RS SigmaD2] 1)) ]); + + +(*** Inverse image of a set under a function/relation ***) + +val vimage_iff = prove_goalw ZF.thy [vimage_def,image_def,converse_def] + "a : r-``B <-> (EX y:B. :r)" + (fn _ => [ fast_tac (pair_cs addIs [rangeI]) 1 ]); + +val vimage_singleton_iff = prove_goal ZF.thy + "a : r-``{b} <-> :r" + (fn _ => [ rtac (vimage_iff RS iff_trans) 1, + fast_tac pair_cs 1 ]); + +val vimageI = prove_goalw ZF.thy [vimage_def] + "!!A B r. [| : r; b:B |] ==> a : r-``B" + (fn _ => [ (REPEAT (ares_tac [converseI RS imageI] 1)) ]); + +val vimageE = prove_goalw ZF.thy [vimage_def] + "[| a: r-``B; !!x.[| : r; x:B |] ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS imageE) 1), + (REPEAT (etac converseD 1 ORELSE ares_tac prems 1)) ]); + +val vimage_subset = prove_goalw ZF.thy [vimage_def] + "!!A B r. [| r <= A*B; C<=B |] ==> r-``C <= A" + (fn _ => [ (REPEAT (ares_tac [converse_type RS image_subset] 1)) ]); + + +(** Theorem-proving for ZF set theory **) + +val ZF_cs = pair_cs + addSIs [converseI] + addIs [imageI, vimageI, domainI, rangeI, fieldCI] + addSEs [imageE, vimageE, domainE, rangeE, fieldE, converseD, converseE]; + +val eq_cs = ZF_cs addSIs [equalityI]; + +(** The Union of a set of relations is a relation -- Lemma for fun_Union **) +goal ZF.thy "!!S. (ALL x:S. EX A B. x <= A*B) ==> \ +\ Union(S) <= domain(Union(S)) * range(Union(S))"; +by (fast_tac ZF_cs 1); +val rel_Union = result(); + +(** The Union of 2 relations is a relation (Lemma for fun_Un) **) +val rel_Un = prove_goal ZF.thy + "!!r s. [| r <= A*B; s <= C*D |] ==> (r Un s) <= (A Un C) * (B Un D)" + (fn _ => [ (fast_tac ZF_cs 1) ]); + + diff -r 000000000000 -r a5a9c433f639 src/ZF/epsilon.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/epsilon.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,325 @@ +(* Title: ZF/epsilon.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For epsilon.thy. Epsilon induction and recursion +*) + +open Epsilon; + +(*** Basic closure properties ***) + +goalw Epsilon.thy [eclose_def] "A <= eclose(A)"; +by (rtac (nat_rec_0 RS equalityD2 RS subset_trans) 1); +br (nat_0I RS UN_upper) 1; +val arg_subset_eclose = result(); + +val arg_into_eclose = arg_subset_eclose RS subsetD; + +goalw Epsilon.thy [eclose_def,Transset_def] "Transset(eclose(A))"; +by (rtac (subsetI RS ballI) 1); +by (etac UN_E 1); +by (rtac (nat_succI RS UN_I) 1); +by (assume_tac 1); +by (etac (nat_rec_succ RS ssubst) 1); +by (etac UnionI 1); +by (assume_tac 1); +val Transset_eclose = result(); + +(* x : eclose(A) ==> x <= eclose(A) *) +val eclose_subset = + standard (rewrite_rule [Transset_def] Transset_eclose RS bspec); + +(* [| A : eclose(B); c : A |] ==> c : eclose(B) *) +val ecloseD = standard (eclose_subset RS subsetD); + +val arg_in_eclose_sing = arg_subset_eclose RS singleton_subsetD; +val arg_into_eclose_sing = arg_in_eclose_sing RS ecloseD; + +(* This is epsilon-induction for eclose(A); see also eclose_induct_down... + [| a: eclose(A); !!x. [| x: eclose(A); ALL y:x. P(y) |] ==> P(x) + |] ==> P(a) +*) +val eclose_induct = standard (Transset_eclose RSN (2, Transset_induct)); + +(*Epsilon induction*) +val prems = goal Epsilon.thy + "[| !!x. ALL y:x. P(y) ==> P(x) |] ==> P(a)"; +by (rtac (arg_in_eclose_sing RS eclose_induct) 1); +by (eresolve_tac prems 1); +val eps_induct = result(); + +(*Perform epsilon-induction on i. *) +fun eps_ind_tac a = + EVERY' [res_inst_tac [("a",a)] eps_induct, + rename_last_tac a ["1"]]; + + +(*** Leastness of eclose ***) + +(** eclose(A) is the least transitive set including A as a subset. **) + +goalw Epsilon.thy [Transset_def] + "!!X A n. [| Transset(X); A<=X; n: nat |] ==> \ +\ nat_rec(n, A, %m r. Union(r)) <= X"; +by (etac nat_induct 1); +by (ASM_SIMP_TAC (ZF_ss addrews [nat_rec_0]) 1); +by (ASM_SIMP_TAC (ZF_ss addrews [nat_rec_succ]) 1); +by (fast_tac ZF_cs 1); +val eclose_least_lemma = result(); + +goalw Epsilon.thy [eclose_def] + "!!X A. [| Transset(X); A<=X |] ==> eclose(A) <= X"; +br (eclose_least_lemma RS UN_least) 1; +by (REPEAT (assume_tac 1)); +val eclose_least = result(); + +(*COMPLETELY DIFFERENT induction principle from eclose_induct!!*) +val [major,base,step] = goal Epsilon.thy + "[| a: eclose(b); \ +\ !!y. [| y: b |] ==> P(y); \ +\ !!y z. [| y: eclose(b); P(y); z: y |] ==> P(z) \ +\ |] ==> P(a)"; +by (rtac (major RSN (3, eclose_least RS subsetD RS CollectD2)) 1); +by (rtac (CollectI RS subsetI) 2); +by (etac (arg_subset_eclose RS subsetD) 2); +by (etac base 2); +by (rewtac Transset_def); +by (fast_tac (ZF_cs addEs [step,ecloseD]) 1); +val eclose_induct_down = result(); + +goal Epsilon.thy "!!X. Transset(X) ==> eclose(X) = X"; +be ([eclose_least, arg_subset_eclose] MRS equalityI) 1; +br subset_refl 1; +val Transset_eclose_eq_arg = result(); + + +(*** Epsilon recursion ***) + +(*Unused...*) +goal Epsilon.thy "!!A B C. [| A: eclose(B); B: eclose(C) |] ==> A: eclose(C)"; +by (rtac ([Transset_eclose, eclose_subset] MRS eclose_least RS subsetD) 1); +by (REPEAT (assume_tac 1)); +val mem_eclose_trans = result(); + +(*Variant of the previous lemma in a useable form for the sequel*) +goal Epsilon.thy + "!!A B C. [| A: eclose({B}); B: eclose({C}) |] ==> A: eclose({C})"; +by (rtac ([Transset_eclose, singleton_subsetI] MRS eclose_least RS subsetD) 1); +by (REPEAT (assume_tac 1)); +val mem_eclose_sing_trans = result(); + +goalw Epsilon.thy [Transset_def] + "!!i j. [| Transset(i); j:i |] ==> Memrel(i)-``{j} = j"; +by (fast_tac (eq_cs addSIs [MemrelI] addSEs [MemrelE]) 1); +val under_Memrel = result(); + +(* j : eclose(A) ==> Memrel(eclose(A)) -`` j = j *) +val under_Memrel_eclose = Transset_eclose RS under_Memrel; + +val wfrec_ssubst = standard (wf_Memrel RS wfrec RS ssubst); + +val [kmemj,jmemi] = goal Epsilon.thy + "[| k:eclose({j}); j:eclose({i}) |] ==> \ +\ wfrec(Memrel(eclose({i})), k, H) = wfrec(Memrel(eclose({j})), k, H)"; +by (rtac (kmemj RS eclose_induct) 1); +by (rtac wfrec_ssubst 1); +by (rtac wfrec_ssubst 1); +by (ASM_SIMP_TAC (wf_ss addrews [under_Memrel_eclose, + jmemi RSN (2,mem_eclose_sing_trans)]) 1); +val wfrec_eclose_eq = result(); + +val [prem] = goal Epsilon.thy + "k: i ==> wfrec(Memrel(eclose({i})),k,H) = wfrec(Memrel(eclose({k})),k,H)"; +by (rtac (arg_in_eclose_sing RS wfrec_eclose_eq) 1); +by (rtac (prem RS arg_into_eclose_sing) 1); +val wfrec_eclose_eq2 = result(); + +goalw Epsilon.thy [transrec_def] + "transrec(a,H) = H(a, lam x:a. transrec(x,H))"; +by (rtac wfrec_ssubst 1); +by (SIMP_TAC (wf_ss addrews [wfrec_eclose_eq2, + arg_in_eclose_sing, under_Memrel_eclose]) 1); +val transrec = result(); + +(*Avoids explosions in proofs; resolve it with a meta-level definition.*) +val rew::prems = goal Epsilon.thy + "[| !!x. f(x)==transrec(x,H) |] ==> f(a) = H(a, lam x:a. f(x))"; +by (rewtac rew); +by (REPEAT (resolve_tac (prems@[transrec]) 1)); +val def_transrec = result(); + +val prems = goal Epsilon.thy + "[| !!x u. [| x:eclose({a}); u: Pi(x,B) |] ==> H(x,u) : B(x) \ +\ |] ==> transrec(a,H) : B(a)"; +by (res_inst_tac [("i", "a")] (arg_in_eclose_sing RS eclose_induct) 1); +by (rtac (transrec RS ssubst) 1); +by (REPEAT (ares_tac (prems @ [lam_type]) 1 ORELSE etac bspec 1)); +val transrec_type = result(); + +goal Epsilon.thy "!!i. Ord(i) ==> eclose({i}) <= succ(i)"; +by (etac (Ord_is_Transset RS Transset_succ RS eclose_least) 1); +by (rtac (succI1 RS singleton_subsetI) 1); +val eclose_sing_Ord = result(); + +val prems = goal Epsilon.thy + "[| j: i; Ord(i); \ +\ !!x u. [| x: i; u: Pi(x,B) |] ==> H(x,u) : B(x) \ +\ |] ==> transrec(j,H) : B(j)"; +by (rtac transrec_type 1); +by (resolve_tac prems 1); +by (rtac (Ord_in_Ord RS eclose_sing_Ord RS subsetD RS succE) 1); +by (DEPTH_SOLVE (ares_tac prems 1 ORELSE eresolve_tac [ssubst,Ord_trans] 1)); +val Ord_transrec_type = result(); + +(*Congruence*) +val prems = goalw Epsilon.thy [transrec_def,Memrel_def] + "[| a=a'; !!x u. H(x,u)=H'(x,u) |] ==> transrec(a,H)=transrec(a',H')"; +val transrec_ss = + ZF_ss addcongs ([wfrec_cong] @ mk_congs Epsilon.thy ["eclose"]) + addrews (prems RL [sym]); +by (SIMP_TAC transrec_ss 1); +val transrec_cong = result(); + +(*** Rank ***) + +val ord_ss = ZF_ss addcongs (mk_congs Ord.thy ["Ord"]); + +(*NOT SUITABLE FOR REWRITING -- RECURSIVE!*) +goal Epsilon.thy "rank(a) = (UN y:a. succ(rank(y)))"; +by (rtac (rank_def RS def_transrec RS ssubst) 1); +by (SIMP_TAC ZF_ss 1); +val rank = result(); + +goal Epsilon.thy "Ord(rank(a))"; +by (eps_ind_tac "a" 1); +by (rtac (rank RS ssubst) 1); +by (rtac (Ord_succ RS Ord_UN) 1); +by (etac bspec 1); +by (assume_tac 1); +val Ord_rank = result(); + +val [major] = goal Epsilon.thy "Ord(i) ==> rank(i) = i"; +by (rtac (major RS trans_induct) 1); +by (rtac (rank RS ssubst) 1); +by (ASM_SIMP_TAC (ord_ss addrews [Ord_equality]) 1); +val rank_of_Ord = result(); + +val [prem] = goal Epsilon.thy "a:b ==> rank(a) : rank(b)"; +by (res_inst_tac [("a1","b")] (rank RS ssubst) 1); +by (rtac (prem RS UN_I) 1); +by (rtac succI1 1); +val rank_lt = result(); + +val [major] = goal Epsilon.thy "a: eclose(b) ==> rank(a) : rank(b)"; +by (rtac (major RS eclose_induct_down) 1); +by (etac rank_lt 1); +by (etac (rank_lt RS Ord_trans) 1); +by (assume_tac 1); +by (rtac Ord_rank 1); +val eclose_rank_lt = result(); + +goal Epsilon.thy "!!a b. a<=b ==> rank(a) <= rank(b)"; +by (rtac (rank RS ssubst) 1); +by (rtac (rank RS ssubst) 1); +by (etac UN_mono 1); +by (rtac subset_refl 1); +val rank_mono = result(); + +goal Epsilon.thy "rank(Pow(a)) = succ(rank(a))"; +by (rtac (rank RS trans) 1); +by (rtac equalityI 1); +by (fast_tac ZF_cs 2); +by (rtac UN_least 1); +by (etac (PowD RS rank_mono RS Ord_succ_mono) 1); +by (rtac Ord_rank 1); +by (rtac Ord_rank 1); +val rank_Pow = result(); + +goal Epsilon.thy "rank(0) = 0"; +by (rtac (rank RS trans) 1); +by (fast_tac (ZF_cs addSIs [equalityI]) 1); +val rank_0 = result(); + +goal Epsilon.thy "rank(succ(x)) = succ(rank(x))"; +by (rtac (rank RS trans) 1); +br ([UN_least, succI1 RS UN_upper] MRS equalityI) 1; +be succE 1; +by (fast_tac ZF_cs 1); +by (REPEAT (ares_tac [Ord_succ_mono,Ord_rank,OrdmemD,rank_lt] 1)); +val rank_succ = result(); + +goal Epsilon.thy "rank(Union(A)) = (UN x:A. rank(x))"; +by (rtac equalityI 1); +by (rtac (rank_mono RS UN_least) 2); +by (etac Union_upper 2); +by (rtac (rank RS ssubst) 1); +by (rtac UN_least 1); +by (etac UnionE 1); +by (rtac subset_trans 1); +by (etac (RepFunI RS Union_upper) 2); +by (etac (rank_lt RS Ord_succ_subsetI) 1); +by (rtac Ord_rank 1); +val rank_Union = result(); + +goal Epsilon.thy "rank(eclose(a)) = rank(a)"; +by (rtac equalityI 1); +by (rtac (arg_subset_eclose RS rank_mono) 2); +by (res_inst_tac [("a1","eclose(a)")] (rank RS ssubst) 1); +by (rtac UN_least 1); +by (etac ([eclose_rank_lt, Ord_rank] MRS Ord_succ_subsetI) 1); +val rank_eclose = result(); + +(* [| i: j; j: rank(a) |] ==> i: rank(a) *) +val rank_trans = Ord_rank RSN (3, Ord_trans); + +goalw Epsilon.thy [Pair_def] "rank(a) : rank()"; +by (rtac (consI1 RS rank_lt RS Ord_trans) 1); +by (rtac (consI1 RS consI2 RS rank_lt) 1); +by (rtac Ord_rank 1); +val rank_pair1 = result(); + +goalw Epsilon.thy [Pair_def] "rank(b) : rank()"; +by (rtac (consI1 RS consI2 RS rank_lt RS Ord_trans) 1); +by (rtac (consI1 RS consI2 RS rank_lt) 1); +by (rtac Ord_rank 1); +val rank_pair2 = result(); + +goalw (merge_theories(Epsilon.thy,Sum.thy)) [Inl_def] "rank(a) : rank(Inl(a))"; +by (rtac rank_pair2 1); +val rank_Inl = result(); + +goalw (merge_theories(Epsilon.thy,Sum.thy)) [Inr_def] "rank(a) : rank(Inr(a))"; +by (rtac rank_pair2 1); +val rank_Inr = result(); + +val [major] = goal Epsilon.thy "i: rank(a) ==> (EX x:a. i<=rank(x))"; +by (resolve_tac ([major] RL [rank RS subst] RL [UN_E]) 1); +by (rtac bexI 1); +by (etac member_succD 1); +by (rtac Ord_rank 1); +by (assume_tac 1); +val rank_implies_mem = result(); + + +(*** Corollaries of leastness ***) + +goal Epsilon.thy "!!A B. A:B ==> eclose(A)<=eclose(B)"; +by (rtac (Transset_eclose RS eclose_least) 1); +by (etac (arg_into_eclose RS eclose_subset) 1); +val mem_eclose_subset = result(); + +goal Epsilon.thy "!!A B. A<=B ==> eclose(A) <= eclose(B)"; +by (rtac (Transset_eclose RS eclose_least) 1); +by (etac subset_trans 1); +by (rtac arg_subset_eclose 1); +val eclose_mono = result(); + +(** Idempotence of eclose **) + +goal Epsilon.thy "eclose(eclose(A)) = eclose(A)"; +by (rtac equalityI 1); +by (rtac ([Transset_eclose, subset_refl] MRS eclose_least) 1); +by (rtac arg_subset_eclose 1); +val eclose_idem = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/epsilon.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/epsilon.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,18 @@ +(* Title: ZF/epsilon.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Epsilon induction and recursion +*) + +Epsilon = Nat + +consts + eclose,rank :: "i=>i" + transrec :: "[i, [i,i]=>i] =>i" + +rules + eclose_def "eclose(A) == UN n:nat. nat_rec(n, A, %m r. Union(r))" + transrec_def "transrec(a,H) == wfrec(Memrel(eclose({a})), a, H)" + rank_def "rank(a) == transrec(a, %x f. UN y:x. succ(f`y))" +end diff -r 000000000000 -r a5a9c433f639 src/ZF/equalities.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/equalities.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,312 @@ +(* Title: ZF/equalities + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Set Theory examples: Union, Intersection, Inclusion, etc. + (Thanks also to Philippe de Groote.) +*) + +(** Finite Sets **) + +goal ZF.thy "cons(a, cons(b, C)) = cons(b, cons(a, C))"; +by (fast_tac eq_cs 1); +val cons_commute = result(); + +goal ZF.thy "!!B. a: B ==> cons(a,B) = B"; +by (fast_tac eq_cs 1); +val cons_absorb = result(); + +goal ZF.thy "!!B. a: B ==> cons(a, B-{a}) = B"; +by (fast_tac eq_cs 1); +val cons_Diff = result(); + +goal ZF.thy "!!C. [| a: C; ALL y:C. y=b |] ==> C = {b}"; +by (fast_tac eq_cs 1); +val equal_singleton_lemma = result(); +val equal_singleton = ballI RSN (2,equal_singleton_lemma); + + +(** Binary Intersection **) + +goal ZF.thy "0 Int A = 0"; +by (fast_tac eq_cs 1); +val Int_0 = result(); + +(*NOT an equality, but it seems to belong here...*) +goal ZF.thy "cons(a,B) Int C <= cons(a, B Int C)"; +by (fast_tac eq_cs 1); +val Int_cons = result(); + +goal ZF.thy "A Int A = A"; +by (fast_tac eq_cs 1); +val Int_absorb = result(); + +goal ZF.thy "A Int B = B Int A"; +by (fast_tac eq_cs 1); +val Int_commute = result(); + +goal ZF.thy "(A Int B) Int C = A Int (B Int C)"; +by (fast_tac eq_cs 1); +val Int_assoc = result(); + +goal ZF.thy "(A Un B) Int C = (A Int C) Un (B Int C)"; +by (fast_tac eq_cs 1); +val Int_Un_distrib = result(); + +goal ZF.thy "A<=B <-> A Int B = A"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val subset_Int_iff = result(); + +(** Binary Union **) + +goal ZF.thy "0 Un A = A"; +by (fast_tac eq_cs 1); +val Un_0 = result(); + +goal ZF.thy "cons(a,B) Un C = cons(a, B Un C)"; +by (fast_tac eq_cs 1); +val Un_cons = result(); + +goal ZF.thy "A Un A = A"; +by (fast_tac eq_cs 1); +val Un_absorb = result(); + +goal ZF.thy "A Un B = B Un A"; +by (fast_tac eq_cs 1); +val Un_commute = result(); + +goal ZF.thy "(A Un B) Un C = A Un (B Un C)"; +by (fast_tac eq_cs 1); +val Un_assoc = result(); + +goal ZF.thy "(A Int B) Un C = (A Un C) Int (B Un C)"; +by (fast_tac eq_cs 1); +val Un_Int_distrib = result(); + +goal ZF.thy "A<=B <-> A Un B = B"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val subset_Un_iff = result(); + +(** Simple properties of Diff -- set difference **) + +goal ZF.thy "A-A = 0"; +by (fast_tac eq_cs 1); +val Diff_cancel = result(); + +goal ZF.thy "0-A = 0"; +by (fast_tac eq_cs 1); +val empty_Diff = result(); + +goal ZF.thy "A-0 = A"; +by (fast_tac eq_cs 1); +val Diff_0 = result(); + +(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*) +goal ZF.thy "A - cons(a,B) = A - B - {a}"; +by (fast_tac eq_cs 1); +val Diff_cons = result(); + +(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*) +goal ZF.thy "A - cons(a,B) = A - {a} - B"; +by (fast_tac eq_cs 1); +val Diff_cons2 = result(); + +goal ZF.thy "A Int (B-A) = 0"; +by (fast_tac eq_cs 1); +val Diff_disjoint = result(); + +goal ZF.thy "!!A B. A<=B ==> A Un (B-A) = B"; +by (fast_tac eq_cs 1); +val Diff_partition = result(); + +goal ZF.thy "!!A B. [| A<=B; B<= C |] ==> (B - (C-A)) = A"; +by (fast_tac eq_cs 1); +val double_complement = result(); + +goal ZF.thy + "(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)"; +by (fast_tac eq_cs 1); +val Un_Int_crazy = result(); + +goal ZF.thy "A - (B Un C) = (A-B) Int (A-C)"; +by (fast_tac eq_cs 1); +val Diff_Un = result(); + +goal ZF.thy "A - (B Int C) = (A-B) Un (A-C)"; +by (fast_tac eq_cs 1); +val Diff_Int = result(); + +(*Halmos, Naive Set Theory, page 16.*) +goal ZF.thy "(A Int B) Un C = A Int (B Un C) <-> C<=A"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val Un_Int_assoc_iff = result(); + + +(** Big Union and Intersection **) + +goal ZF.thy "Union(0) = 0"; +by (fast_tac eq_cs 1); +val Union_0 = result(); + +goal ZF.thy "Union(cons(a,B)) = a Un Union(B)"; +by (fast_tac eq_cs 1); +val Union_cons = result(); + +goal ZF.thy "Union(A Un B) = Union(A) Un Union(B)"; +by (fast_tac eq_cs 1); +val Union_Un_distrib = result(); + +goal ZF.thy "Union(C) Int A = 0 <-> (ALL B:C. B Int A = 0)"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val Union_disjoint = result(); + +(* A good challenge: Inter is ill-behaved on the empty set *) +goal ZF.thy "!!A B. [| a:A; b:B |] ==> Inter(A Un B) = Inter(A) Int Inter(B)"; +by (fast_tac eq_cs 1); +val Inter_Un_distrib = result(); + +goal ZF.thy "Union({b}) = b"; +by (fast_tac eq_cs 1); +val Union_singleton = result(); + +goal ZF.thy "Inter({b}) = b"; +by (fast_tac eq_cs 1); +val Inter_singleton = result(); + +(** Unions and Intersections of Families **) + +goal ZF.thy "Union(A) = (UN x:A. x)"; +by (fast_tac eq_cs 1); +val Union_eq_UN = result(); + +goalw ZF.thy [Inter_def] "Inter(A) = (INT x:A. x)"; +by (fast_tac eq_cs 1); +val Inter_eq_INT = result(); + +(*Halmos, Naive Set Theory, page 35.*) +goal ZF.thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))"; +by (fast_tac eq_cs 1); +val Int_UN_distrib = result(); + +goal ZF.thy "!!A B. i:I ==> B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))"; +by (fast_tac eq_cs 1); +val Un_INT_distrib = result(); + +goal ZF.thy + "(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))"; +by (fast_tac eq_cs 1); +val Int_UN_distrib2 = result(); + +goal ZF.thy + "!!I J. [| i:I; j:J |] ==> \ +\ (INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))"; +by (fast_tac eq_cs 1); +val Un_INT_distrib2 = result(); + +goal ZF.thy "!!A. [| a: A; ALL y:A. b(y)=b(a) |] ==> (UN y:A. b(y)) = b(a)"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val UN_singleton_lemma = result(); +val UN_singleton = ballI RSN (2,UN_singleton_lemma); + +goal ZF.thy "!!A. [| a: A; ALL y:A. b(y)=b(a) |] ==> (INT y:A. b(y)) = b(a)"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val INT_singleton_lemma = result(); +val INT_singleton = ballI RSN (2,INT_singleton_lemma); + +(** Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: + Union of a family of unions **) + +goal ZF.thy "(UN i:I. A(x) Un B(x)) = (UN i:I. A(x)) Un (UN i:I. B(x))"; +by (fast_tac eq_cs 1); +val UN_Un_distrib = result(); + +goal ZF.thy + "!!A B. i:I ==> \ +\ (INT i:I. A(x) Int B(x)) = (INT i:I. A(x)) Int (INT i:I. B(x))"; +by (fast_tac eq_cs 1); +val INT_Int_distrib = result(); + +(** Devlin, page 12, exercise 5: Complements **) + +goal ZF.thy "!!A B. i:I ==> B - (UN i:I. A(i)) = (INT i:I. B - A(i))"; +by (fast_tac eq_cs 1); +val Diff_UN = result(); + +goal ZF.thy "!!A B. i:I ==> B - (INT i:I. A(i)) = (UN i:I. B - A(i))"; +by (fast_tac eq_cs 1); +val Diff_INT = result(); + +(** Unions and Intersections with General Sum **) + +goal ZF.thy "(SUM x:A Un B. C(x)) = (SUM x:A. C(x)) Un (SUM x:B. C(x))"; +by (fast_tac eq_cs 1); +val SUM_Un_distrib1 = result(); + +goal ZF.thy + "(SUM i:I. A(x) Un B(x)) = (SUM i:I. A(x)) Un (SUM i:I. B(x))"; +by (fast_tac eq_cs 1); +val SUM_Un_distrib2 = result(); + +goal ZF.thy "(SUM x:A Int B. C(x)) = (SUM x:A. C(x)) Int (SUM x:B. C(x))"; +by (fast_tac eq_cs 1); +val SUM_Int_distrib1 = result(); + +goal ZF.thy + "(SUM i:I. A(x) Int B(x)) = (SUM i:I. A(x)) Int (SUM i:I. B(x))"; +by (fast_tac eq_cs 1); +val SUM_Int_distrib2 = result(); + +(** Domain, Range and Field **) + +goal ZF.thy "domain(A Un B) = domain(A) Un domain(B)"; +by (fast_tac eq_cs 1); +val domain_Un_eq = result(); + +goal ZF.thy "domain(A Int B) <= domain(A) Int domain(B)"; +by (fast_tac eq_cs 1); +val domain_Int_subset = result(); + +goal ZF.thy "domain(A) - domain(B) <= domain(A - B)"; +by (fast_tac eq_cs 1); +val domain_diff_subset = result(); + +goal ZF.thy "range(A Un B) = range(A) Un range(B)"; +by (fast_tac eq_cs 1); +val range_Un_eq = result(); + +goal ZF.thy "range(A Int B) <= range(A) Int range(B)"; +by (fast_tac eq_cs 1); +val range_Int_subset = result(); + +goal ZF.thy "range(A) - range(B) <= range(A - B)"; +by (fast_tac eq_cs 1); +val range_diff_subset = result(); + +goal ZF.thy "field(A Un B) = field(A) Un field(B)"; +by (fast_tac eq_cs 1); +val field_Un_eq = result(); + +goal ZF.thy "field(A Int B) <= field(A) Int field(B)"; +by (fast_tac eq_cs 1); +val field_Int_subset = result(); + +goal ZF.thy "field(A) - field(B) <= field(A - B)"; +by (fast_tac eq_cs 1); +val field_diff_subset = result(); + + +(** Converse **) + +goal ZF.thy "converse(A Un B) = converse(A) Un converse(B)"; +by (fast_tac eq_cs 1); +val converse_Un = result(); + +goal ZF.thy "converse(A Int B) = converse(A) Int converse(B)"; +by (fast_tac eq_cs 1); +val converse_Int = result(); + +goal ZF.thy "converse(A) - converse(B) = converse(A - B)"; +by (fast_tac eq_cs 1); +val converse_diff = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Acc.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Acc.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,65 @@ +(* Title: ZF/ex/acc + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Inductive definition of acc(r) + +See Ch. Paulin-Mohring, Inductive Definitions in the System Coq. +Research Report 92-49, LIP, ENS Lyon. Dec 1992. +*) + +structure Acc = Inductive_Fun + (val thy = WF.thy addconsts [(["acc"],"i=>i")]; + val rec_doms = [("acc", "field(r)")]; + val sintrs = + ["[| r-``{b} : Pow(acc(r)); b : field(r) |] ==> b : acc(r)"]; + val monos = [Pow_mono]; + val con_defs = []; + val type_intrs = []; + val type_elims = []); + +goal Acc.thy "!!a b r. [| b: acc(r); : r |] ==> a: acc(r)"; +by (etac Acc.elim 1); +by (fast_tac ZF_cs 1); +val acc_downward = result(); + +val [major] = goal Acc.thy "field(r) <= acc(r) ==> wf(r)"; +by (rtac (major RS wfI2) 1); +by (rtac subsetI 1); +by (etac Acc.induct 1); +by (etac (bspec RS mp) 1); +by (resolve_tac Acc.intrs 1); +by (assume_tac 2); +by (ALLGOALS (fast_tac ZF_cs)); +val acc_wfI = result(); + +goal ZF.thy "!!r A. field(r Int A*A) <= field(r) Int A"; +by (fast_tac ZF_cs 1); +val field_Int_prodself = result(); + +goal Acc.thy "wf(r Int (acc(r)*acc(r)))"; +by (rtac (field_Int_prodself RS wfI2) 1); +by (rtac subsetI 1); +by (etac IntE 1); +by (etac Acc.induct 1); +by (etac (bspec RS mp) 1); +by (rtac IntI 1); +by (assume_tac 1); +by (resolve_tac Acc.intrs 1); +by (assume_tac 2); +by (ALLGOALS (fast_tac ZF_cs)); +val wf_acc_Int = result(); + +val [major] = goal Acc.thy "wf(r) ==> field(r) <= acc(r)"; +by (rtac subsetI 1); +by (etac (major RS wf_induct2) 1); +by (rtac subset_refl 1); +by (resolve_tac Acc.intrs 1); +by (assume_tac 2); +by (fast_tac ZF_cs 1); +val acc_wfD = result(); + +goal Acc.thy "wf(r) <-> field(r) <= acc(r)"; +by (EVERY1 [rtac iffI, etac acc_wfD, etac acc_wfI]); +val wf_acc_iff = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/BT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/BT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,49 @@ +(* Title: ZF/ex/bt.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Datatype definition of binary trees +*) + +structure BT = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("bt", "univ(A)", + [(["Lf"],"i"), (["Br"],"[i,i,i]=>i")])]; + val rec_styp = "i=>i"; + val ext = None + val sintrs = + ["Lf : bt(A)", + "[| a: A; t1: bt(A); t2: bt(A) |] ==> Br(a,t1,t2) : bt(A)"]; + val monos = []; + val type_intrs = data_typechecks + val type_elims = []); + +val [LfI, BrI] = BT.intrs; + +(*Perform induction on l, then prove the major premise using prems. *) +fun bt_ind_tac a prems i = + EVERY [res_inst_tac [("x",a)] BT.induct i, + rename_last_tac a ["1","2"] (i+2), + ares_tac prems i]; + + +(** Lemmas to justify using "bt" in other recursive type definitions **) + +goalw BT.thy BT.defs "!!A B. A<=B ==> bt(A) <= bt(B)"; +by (rtac lfp_mono 1); +by (REPEAT (rtac BT.bnd_mono 1)); +by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); +val bt_mono = result(); + +goalw BT.thy (BT.defs@BT.con_defs) "bt(univ(A)) <= univ(A)"; +by (rtac lfp_lowerbound 1); +by (rtac (A_subset_univ RS univ_mono) 2); +by (fast_tac (ZF_cs addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ, + Pair_in_univ]) 1); +val bt_univ = result(); + +val bt_subset_univ = standard (bt_mono RS (bt_univ RSN (2,subset_trans))); + + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/BT_Fn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/BT_Fn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,138 @@ +(* Title: ZF/bt.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For bt.thy. Binary trees +*) + +open BT_Fn; + + + +(** bt_rec -- by Vset recursion **) + +(*Used to verify bt_rec*) +val bt_rec_ss = ZF_ss + addcongs (mk_typed_congs BT_Fn.thy [("h", "[i,i,i,i,i]=>i")]) + addrews BT.case_eqns; + +goalw BT.thy BT.con_defs "rank(l) : rank(Br(a,l,r))"; +by (SIMP_TAC rank_ss 1); +val rank_Br1 = result(); + +goalw BT.thy BT.con_defs "rank(r) : rank(Br(a,l,r))"; +by (SIMP_TAC rank_ss 1); +val rank_Br2 = result(); + +goal BT_Fn.thy "bt_rec(Lf,c,h) = c"; +by (rtac (bt_rec_def RS def_Vrec RS trans) 1); +by (SIMP_TAC bt_rec_ss 1); +val bt_rec_Lf = result(); + +goal BT_Fn.thy + "bt_rec(Br(a,l,r), c, h) = h(a, l, r, bt_rec(l,c,h), bt_rec(r,c,h))"; +by (rtac (bt_rec_def RS def_Vrec RS trans) 1); +by (SIMP_TAC (bt_rec_ss addrews [Vset_rankI, rank_Br1, rank_Br2]) 1); +val bt_rec_Br = result(); + +(*Type checking -- proved by induction, as usual*) +val prems = goal BT_Fn.thy + "[| t: bt(A); \ +\ c: C(Lf); \ +\ !!x y z r s. [| x:A; y:bt(A); z:bt(A); r:C(y); s:C(z) |] ==> \ +\ h(x,y,z,r,s): C(Br(x,y,z)) \ +\ |] ==> bt_rec(t,c,h) : C(t)"; +by (bt_ind_tac "t" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (ZF_ss addrews + (prems@[bt_rec_Lf,bt_rec_Br])))); +val bt_rec_type = result(); + +(** Versions for use with definitions **) + +val [rew] = goal BT_Fn.thy "[| !!t. j(t)==bt_rec(t, c, h) |] ==> j(Lf) = c"; +by (rewtac rew); +by (rtac bt_rec_Lf 1); +val def_bt_rec_Lf = result(); + +val [rew] = goal BT_Fn.thy + "[| !!t. j(t)==bt_rec(t, c, h) |] ==> j(Br(a,l,r)) = h(a,l,r,j(l),j(r))"; +by (rewtac rew); +by (rtac bt_rec_Br 1); +val def_bt_rec_Br = result(); + +fun bt_recs def = map standard ([def] RL [def_bt_rec_Lf, def_bt_rec_Br]); + +(** n_nodes **) + +val [n_nodes_Lf,n_nodes_Br] = bt_recs n_nodes_def; + +val prems = goalw BT_Fn.thy [n_nodes_def] + "xs: bt(A) ==> n_nodes(xs) : nat"; +by (REPEAT (ares_tac (prems @ [bt_rec_type, nat_0I, nat_succI, add_type]) 1)); +val n_nodes_type = result(); + + +(** n_leaves **) + +val [n_leaves_Lf,n_leaves_Br] = bt_recs n_leaves_def; + +val prems = goalw BT_Fn.thy [n_leaves_def] + "xs: bt(A) ==> n_leaves(xs) : nat"; +by (REPEAT (ares_tac (prems @ [bt_rec_type, nat_0I, nat_succI, add_type]) 1)); +val n_leaves_type = result(); + +(** bt_reflect **) + +val [bt_reflect_Lf, bt_reflect_Br] = bt_recs bt_reflect_def; + +val prems = goalw BT_Fn.thy [bt_reflect_def] + "xs: bt(A) ==> bt_reflect(xs) : bt(A)"; +by (REPEAT (ares_tac (prems @ [bt_rec_type, LfI, BrI]) 1)); +val bt_reflect_type = result(); + + +(** BT_Fn simplification **) + + +val bt_typechecks = + [LfI, BrI, bt_rec_type, n_nodes_type, n_leaves_type, bt_reflect_type]; + +val bt_congs = + BT.congs @ + mk_congs BT_Fn.thy ["bt_case","bt_rec","n_nodes","n_leaves","bt_reflect"]; + +val bt_ss = arith_ss + addcongs bt_congs + addrews BT.case_eqns + addrews bt_typechecks + addrews [bt_rec_Lf, bt_rec_Br, + n_nodes_Lf, n_nodes_Br, + n_leaves_Lf, n_leaves_Br, + bt_reflect_Lf, bt_reflect_Br]; + + +(*** theorems about n_leaves ***) + +val prems = goal BT_Fn.thy + "t: bt(A) ==> n_leaves(bt_reflect(t)) = n_leaves(t)"; +by (bt_ind_tac "t" prems 1); +by (ALLGOALS (ASM_SIMP_TAC bt_ss)); +by (REPEAT (ares_tac [add_commute, n_leaves_type] 1)); +val n_leaves_reflect = result(); + +val prems = goal BT_Fn.thy + "t: bt(A) ==> n_leaves(t) = succ(n_nodes(t))"; +by (bt_ind_tac "t" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (bt_ss addrews [add_succ_right]))); +val n_leaves_nodes = result(); + +(*** theorems about bt_reflect ***) + +val prems = goal BT_Fn.thy + "t: bt(A) ==> bt_reflect(bt_reflect(t))=t"; +by (bt_ind_tac "t" prems 1); +by (ALLGOALS (ASM_SIMP_TAC bt_ss)); +val bt_reflect_bt_reflect_ident = result(); + + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/BT_Fn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/BT_Fn.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,24 @@ +(* Title: ZF/ex/bt-fn.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Binary trees +*) + +BT_Fn = BT + +consts + bt_rec :: "[i, i, [i,i,i,i,i]=>i] => i" + n_nodes :: "i=>i" + n_leaves :: "i=>i" + bt_reflect :: "i=>i" + +rules + bt_rec_def + "bt_rec(t,c,h) == Vrec(t, %t g.bt_case(c, %x y z. h(x,y,z,g`y,g`z), t))" + + n_nodes_def "n_nodes(t) == bt_rec(t, 0, %x y z r s. succ(r#+s))" + n_leaves_def "n_leaves(t) == bt_rec(t, succ(0), %x y z r s. r#+s)" + bt_reflect_def "bt_reflect(t) == bt_rec(t, Lf, %x y z r s. Br(x,s,r))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Bin.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Bin.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,42 @@ +(* Title: ZF/ex/bin.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Datatype of binary integers +*) + +(*Example of a datatype with an infix constructor*) +structure Bin = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("bin", "univ(0)", + [(["Plus", "Minus"], "i"), + (["op $$"], "[i,i]=>i")])]; + val rec_styp = "i"; + val ext = Some (NewSext { + mixfix = + [Infixl("$$", "[i,i] => i", 60)], + xrules = [], + parse_ast_translation = [], + parse_preproc = None, + parse_postproc = None, + parse_translation = [], + print_translation = [], + print_preproc = None, + print_postproc = None, + print_ast_translation = []}); + val sintrs = + ["Plus : bin", + "Minus : bin", + "[| w: bin; b: bool |] ==> w$$b : bin"]; + val monos = []; + val type_intrs = bool_into_univ::data_typechecks; + val type_elims = []); + +(*Perform induction on l, then prove the major premise using prems. *) +fun bin_ind_tac a prems i = + EVERY [res_inst_tac [("x",a)] Bin.induct i, + rename_last_tac a ["1"] (i+3), + ares_tac prems i]; + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/BinFn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/BinFn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,463 @@ +(* Title: ZF/ex/bin.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For bin.thy. Arithmetic on binary integers. +*) + +open BinFn; + + +(** bin_rec -- by Vset recursion **) + +goal BinFn.thy "bin_rec(Plus,a,b,h) = a"; +by (rtac (bin_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Bin.con_defs); +by (SIMP_TAC rank_ss 1); +val bin_rec_Plus = result(); + +goal BinFn.thy "bin_rec(Minus,a,b,h) = b"; +by (rtac (bin_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Bin.con_defs); +by (SIMP_TAC rank_ss 1); +val bin_rec_Minus = result(); + +goal BinFn.thy "bin_rec(w$$x,a,b,h) = h(w, x, bin_rec(w,a,b,h))"; +by (rtac (bin_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Bin.con_defs); +by (SIMP_TAC (rank_ss addcongs + (mk_typed_congs BinFn.thy [("h", "[i,i,i]=>i")])) 1); +val bin_rec_Bcons = result(); + +(*Type checking*) +val prems = goal BinFn.thy + "[| w: bin; \ +\ a: C(Plus); b: C(Minus); \ +\ !!w x r. [| w: bin; x: bool; r: C(w) |] ==> h(w,x,r): C(w$$x) \ +\ |] ==> bin_rec(w,a,b,h) : C(w)"; +by (bin_ind_tac "w" prems 1); +by (ALLGOALS + (ASM_SIMP_TAC (ZF_ss addrews (prems@[bin_rec_Plus,bin_rec_Minus, + bin_rec_Bcons])))); +val bin_rec_type = result(); + +(** Versions for use with definitions **) + +val [rew] = goal BinFn.thy + "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(Plus) = a"; +by (rewtac rew); +by (rtac bin_rec_Plus 1); +val def_bin_rec_Plus = result(); + +val [rew] = goal BinFn.thy + "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(Minus) = b"; +by (rewtac rew); +by (rtac bin_rec_Minus 1); +val def_bin_rec_Minus = result(); + +val [rew] = goal BinFn.thy + "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(w$$x) = h(w,x,j(w))"; +by (rewtac rew); +by (rtac bin_rec_Bcons 1); +val def_bin_rec_Bcons = result(); + +fun bin_recs def = map standard + ([def] RL [def_bin_rec_Plus, def_bin_rec_Minus, def_bin_rec_Bcons]); + +(** Type checking **) + +val bin_typechecks0 = bin_rec_type :: Bin.intrs; + +goalw BinFn.thy [integ_of_bin_def] + "!!w. w: bin ==> integ_of_bin(w) : integ"; +by (typechk_tac (bin_typechecks0@integ_typechecks@ + nat_typechecks@[bool_into_nat])); +val integ_of_bin_type = result(); + +goalw BinFn.thy [bin_succ_def] + "!!w. w: bin ==> bin_succ(w) : bin"; +by (typechk_tac (bin_typechecks0@bool_typechecks)); +val bin_succ_type = result(); + +goalw BinFn.thy [bin_pred_def] + "!!w. w: bin ==> bin_pred(w) : bin"; +by (typechk_tac (bin_typechecks0@bool_typechecks)); +val bin_pred_type = result(); + +goalw BinFn.thy [bin_minus_def] + "!!w. w: bin ==> bin_minus(w) : bin"; +by (typechk_tac ([bin_pred_type]@bin_typechecks0@bool_typechecks)); +val bin_minus_type = result(); + +goalw BinFn.thy [bin_add_def] + "!!v w. [| v: bin; w: bin |] ==> bin_add(v,w) : bin"; +by (typechk_tac ([bin_succ_type,bin_pred_type]@bin_typechecks0@ + bool_typechecks@ZF_typechecks)); +val bin_add_type = result(); + +goalw BinFn.thy [bin_mult_def] + "!!v w. [| v: bin; w: bin |] ==> bin_mult(v,w) : bin"; +by (typechk_tac ([bin_minus_type,bin_add_type]@bin_typechecks0@ + bool_typechecks)); +val bin_mult_type = result(); + +val bin_typechecks = bin_typechecks0 @ + [integ_of_bin_type, bin_succ_type, bin_pred_type, + bin_minus_type, bin_add_type, bin_mult_type]; + +val bin_congs = mk_congs BinFn.thy + ["bin_rec","op $$","integ_of_bin","bin_succ","bin_pred", + "bin_minus","bin_add","bin_mult"]; + +val bin_ss = integ_ss + addcongs (bin_congs@bool_congs) + addrews([bool_1I, bool_0I, + bin_rec_Plus, bin_rec_Minus, bin_rec_Bcons] @ + bin_recs integ_of_bin_def @ bool_rews @ bin_typechecks); + +val typechecks = bin_typechecks @ integ_typechecks @ nat_typechecks @ + [bool_subset_nat RS subsetD]; + +(**** The carry/borrow functions, bin_succ and bin_pred ****) + +(** Lemmas **) + +goal Integ.thy + "!!z v. [| z $+ v = z' $+ v'; \ +\ z: integ; z': integ; v: integ; v': integ; w: integ |] \ +\ ==> z $+ (v $+ w) = z' $+ (v' $+ w)"; +by (ASM_SIMP_TAC (integ_ss addrews ([zadd_assoc RS sym])) 1); +val zadd_assoc_cong = result(); + +goal Integ.thy + "!!z v w. [| z: integ; v: integ; w: integ |] \ +\ ==> z $+ (v $+ w) = v $+ (z $+ w)"; +by (REPEAT (ares_tac [zadd_commute RS zadd_assoc_cong] 1)); +val zadd_assoc_swap = result(); + +val [zadd_cong] = mk_congs Integ.thy ["op $+"]; + +val zadd_kill = (refl RS zadd_cong); +val zadd_assoc_swap_kill = zadd_kill RSN (4, zadd_assoc_swap RS trans); + +(*Pushes 'constants' of the form $#m to the right -- LOOPS if two!*) +val zadd_assoc_znat = standard (znat_type RS zadd_assoc_swap); + +goal Integ.thy + "!!z w. [| z: integ; w: integ |] \ +\ ==> w $+ (z $+ (w $+ z)) = w $+ (w $+ (z $+ z))"; +by (REPEAT (ares_tac [zadd_kill, zadd_assoc_swap] 1)); +val zadd_swap_pairs = result(); + + +val carry_ss = bin_ss addrews + (bin_recs bin_succ_def @ bin_recs bin_pred_def); + +goal BinFn.thy + "!!w. w: bin ==> integ_of_bin(bin_succ(w)) = $#1 $+ integ_of_bin(w)"; +by (etac Bin.induct 1); +by (SIMP_TAC (carry_ss addrews [zadd_0_right]) 1); +by (SIMP_TAC (carry_ss addrews [zadd_zminus_inverse]) 1); +by (etac boolE 1); +by (ALLGOALS (ASM_SIMP_TAC (carry_ss addrews [zadd_assoc]))); +by (REPEAT (ares_tac (zadd_swap_pairs::typechecks) 1)); +val integ_of_bin_succ = result(); + +goal BinFn.thy + "!!w. w: bin ==> integ_of_bin(bin_pred(w)) = $~ ($#1) $+ integ_of_bin(w)"; +by (etac Bin.induct 1); +by (SIMP_TAC (carry_ss addrews [zadd_0_right]) 1); +by (SIMP_TAC (carry_ss addrews [zadd_zminus_inverse]) 1); +by (etac boolE 1); +by (ALLGOALS + (ASM_SIMP_TAC + (carry_ss addrews [zadd_assoc RS sym, + zadd_zminus_inverse, zadd_zminus_inverse2]))); +by (REPEAT (ares_tac ([zadd_commute, zadd_cong, refl]@typechecks) 1)); +val integ_of_bin_pred = result(); + +(*These two results replace the definitions of bin_succ and bin_pred*) + + +(*** bin_minus: (unary!) negation of binary integers ***) + +val bin_minus_ss = + bin_ss addrews (bin_recs bin_minus_def @ + [integ_of_bin_succ, integ_of_bin_pred]); + +goal BinFn.thy + "!!w. w: bin ==> integ_of_bin(bin_minus(w)) = $~ integ_of_bin(w)"; +by (etac Bin.induct 1); +by (SIMP_TAC (bin_minus_ss addrews [zminus_0]) 1); +by (SIMP_TAC (bin_minus_ss addrews [zadd_0_right]) 1); +by (etac boolE 1); +by (ALLGOALS + (ASM_SIMP_TAC (bin_minus_ss addrews [zminus_zadd_distrib, zadd_assoc]))); +val integ_of_bin_minus = result(); + + +(*** bin_add: binary addition ***) + +goalw BinFn.thy [bin_add_def] "!!w. w: bin ==> bin_add(Plus,w) = w"; +by (ASM_SIMP_TAC bin_ss 1); +val bin_add_Plus = result(); + +goalw BinFn.thy [bin_add_def] "!!w. w: bin ==> bin_add(Minus,w) = bin_pred(w)"; +by (ASM_SIMP_TAC bin_ss 1); +val bin_add_Minus = result(); + +goalw BinFn.thy [bin_add_def] "bin_add(v$$x,Plus) = v$$x"; +by (SIMP_TAC bin_ss 1); +val bin_add_Bcons_Plus = result(); + +goalw BinFn.thy [bin_add_def] "bin_add(v$$x,Minus) = bin_pred(v$$x)"; +by (SIMP_TAC bin_ss 1); +val bin_add_Bcons_Minus = result(); + +goalw BinFn.thy [bin_add_def] + "!!w y. [| w: bin; y: bool |] ==> \ +\ bin_add(v$$x, w$$y) = \ +\ bin_add(v, cond(x and y, bin_succ(w), w)) $$ (x xor y)"; +by (ASM_SIMP_TAC bin_ss 1); +val bin_add_Bcons_Bcons = result(); + +val bin_add_rews = [bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, + bin_add_Bcons_Minus, bin_add_Bcons_Bcons, + integ_of_bin_succ, integ_of_bin_pred]; + +val bin_add_ss = bin_ss addrews ([bool_subset_nat RS subsetD] @ bin_add_rews); + +goal BinFn.thy + "!!v. v: bin ==> \ +\ ALL w: bin. integ_of_bin(bin_add(v,w)) = \ +\ integ_of_bin(v) $+ integ_of_bin(w)"; +by (etac Bin.induct 1); +by (SIMP_TAC bin_add_ss 1); +by (SIMP_TAC bin_add_ss 1); +by (rtac ballI 1); +by (bin_ind_tac "wa" [] 1); +by (ASM_SIMP_TAC (bin_add_ss addrews [zadd_0_right]) 1); +by (ASM_SIMP_TAC bin_add_ss 1); +by (REPEAT (ares_tac (zadd_commute::typechecks) 1)); +by (etac boolE 1); +by (ASM_SIMP_TAC (bin_add_ss addrews [zadd_assoc, zadd_swap_pairs]) 2); +by (REPEAT (ares_tac ([refl, zadd_kill, zadd_assoc_swap_kill]@typechecks) 2)); +by (etac boolE 1); +by (ALLGOALS (ASM_SIMP_TAC (bin_add_ss addrews [zadd_assoc,zadd_swap_pairs]))); +by (REPEAT (ares_tac ([refl, zadd_kill, zadd_assoc_swap_kill RS sym]@ + typechecks) 1)); +val integ_of_bin_add_lemma = result(); + +val integ_of_bin_add = integ_of_bin_add_lemma RS bspec; + + +(*** bin_add: binary multiplication ***) + +val bin_mult_ss = + bin_ss addrews (bin_recs bin_mult_def @ + [integ_of_bin_minus, integ_of_bin_add]); + + +val major::prems = goal BinFn.thy + "[| v: bin; w: bin |] ==> \ +\ integ_of_bin(bin_mult(v,w)) = \ +\ integ_of_bin(v) $* integ_of_bin(w)"; +by (cut_facts_tac prems 1); +by (bin_ind_tac "v" [major] 1); +by (SIMP_TAC (bin_mult_ss addrews [zmult_0]) 1); +by (SIMP_TAC (bin_mult_ss addrews [zmult_1,zmult_zminus]) 1); +by (etac boolE 1); +by (ASM_SIMP_TAC (bin_mult_ss addrews [zadd_zmult_distrib]) 2); +by (ASM_SIMP_TAC + (bin_mult_ss addrews [zadd_zmult_distrib, zmult_1, zadd_assoc]) 1); +by (REPEAT (ares_tac ([zadd_commute, zadd_assoc_swap_kill RS sym]@ + typechecks) 1)); +val integ_of_bin_mult = result(); + +(**** Computations ****) + +(** extra rules for bin_succ, bin_pred **) + +val [bin_succ_Plus, bin_succ_Minus, _] = bin_recs bin_succ_def; +val [bin_pred_Plus, bin_pred_Minus, _] = bin_recs bin_pred_def; + +goal BinFn.thy "bin_succ(w$$1) = bin_succ(w) $$ 0"; +by (SIMP_TAC carry_ss 1); +val bin_succ_Bcons1 = result(); + +goal BinFn.thy "bin_succ(w$$0) = w$$1"; +by (SIMP_TAC carry_ss 1); +val bin_succ_Bcons0 = result(); + +goal BinFn.thy "bin_pred(w$$1) = w$$0"; +by (SIMP_TAC carry_ss 1); +val bin_pred_Bcons1 = result(); + +goal BinFn.thy "bin_pred(w$$0) = bin_pred(w) $$ 1"; +by (SIMP_TAC carry_ss 1); +val bin_pred_Bcons0 = result(); + +(** extra rules for bin_minus **) + +val [bin_minus_Plus, bin_minus_Minus, _] = bin_recs bin_minus_def; + +goal BinFn.thy "bin_minus(w$$1) = bin_pred(bin_minus(w) $$ 0)"; +by (SIMP_TAC bin_minus_ss 1); +val bin_minus_Bcons1 = result(); + +goal BinFn.thy "bin_minus(w$$0) = bin_minus(w) $$ 0"; +by (SIMP_TAC bin_minus_ss 1); +val bin_minus_Bcons0 = result(); + +(** extra rules for bin_add **) + +goal BinFn.thy + "!!w. w: bin ==> bin_add(v$$1, w$$1) = bin_add(v, bin_succ(w)) $$ 0"; +by (ASM_SIMP_TAC bin_add_ss 1); +val bin_add_Bcons_Bcons11 = result(); + +goal BinFn.thy + "!!w. w: bin ==> bin_add(v$$1, w$$0) = bin_add(v,w) $$ 1"; +by (ASM_SIMP_TAC bin_add_ss 1); +val bin_add_Bcons_Bcons10 = result(); + +goal BinFn.thy + "!!w y.[| w: bin; y: bool |] ==> bin_add(v$$0, w$$y) = bin_add(v,w) $$ y"; +by (ASM_SIMP_TAC bin_add_ss 1); +val bin_add_Bcons_Bcons0 = result(); + +(** extra rules for bin_mult **) + +val [bin_mult_Plus, bin_mult_Minus, _] = bin_recs bin_mult_def; + +goal BinFn.thy "bin_mult(v$$1, w) = bin_add(bin_mult(v,w)$$0, w)"; +by (SIMP_TAC bin_mult_ss 1); +val bin_mult_Bcons1 = result(); + +goal BinFn.thy "bin_mult(v$$0, w) = bin_mult(v,w)$$0"; +by (SIMP_TAC bin_mult_ss 1); +val bin_mult_Bcons0 = result(); + + +(*** The computation simpset ***) + +val bin_comp_ss = carry_ss addrews + [bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, bin_add_Bcons_Minus, + bin_add_Bcons_Bcons0, bin_add_Bcons_Bcons10, bin_add_Bcons_Bcons11] + setauto (type_auto_tac bin_typechecks0); + +val bin_comp_ss = integ_ss + addcongs bin_congs + addrews [bin_succ_Plus, bin_succ_Minus, + bin_succ_Bcons1, bin_succ_Bcons0, + bin_pred_Plus, bin_pred_Minus, + bin_pred_Bcons1, bin_pred_Bcons0, + bin_minus_Plus, bin_minus_Minus, + bin_minus_Bcons1, bin_minus_Bcons0, + bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, + bin_add_Bcons_Minus, bin_add_Bcons_Bcons0, + bin_add_Bcons_Bcons10, bin_add_Bcons_Bcons11, + bin_mult_Plus, bin_mult_Minus, + bin_mult_Bcons1, bin_mult_Bcons0] + setauto (type_auto_tac ([bool_1I, bool_0I] @ bin_typechecks0)); + +(*** Examples of performing binary arithmetic by simplification ***) + +proof_timing := true; +(*All runtimes below are on a SPARCserver 10*) + +(* 13+19 = 32 *) +goal BinFn.thy + "bin_add(Plus$$1$$1$$0$$1, Plus$$1$$0$$0$$1$$1) = Plus$$1$$0$$0$$0$$0$$0"; +by (SIMP_TAC bin_comp_ss 1); (*1.9 secs*) +result(); + +bin_add(binary_of_int 13, binary_of_int 19); + +(* 1234+5678 = 6912 *) +goal BinFn.thy + "bin_add(Plus$$1$$0$$0$$1$$1$$0$$1$$0$$0$$1$$0, \ +\ Plus$$1$$0$$1$$1$$0$$0$$0$$1$$0$$1$$1$$1$$0) = \ +\ Plus$$1$$1$$0$$1$$1$$0$$0$$0$$0$$0$$0$$0$$0"; +by (SIMP_TAC bin_comp_ss 1); (*8.9 secs*) +result(); + +bin_add(binary_of_int 1234, binary_of_int 5678); + +(* 1359-2468 = ~1109 *) +goal BinFn.thy + "bin_add(Plus$$1$$0$$1$$0$$1$$0$$0$$1$$1$$1$$1, \ +\ Minus$$0$$1$$1$$0$$0$$1$$0$$1$$1$$1$$0$$0) = \ +\ Minus$$1$$0$$1$$1$$1$$0$$1$$0$$1$$0$$1$$1"; +by (SIMP_TAC bin_comp_ss 1); (*7.4 secs*) +result(); + +bin_add(binary_of_int 1359, binary_of_int ~2468); + +(* 93746-46375 = 47371 *) +goal BinFn.thy + "bin_add(Plus$$1$$0$$1$$1$$0$$1$$1$$1$$0$$0$$0$$1$$1$$0$$0$$1$$0, \ +\ Minus$$0$$1$$0$$0$$1$$0$$1$$0$$1$$1$$0$$1$$1$$0$$0$$1) = \ +\ Plus$$0$$1$$0$$1$$1$$1$$0$$0$$1$$0$$0$$0$$0$$1$$0$$1$$1"; +by (SIMP_TAC bin_comp_ss 1); (*13.7 secs*) +result(); + +bin_add(binary_of_int 93746, binary_of_int ~46375); + +(* negation of 65745 *) +goal BinFn.thy + "bin_minus(Plus$$1$$0$$0$$0$$0$$0$$0$$0$$0$$1$$1$$0$$1$$0$$0$$0$$1) = \ +\ Minus$$0$$1$$1$$1$$1$$1$$1$$1$$1$$0$$0$$1$$0$$1$$1$$1$$1"; +by (SIMP_TAC bin_comp_ss 1); (*2.8 secs*) +result(); + +bin_minus(binary_of_int 65745); + +(* negation of ~54321 *) +goal BinFn.thy + "bin_minus(Minus$$0$$0$$1$$0$$1$$0$$1$$1$$1$$1$$0$$0$$1$$1$$1$$1) = \ +\ Plus$$0$$1$$1$$0$$1$$0$$1$$0$$0$$0$$0$$1$$1$$0$$0$$0$$1"; +by (SIMP_TAC bin_comp_ss 1); (*3.3 secs*) +result(); + +bin_minus(binary_of_int ~54321); + +(* 13*19 = 247 *) +goal BinFn.thy "bin_mult(Plus$$1$$1$$0$$1, Plus$$1$$0$$0$$1$$1) = \ +\ Plus$$1$$1$$1$$1$$0$$1$$1$$1"; +by (SIMP_TAC bin_comp_ss 1); (*4.4 secs*) +result(); + +bin_mult(binary_of_int 13, binary_of_int 19); + +(* ~84 * 51 = ~4284 *) +goal BinFn.thy + "bin_mult(Minus$$0$$1$$0$$1$$1$$0$$0, Plus$$1$$1$$0$$0$$1$$1) = \ +\ Minus$$0$$1$$1$$1$$1$$0$$1$$0$$0$$0$$1$$0$$0"; +by (SIMP_TAC bin_comp_ss 1); (*9.2 secs*) +result(); + +bin_mult(binary_of_int ~84, binary_of_int 51); + +(* 255*255 = 65025; the worst case for 8-bit operands *) +goal BinFn.thy + "bin_mult(Plus$$1$$1$$1$$1$$1$$1$$1$$1, \ +\ Plus$$1$$1$$1$$1$$1$$1$$1$$1) = \ +\ Plus$$1$$1$$1$$1$$1$$1$$1$$0$$0$$0$$0$$0$$0$$0$$0$$1"; +by (SIMP_TAC bin_comp_ss 1); (*38.4 secs*) +result(); + +bin_mult(binary_of_int 255, binary_of_int 255); + +(***************** TOO SLOW TO INCLUDE IN TEST RUNS +(* 1359 * ~2468 = ~3354012 *) +goal BinFn.thy + "bin_mult(Plus$$1$$0$$1$$0$$1$$0$$0$$1$$1$$1$$1, \ +\ Minus$$0$$1$$1$$0$$0$$1$$0$$1$$1$$1$$0$$0) = \ +\ Minus$$0$$0$$1$$1$$0$$0$$1$$1$$0$$1$$0$$0$$1$$0$$0$$1$$1$$0$$0$$1$$0$$0"; +by (SIMP_TAC bin_comp_ss 1); (*54.8 secs*) +result(); +****************) + +bin_mult(binary_of_int 1359, binary_of_int ~2468); diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/BinFn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/BinFn.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,52 @@ +(* Title: ZF/bin + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Arithmetic on binary integers. +*) + +BinFn = Integ + Bin + +consts + bin_rec :: "[i, i, i, [i,i,i]=>i] => i" + integ_of_bin :: "i=>i" + bin_succ :: "i=>i" + bin_pred :: "i=>i" + bin_minus :: "i=>i" + bin_add,bin_mult :: "[i,i]=>i" + +rules + + bin_rec_def + "bin_rec(z,a,b,h) == \ +\ Vrec(z, %z g. bin_case(a, b, %w x. h(w, x, g`w), z))" + + integ_of_bin_def + "integ_of_bin(w) == bin_rec(w, $#0, $~($#1), %w x r. $#x $+ r $+ r)" + + bin_succ_def + "bin_succ(w0) == bin_rec(w0, Plus$$1, Plus, %w x r. cond(x, r$$0, w$$1))" + + bin_pred_def + "bin_pred(w0) == \ +\ bin_rec(w0, Minus, Minus$$0, %w x r. cond(x, w$$0, r$$1))" + + bin_minus_def + "bin_minus(w0) == \ +\ bin_rec(w0, Plus, Plus$$1, %w x r. cond(x, bin_pred(r$$0), r$$0))" + + bin_add_def + "bin_add(v0,w0) == \ +\ bin_rec(v0, \ +\ lam w:bin. w, \ +\ lam w:bin. bin_pred(w), \ +\ %v x r. lam w1:bin. \ +\ bin_rec(w1, v$$x, bin_pred(v$$x), \ +\ %w y s. (r`cond(x and y, bin_succ(w), w)) \ +\ $$ (x xor y))) ` w0" + + bin_mult_def + "bin_mult(v0,w) == \ +\ bin_rec(v0, Plus, bin_minus(w), \ +\ %v x r. cond(x, bin_add(r$$0,w), r$$0))" +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Comb.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Comb.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,45 @@ +(* Title: ZF/ex/comb.ML + ID: $Id$ + Author: Lawrence C Paulson + Copyright 1993 University of Cambridge + +Datatype definition of combinators S and K + +J. Camilleri and T. F. Melham. +Reasoning with Inductively Defined Relations in the HOL Theorem Prover. +Report 265, University of Cambridge Computer Laboratory, 1992. +*) + + +(*Example of a datatype with mixfix syntax for some constructors*) +structure Comb = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("comb", "univ(0)", + [(["K","S"], "i"), + (["op #"], "[i,i]=>i")])]; + val rec_styp = "i"; + val ext = Some (NewSext { + mixfix = + [Infixl("#", "[i,i] => i", 90)], + xrules = [], + parse_ast_translation = [], + parse_preproc = None, + parse_postproc = None, + parse_translation = [], + print_translation = [], + print_preproc = None, + print_postproc = None, + print_ast_translation = []}); + val sintrs = + ["K : comb", + "S : comb", + "[| p: comb; q: comb |] ==> p#q : comb"]; + val monos = []; + val type_intrs = data_typechecks; + val type_elims = []); + +val [K_comb,S_comb,Ap_comb] = Comb.intrs; + +val Ap_E = Comb.mk_cases Comb.con_defs "p#q : comb"; + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Contract0.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Contract0.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,117 @@ +(* Title: ZF/ex/contract.ML + ID: $Id$ + Author: Tobias Nipkow & Lawrence C Paulson + Copyright 1992 University of Cambridge + +For ex/contract.thy. +*) + +open Contract0; + +structure Contract = Inductive_Fun + (val thy = Contract0.thy; + val rec_doms = [("contract","comb*comb")]; + val sintrs = + ["[| p:comb; q:comb |] ==> K#p#q -1-> p", + "[| p:comb; q:comb; r:comb |] ==> S#p#q#r -1-> (p#r)#(q#r)", + "[| p-1->q; r:comb |] ==> p#r -1-> q#r", + "[| p-1->q; r:comb |] ==> r#p -1-> r#q"]; + val monos = []; + val con_defs = []; + val type_intrs = Comb.intrs@[SigmaI]; + val type_elims = [SigmaE2]); + +val [K_contract,S_contract,Ap_contract1,Ap_contract2] = Contract.intrs; + +val contract_induct = standard + (Contract.mutual_induct RS spec RS spec RSN (2,rev_mp)); + +(*For type checking: replaces a-1->b by a,b:comb *) +val contract_combE2 = Contract.dom_subset RS subsetD RS SigmaE2; +val contract_combD1 = Contract.dom_subset RS subsetD RS SigmaD1; +val contract_combD2 = Contract.dom_subset RS subsetD RS SigmaD2; + +goal Contract.thy "field(contract) = comb"; +by (fast_tac (ZF_cs addIs [equalityI,K_contract] addSEs [contract_combE2]) 1); +val field_contract_eq = result(); + +val reduction_refl = standard + (field_contract_eq RS equalityD2 RS subsetD RS rtrancl_refl); + +val rtrancl_into_rtrancl2 = standard + (r_into_rtrancl RS (trans_rtrancl RS transD)); + +val reduction_rls = [reduction_refl, K_contract, S_contract, + K_contract RS rtrancl_into_rtrancl2, + S_contract RS rtrancl_into_rtrancl2, + Ap_contract1 RS rtrancl_into_rtrancl2, + Ap_contract2 RS rtrancl_into_rtrancl2]; + +goalw Contract.thy [I_def] "!!p. p:comb ==> I#p ---> p"; +by (REPEAT (ares_tac (Comb.intrs @ reduction_rls) 1)); +val I_reduce = result(); + +goalw Contract.thy [I_def] "I: comb"; +by (REPEAT (ares_tac Comb.intrs 1)); +val I_comb = result(); + +(** Non-contraction results **) + +(*Derive a case for each combinator constructor*) +val K_contract_case = Contract.mk_cases Comb.con_defs "K -1-> r"; +val S_contract_case = Contract.mk_cases Comb.con_defs "S -1-> r"; +val Ap_contract_case = Contract.mk_cases Comb.con_defs "p#q -1-> r"; + +val contract_cs = + ZF_cs addSIs Comb.intrs + addIs Contract.intrs + addSEs [contract_combD1,contract_combD2] (*type checking*) + addSEs [K_contract_case, S_contract_case, Ap_contract_case] + addSEs Comb.free_SEs; + +goalw Contract.thy [I_def] "!!r. I -1-> r ==> P"; +by (fast_tac contract_cs 1); +val I_contract_case = result(); + +goal Contract.thy "!!p r. K#p -1-> r ==> (EX q. r = K#q & p -1-> q)"; +by (fast_tac contract_cs 1); +val K1_contractD = result(); + +goal Contract.thy "!!p r. [| p ---> q; r: comb |] ==> p#r ---> q#r"; +by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1); +by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1); +by (etac rtrancl_induct 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +by (etac (trans_rtrancl RS transD) 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +val Ap_reduce1 = result(); + +goal Contract.thy "!!p r. [| p ---> q; r: comb |] ==> r#p ---> r#q"; +by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1); +by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1); +by (etac rtrancl_induct 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +by (etac (trans_rtrancl RS transD) 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +val Ap_reduce2 = result(); + +(** Counterexample to the diamond property for -1-> **) + +goal Contract.thy "K#I#(I#I) -1-> I"; +by (REPEAT (ares_tac [K_contract, I_comb, Ap_comb] 1)); +val KIII_contract1 = result(); + +goalw Contract.thy [I_def] "K#I#(I#I) -1-> K#I#((K#I)#(K#I))"; +by (DEPTH_SOLVE (resolve_tac (Comb.intrs @ Contract.intrs) 1)); +val KIII_contract2 = result(); + +goal Contract.thy "K#I#((K#I)#(K#I)) -1-> I"; +by (REPEAT (ares_tac (Comb.intrs @ [K_contract, I_comb]) 1)); +val KIII_contract3 = result(); + +goalw Contract.thy [diamond_def] "~ diamond(contract)"; +by (fast_tac (ZF_cs addIs [KIII_contract1,KIII_contract2,KIII_contract3] + addSEs [I_contract_case]) 1); +val not_diamond_contract = result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Contract0.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Contract0.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,36 @@ +(* Title: ZF/ex/contract.thy + ID: $Id$ + Author: Tobias Nipkow & Lawrence C Paulson + Copyright 1993 University of Cambridge + +Inductive definition of (1-step) contractions and (mult-step) reductions +*) + +Contract0 = Comb + +consts + diamond :: "i => o" + I :: "i" + + contract :: "i" + "-1->" :: "[i,i] => o" (infixl 50) + "--->" :: "[i,i] => o" (infixl 50) + + parcontract :: "i" + "=1=>" :: "[i,i] => o" (infixl 50) + "===>" :: "[i,i] => o" (infixl 50) + +translations + "p -1-> q" == " : contract" + "p ---> q" == " : contract^*" + "p =1=> q" == " : parcontract" + "p ===> q" == " : parcontract^+" + +rules + + diamond_def "diamond(r) == ALL x y. :r --> \ +\ (ALL y'. :r --> \ +\ (EX z. :r & : r))" + + I_def "I == S#K#K" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Enum.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Enum.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,33 @@ +(* Title: ZF/ex/enum + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Example of a BIG enumeration type + +Can go up to at least 100 constructors, but it takes over 10 minutes... +*) + + +(*An enumeration type with 60 contructors! -- takes about 214 seconds!*) +fun mk_ids a 0 = [] + | mk_ids a n = a :: mk_ids (bump_string a) (n-1); + +val consts = mk_ids "con1" 60; + +structure Enum = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("enum", "univ(0)", + [(consts, "i")])]; + val rec_styp = "i"; + val ext = None + val sintrs = map (fn const => const ^ " : enum") consts; + val monos = []; + val type_intrs = data_typechecks + val type_elims = []); + +goal Enum.thy "~ con59=con60"; +by (SIMP_TAC (ZF_ss addrews Enum.free_iffs) 1); (*2.3 secs*) +result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Equiv.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Equiv.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,268 @@ +(* Title: ZF/ex/equiv.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For equiv.thy. Equivalence relations in Zermelo-Fraenkel Set Theory +*) + +val RSLIST = curry (op MRS); + +open Equiv; + +(*** Suppes, Theorem 70: r is an equiv relation iff converse(r) O r = r ***) + +(** first half: equiv(A,r) ==> converse(r) O r = r **) + +goalw Equiv.thy [trans_def,sym_def] + "!!r. [| sym(r); trans(r) |] ==> converse(r) O r <= r"; +by (fast_tac (ZF_cs addSEs [converseD,compE]) 1); +val sym_trans_comp_subset = result(); + +goalw Equiv.thy [refl_def] + "!!A r. refl(A,r) ==> r <= converse(r) O r"; +by (fast_tac (ZF_cs addSIs [converseI] addIs [compI]) 1); +val refl_comp_subset = result(); + +goalw Equiv.thy [equiv_def] + "!!A r. equiv(A,r) ==> converse(r) O r = r"; +by (rtac equalityI 1); +by (REPEAT (ares_tac [sym_trans_comp_subset, refl_comp_subset] 1 + ORELSE etac conjE 1)); +val equiv_comp_eq = result(); + +(*second half*) +goalw Equiv.thy [equiv_def,refl_def,sym_def,trans_def] + "!!A r. [| converse(r) O r = r; domain(r) = A |] ==> equiv(A,r)"; +by (etac equalityE 1); +by (subgoal_tac "ALL x y. : r --> : r" 1); +by (safe_tac ZF_cs); +by (fast_tac (ZF_cs addSIs [converseI] addIs [compI]) 3); +by (ALLGOALS (fast_tac + (ZF_cs addSIs [converseI] addIs [compI] addSEs [compE]))); +by flexflex_tac; +val comp_equivI = result(); + +(** Equivalence classes **) + +(*Lemma for the next result*) +goalw Equiv.thy [equiv_def,trans_def,sym_def] + "!!A r. [| equiv(A,r); : r |] ==> r``{a} <= r``{b}"; +by (fast_tac ZF_cs 1); +val equiv_class_subset = result(); + +goal Equiv.thy "!!A r. [| equiv(A,r); : r |] ==> r``{a} = r``{b}"; +by (REPEAT (ares_tac [equalityI, equiv_class_subset] 1)); +by (rewrite_goals_tac [equiv_def,sym_def]); +by (fast_tac ZF_cs 1); +val equiv_class_eq = result(); + +val prems = goalw Equiv.thy [equiv_def,refl_def] + "[| equiv(A,r); a: A |] ==> a: r``{a}"; +by (cut_facts_tac prems 1); +by (fast_tac ZF_cs 1); +val equiv_class_self = result(); + +(*Lemma for the next result*) +goalw Equiv.thy [equiv_def,refl_def] + "!!A r. [| equiv(A,r); r``{b} <= r``{a}; b: A |] ==> : r"; +by (fast_tac ZF_cs 1); +val subset_equiv_class = result(); + +val prems = goal Equiv.thy + "[| r``{a} = r``{b}; equiv(A,r); b: A |] ==> : r"; +by (REPEAT (resolve_tac (prems @ [equalityD2, subset_equiv_class]) 1)); +val eq_equiv_class = result(); + +(*thus r``{a} = r``{b} as well*) +goalw Equiv.thy [equiv_def,trans_def,sym_def] + "!!A r. [| equiv(A,r); x: (r``{a} Int r``{b}) |] ==> : r"; +by (fast_tac ZF_cs 1); +val equiv_class_nondisjoint = result(); + +val [major] = goalw Equiv.thy [equiv_def,refl_def] + "equiv(A,r) ==> r <= A*A"; +by (rtac (major RS conjunct1 RS conjunct1) 1); +val equiv_type = result(); + +goal Equiv.thy + "!!A r. equiv(A,r) ==> : r <-> r``{x} = r``{y} & x:A & y:A"; +by (fast_tac (ZF_cs addIs [eq_equiv_class, equiv_class_eq] + addDs [equiv_type]) 1); +val equiv_class_eq_iff = result(); + +goal Equiv.thy + "!!A r. [| equiv(A,r); x: A; y: A |] ==> r``{x} = r``{y} <-> : r"; +by (fast_tac (ZF_cs addIs [eq_equiv_class, equiv_class_eq] + addDs [equiv_type]) 1); +val eq_equiv_class_iff = result(); + +(*** Quotients ***) + +(** Introduction/elimination rules -- needed? **) + +val prems = goalw Equiv.thy [quotient_def] "x:A ==> r``{x}: A/r"; +by (rtac RepFunI 1); +by (resolve_tac prems 1); +val quotientI = result(); + +val major::prems = goalw Equiv.thy [quotient_def] + "[| X: A/r; !!x. [| X = r``{x}; x:A |] ==> P |] \ +\ ==> P"; +by (rtac (major RS RepFunE) 1); +by (eresolve_tac prems 1); +by (assume_tac 1); +val quotientE = result(); + +goalw Equiv.thy [equiv_def,refl_def,quotient_def] + "!!A r. equiv(A,r) ==> Union(A/r) = A"; +by (fast_tac eq_cs 1); +val Union_quotient = result(); + +goalw Equiv.thy [quotient_def] + "!!A r. [| equiv(A,r); X: A/r; Y: A/r |] ==> X=Y | (X Int Y <= 0)"; +by (safe_tac (ZF_cs addSIs [equiv_class_eq])); +by (assume_tac 1); +by (rewrite_goals_tac [equiv_def,trans_def,sym_def]); +by (fast_tac ZF_cs 1); +val quotient_disj = result(); + +(**** Defining unary operations upon equivalence classes ****) + +(** These proofs really require as local premises + equiv(A,r); congruent(r,b) +**) + +(*Conversion rule*) +val prems as [equivA,bcong,_] = goal Equiv.thy + "[| equiv(A,r); congruent(r,b); a: A |] ==> (UN x:r``{a}. b(x)) = b(a)"; +by (cut_facts_tac prems 1); +by (rtac UN_singleton 1); +by (etac equiv_class_self 1); +by (assume_tac 1); +by (rewrite_goals_tac [equiv_def,sym_def,congruent_def]); +by (fast_tac ZF_cs 1); +val UN_equiv_class = result(); + +(*Resolve th against the "local" premises*) +val localize = RSLIST [equivA,bcong]; + +(*type checking of UN x:r``{a}. b(x) *) +val _::_::prems = goalw Equiv.thy [quotient_def] + "[| equiv(A,r); congruent(r,b); X: A/r; \ +\ !!x. x : A ==> b(x) : B |] \ +\ ==> (UN x:X. b(x)) : B"; +by (cut_facts_tac prems 1); +by (safe_tac ZF_cs); +by (rtac (localize UN_equiv_class RS ssubst) 1); +by (REPEAT (ares_tac prems 1)); +val UN_equiv_class_type = result(); + +(*Sufficient conditions for injectiveness. Could weaken premises! + major premise could be an inclusion; bcong could be !!y. y:A ==> b(y):B +*) +val _::_::prems = goalw Equiv.thy [quotient_def] + "[| equiv(A,r); congruent(r,b); \ +\ (UN x:X. b(x))=(UN y:Y. b(y)); X: A/r; Y: A/r; \ +\ !!x y. [| x:A; y:A; b(x)=b(y) |] ==> :r |] \ +\ ==> X=Y"; +by (cut_facts_tac prems 1); +by (safe_tac ZF_cs); +by (rtac (equivA RS equiv_class_eq) 1); +by (REPEAT (ares_tac prems 1)); +by (etac box_equals 1); +by (REPEAT (ares_tac [localize UN_equiv_class] 1)); +val UN_equiv_class_inject = result(); + + +(**** Defining binary operations upon equivalence classes ****) + + +goalw Equiv.thy [congruent_def,congruent2_def,equiv_def,refl_def] + "!!A r. [| equiv(A,r); congruent2(r,b); a: A |] ==> congruent(r,b(a))"; +by (fast_tac ZF_cs 1); +val congruent2_implies_congruent = result(); + +val equivA::prems = goalw Equiv.thy [congruent_def] + "[| equiv(A,r); congruent2(r,b); a: A |] ==> \ +\ congruent(r, %x1. UN x2:r``{a}. b(x1,x2))"; +by (cut_facts_tac (equivA::prems) 1); +by (safe_tac ZF_cs); +by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1); +by (assume_tac 1); +by (ASM_SIMP_TAC (ZF_ss addrews [equivA RS UN_equiv_class, + congruent2_implies_congruent]) 1); +by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]); +by (fast_tac ZF_cs 1); +val congruent2_implies_congruent_UN = result(); + +val prems as equivA::_ = goal Equiv.thy + "[| equiv(A,r); congruent2(r,b); a1: A; a2: A |] \ +\ ==> (UN x1:r``{a1}. UN x2:r``{a2}. b(x1,x2)) = b(a1,a2)"; +by (cut_facts_tac prems 1); +by (ASM_SIMP_TAC (ZF_ss addrews [equivA RS UN_equiv_class, + congruent2_implies_congruent, + congruent2_implies_congruent_UN]) 1); +val UN_equiv_class2 = result(); + +(*type checking*) +val prems = goalw Equiv.thy [quotient_def] + "[| equiv(A,r); congruent2(r,b); \ +\ X1: A/r; X2: A/r; \ +\ !!x1 x2. [| x1: A; x2: A |] ==> b(x1,x2) : B |] \ +\ ==> (UN x1:X1. UN x2:X2. b(x1,x2)) : B"; +by (cut_facts_tac prems 1); +by (safe_tac ZF_cs); +by (REPEAT (ares_tac (prems@[UN_equiv_class_type, + congruent2_implies_congruent_UN, + congruent2_implies_congruent, quotientI]) 1)); +val UN_equiv_class_type2 = result(); + + +(*Suggested by John Harrison -- the two subproofs may be MUCH simpler + than the direct proof*) +val prems = goalw Equiv.thy [congruent2_def,equiv_def,refl_def] + "[| equiv(A,r); \ +\ !! y z w. [| w: A; : r |] ==> b(y,w) = b(z,w); \ +\ !! y z w. [| w: A; : r |] ==> b(w,y) = b(w,z) \ +\ |] ==> congruent2(r,b)"; +by (cut_facts_tac prems 1); +by (safe_tac ZF_cs); +by (rtac trans 1); +by (REPEAT (ares_tac prems 1 + ORELSE etac (subsetD RS SigmaE2) 1 THEN assume_tac 2 THEN assume_tac 1)); +val congruent2I = result(); + +val [equivA,commute,congt] = goal Equiv.thy + "[| equiv(A,r); \ +\ !! y z w. [| y: A; z: A |] ==> b(y,z) = b(z,y); \ +\ !! y z w. [| w: A; : r |] ==> b(w,y) = b(w,z) \ +\ |] ==> congruent2(r,b)"; +by (resolve_tac [equivA RS congruent2I] 1); +by (rtac (commute RS trans) 1); +by (rtac (commute RS trans RS sym) 3); +by (rtac sym 5); +by (REPEAT (ares_tac [congt] 1 + ORELSE etac (equivA RS equiv_type RS subsetD RS SigmaE2) 1)); +val congruent2_commuteI = result(); + +(***OBSOLETE VERSION +(*Rules congruentI and congruentD would simplify use of rewriting below*) +val [equivA,ZinA,congt,commute] = goalw Equiv.thy [quotient_def] + "[| equiv(A,r); Z: A/r; \ +\ !!w. [| w: A |] ==> congruent(r, %z.b(w,z)); \ +\ !!x y. [| x: A; y: A |] ==> b(y,x) = b(x,y) \ +\ |] ==> congruent(r, %w. UN z: Z. b(w,z))"; +val congt' = rewrite_rule [congruent_def] congt; +by (cut_facts_tac [ZinA,congt] 1); +by (rewtac congruent_def); +by (safe_tac ZF_cs); +by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1); +by (assume_tac 1); +by (ASM_SIMP_TAC (ZF_ss addrews [congt RS (equivA RS UN_equiv_class)]) 1); +by (rtac (commute RS trans) 1); +by (rtac (commute RS trans RS sym) 3); +by (rtac sym 5); +by (REPEAT (ares_tac [congt' RS spec RS spec RS mp] 1)); +val congruent_commuteI = result(); +***) diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Equiv.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Equiv.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,28 @@ +(* Title: ZF/ex/equiv.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Equivalence relations in Zermelo-Fraenkel Set Theory +*) + +Equiv = Trancl + +consts + refl,equiv :: "[i,i]=>o" + sym :: "i=>o" + "'/" :: "[i,i]=>i" (infixl 90) (*set of equiv classes*) + congruent :: "[i,i=>i]=>o" + congruent2 :: "[i,[i,i]=>i]=>o" + +rules + refl_def "refl(A,r) == r <= (A*A) & (ALL x: A. : r)" + sym_def "sym(r) == ALL x y. : r --> : r" + equiv_def "equiv(A,r) == refl(A,r) & sym(r) & trans(r)" + quotient_def "A/r == {r``{x} . x:A}" + congruent_def "congruent(r,b) == ALL y z. :r --> b(y)=b(z)" + + congruent2_def + "congruent2(r,b) == ALL y1 z1 y2 z2. \ +\ :r --> :r --> b(y1,y2) = b(z1,z2)" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Integ.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Integ.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,437 @@ +(* Title: ZF/ex/integ.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For integ.thy. The integers as equivalence classes over nat*nat. + +Could also prove... +"znegative(z) ==> $# zmagnitude(z) = $~ z" +"~ znegative(z) ==> $# zmagnitude(z) = z" +$< is a linear ordering +$+ and $* are monotonic wrt $< +*) + +open Integ; + +val [add_cong] = mk_congs Arith.thy ["op #+"]; + +(*** Proving that intrel is an equivalence relation ***) + +val prems = goal Arith.thy + "[| m #+ n = m' #+ n'; m: nat; m': nat |] \ +\ ==> m #+ (n #+ k) = m' #+ (n' #+ k)"; +by (ASM_SIMP_TAC (arith_ss addrews ([add_assoc RS sym] @ prems)) 1); +val add_assoc_cong = result(); + +val prems = goal Arith.thy + "[| m: nat; n: nat |] \ +\ ==> m #+ (n #+ k) = n #+ (m #+ k)"; +by (REPEAT (resolve_tac ([add_commute RS add_assoc_cong] @ prems) 1)); +val add_assoc_swap = result(); + +val add_kill = (refl RS add_cong); +val add_assoc_swap_kill = add_kill RSN (3, add_assoc_swap RS trans); + +(*By luck, requires no typing premises for y1, y2,y3*) +val eqa::eqb::prems = goal Arith.thy + "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2; \ +\ x1: nat; x2: nat; x3: nat |] ==> x1 #+ y3 = x3 #+ y1"; +by (res_inst_tac [("k","x2")] add_left_cancel 1); +by (resolve_tac prems 1); +by (rtac (add_assoc_swap RS trans) 1 THEN typechk_tac prems); +by (rtac (eqb RS ssubst) 1); +by (rtac (add_assoc_swap RS trans) 1 THEN typechk_tac prems); +by (rtac (eqa RS ssubst) 1); +by (rtac (add_assoc_swap) 1 THEN typechk_tac prems); +val integ_trans_lemma = result(); + +(** Natural deduction for intrel **) + +val prems = goalw Integ.thy [intrel_def] + "[| x1#+y2 = x2#+y1; x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ +\ <,>: intrel"; +by (fast_tac (ZF_cs addIs prems) 1); +val intrelI = result(); + +(*intrelE is hard to derive because fast_tac tries hyp_subst_tac so soon*) +goalw Integ.thy [intrel_def] + "p: intrel --> (EX x1 y1 x2 y2. \ +\ p = <,> & x1#+y2 = x2#+y1 & \ +\ x1: nat & y1: nat & x2: nat & y2: nat)"; +by (fast_tac ZF_cs 1); +val intrelE_lemma = result(); + +val [major,minor] = goal Integ.thy + "[| p: intrel; \ +\ !!x1 y1 x2 y2. [| p = <,>; x1#+y2 = x2#+y1; \ +\ x1: nat; y1: nat; x2: nat; y2: nat |] ==> Q |] \ +\ ==> Q"; +by (cut_facts_tac [major RS (intrelE_lemma RS mp)] 1); +by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1)); +val intrelE = result(); + +val intrel_cs = ZF_cs addSIs [intrelI] addSEs [intrelE]; + +goal Integ.thy + "<,>: intrel <-> \ +\ x1#+y2 = x2#+y1 & x1: nat & y1: nat & x2: nat & y2: nat"; +by (fast_tac intrel_cs 1); +val intrel_iff = result(); + +val prems = goalw Integ.thy [equiv_def] "equiv(nat*nat, intrel)"; +by (safe_tac intrel_cs); +by (rewtac refl_def); +by (fast_tac intrel_cs 1); +by (rewtac sym_def); +by (fast_tac (intrel_cs addSEs [sym]) 1); +by (rewtac trans_def); +by (fast_tac (intrel_cs addSEs [integ_trans_lemma]) 1); +val equiv_intrel = result(); + + +val integ_congs = mk_congs Integ.thy + ["znat", "zminus", "znegative", "zmagnitude", "op $+", "op $-", "op $*"]; + +val intrel_ss0 = arith_ss addcongs integ_congs; + +val intrel_ss = + intrel_ss0 addrews [equiv_intrel RS eq_equiv_class_iff, intrel_iff]; + +(*More than twice as fast as simplifying with intrel_ss*) +fun INTEG_SIMP_TAC ths = + let val ss = intrel_ss0 addrews ths + in fn i => + EVERY [ASM_SIMP_TAC ss i, + rtac (intrelI RS (equiv_intrel RS equiv_class_eq)) i, + typechk_tac (ZF_typechecks@nat_typechecks@arith_typechecks), + ASM_SIMP_TAC ss i] + end; + + +(** znat: the injection from nat to integ **) + +val prems = goalw Integ.thy [integ_def,quotient_def,znat_def] + "m : nat ==> $#m : integ"; +by (fast_tac (ZF_cs addSIs (nat_0I::prems)) 1); +val znat_type = result(); + +val [major,nnat] = goalw Integ.thy [znat_def] + "[| $#m = $#n; n: nat |] ==> m=n"; +by (rtac (make_elim (major RS eq_equiv_class)) 1); +by (rtac equiv_intrel 1); +by (typechk_tac [nat_0I,nnat,SigmaI]); +by (safe_tac (intrel_cs addSEs [box_equals,add_0_right])); +val znat_inject = result(); + + +(**** zminus: unary negation on integ ****) + +goalw Integ.thy [congruent_def] + "congruent(intrel, split(%x y. intrel``{}))"; +by (safe_tac intrel_cs); +by (ALLGOALS (ASM_SIMP_TAC intrel_ss)); +by (etac (box_equals RS sym) 1); +by (REPEAT (ares_tac [add_commute] 1)); +val zminus_congruent = result(); + +(*Resolve th against the corresponding facts for zminus*) +val zminus_ize = RSLIST [equiv_intrel, zminus_congruent]; + +val [prem] = goalw Integ.thy [integ_def,zminus_def] + "z : integ ==> $~z : integ"; +by (typechk_tac [split_type, SigmaI, prem, zminus_ize UN_equiv_class_type, + quotientI]); +val zminus_type = result(); + +val major::prems = goalw Integ.thy [integ_def,zminus_def] + "[| $~z = $~w; z: integ; w: integ |] ==> z=w"; +by (rtac (major RS zminus_ize UN_equiv_class_inject) 1); +by (REPEAT (ares_tac prems 1)); +by (REPEAT (etac SigmaE 1)); +by (etac rev_mp 1); +by (ASM_SIMP_TAC ZF_ss 1); +by (fast_tac (intrel_cs addSIs [SigmaI, equiv_intrel] + addSEs [box_equals RS sym, add_commute, + make_elim eq_equiv_class]) 1); +val zminus_inject = result(); + +val prems = goalw Integ.thy [zminus_def] + "[| x: nat; y: nat |] ==> $~ (intrel``{}) = intrel `` {}"; +by (ASM_SIMP_TAC + (ZF_ss addrews (prems@[zminus_ize UN_equiv_class, SigmaI])) 1); +val zminus = result(); + +goalw Integ.thy [integ_def] "!!z. z : integ ==> $~ ($~ z) = z"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (ASM_SIMP_TAC (ZF_ss addrews [zminus] addcongs integ_congs) 1); +val zminus_zminus = result(); + +goalw Integ.thy [integ_def, znat_def] "$~ ($#0) = $#0"; +by (SIMP_TAC (arith_ss addrews [zminus]) 1); +val zminus_0 = result(); + + +(**** znegative: the test for negative integers ****) + +goalw Integ.thy [znegative_def, znat_def] + "~ znegative($# n)"; +by (safe_tac intrel_cs); +by (rtac (add_not_less_self RS notE) 1); +by (etac ssubst 3); +by (ASM_SIMP_TAC (arith_ss addrews [add_0_right]) 3); +by (REPEAT (assume_tac 1)); +val not_znegative_znat = result(); + +val [nnat] = goalw Integ.thy [znegative_def, znat_def] + "n: nat ==> znegative($~ $# succ(n))"; +by (SIMP_TAC (intrel_ss addrews [zminus,nnat]) 1); +by (REPEAT + (resolve_tac [refl, exI, conjI, naturals_are_ordinals RS Ord_0_mem_succ, + refl RS intrelI RS imageI, consI1, nnat, nat_0I, + nat_succI] 1)); +val znegative_zminus_znat = result(); + + +(**** zmagnitude: magnitide of an integer, as a natural number ****) + +goalw Integ.thy [congruent_def] + "congruent(intrel, split(%x y. (y#-x) #+ (x#-y)))"; +by (safe_tac intrel_cs); +by (ALLGOALS (ASM_SIMP_TAC intrel_ss)); +by (etac rev_mp 1); +by (res_inst_tac [("m","x1"),("n","y1")] diff_induct 1); +by (REPEAT (assume_tac 1)); +by (ASM_SIMP_TAC (arith_ss addrews [add_succ_right,succ_inject_iff]) 3); +by (ASM_SIMP_TAC + (arith_ss addrews [diff_add_inverse,diff_add_0,add_0_right]) 2); +by (ASM_SIMP_TAC (arith_ss addrews [add_0_right]) 1); +by (rtac impI 1); +by (etac subst 1); +by (res_inst_tac [("m1","x")] (add_commute RS ssubst) 1); +by (REPEAT (assume_tac 1)); +by (ASM_SIMP_TAC (arith_ss addrews [diff_add_inverse,diff_add_0]) 1); +val zmagnitude_congruent = result(); + +(*Resolve th against the corresponding facts for zmagnitude*) +val zmagnitude_ize = RSLIST [equiv_intrel, zmagnitude_congruent]; + +val [prem] = goalw Integ.thy [integ_def,zmagnitude_def] + "z : integ ==> zmagnitude(z) : nat"; +by (typechk_tac [split_type, prem, zmagnitude_ize UN_equiv_class_type, + add_type, diff_type]); +val zmagnitude_type = result(); + +val prems = goalw Integ.thy [zmagnitude_def] + "[| x: nat; y: nat |] ==> \ +\ zmagnitude (intrel``{}) = (y #- x) #+ (x #- y)"; +by (ASM_SIMP_TAC + (ZF_ss addrews (prems@[zmagnitude_ize UN_equiv_class, SigmaI])) 1); +val zmagnitude = result(); + +val [nnat] = goalw Integ.thy [znat_def] + "n: nat ==> zmagnitude($# n) = n"; +by (SIMP_TAC (intrel_ss addrews [zmagnitude,nnat]) 1); +val zmagnitude_znat = result(); + +val [nnat] = goalw Integ.thy [znat_def] + "n: nat ==> zmagnitude($~ $# n) = n"; +by (SIMP_TAC (intrel_ss addrews [zmagnitude,zminus,nnat,add_0_right]) 1); +val zmagnitude_zminus_znat = result(); + + +(**** zadd: addition on integ ****) + +(** Congruence property for addition **) + +goalw Integ.thy [congruent2_def] + "congruent2(intrel, %p1 p2. \ +\ split(%x1 y1. split(%x2 y2. intrel `` {}, p2), p1))"; +(*Proof via congruent2_commuteI seems longer*) +by (safe_tac intrel_cs); +by (INTEG_SIMP_TAC [add_assoc] 1); +(*The rest should be trivial, but rearranging terms is hard*) +by (res_inst_tac [("m1","x1a")] (add_assoc_swap RS ssubst) 1); +by (res_inst_tac [("m1","x2a")] (add_assoc_swap RS ssubst) 3); +by (typechk_tac [add_type]); +by (ASM_SIMP_TAC (arith_ss addrews [add_assoc RS sym]) 1); +val zadd_congruent2 = result(); + +(*Resolve th against the corresponding facts for zadd*) +val zadd_ize = RSLIST [equiv_intrel, zadd_congruent2]; + +val prems = goalw Integ.thy [integ_def,zadd_def] + "[| z: integ; w: integ |] ==> z $+ w : integ"; +by (REPEAT (ares_tac (prems@[zadd_ize UN_equiv_class_type2, + split_type, add_type, quotientI, SigmaI]) 1)); +val zadd_type = result(); + +val prems = goalw Integ.thy [zadd_def] + "[| x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ +\ (intrel``{}) $+ (intrel``{}) = intrel `` {}"; +by (ASM_SIMP_TAC (ZF_ss addrews + (prems @ [zadd_ize UN_equiv_class2, SigmaI])) 1); +val zadd = result(); + +goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $+ z = z"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (ASM_SIMP_TAC (arith_ss addrews [zadd]) 1); +val zadd_0 = result(); + +goalw Integ.thy [integ_def] + "!!z w. [| z: integ; w: integ |] ==> $~ (z $+ w) = $~ z $+ $~ w"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (ASM_SIMP_TAC (arith_ss addrews [zminus,zadd] addcongs integ_congs) 1); +val zminus_zadd_distrib = result(); + +goalw Integ.thy [integ_def] + "!!z w. [| z: integ; w: integ |] ==> z $+ w = w $+ z"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (INTEG_SIMP_TAC [zadd] 1); +by (REPEAT (ares_tac [add_commute,add_cong] 1)); +val zadd_commute = result(); + +goalw Integ.thy [integ_def] + "!!z1 z2 z3. [| z1: integ; z2: integ; z3: integ |] ==> \ +\ (z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +(*rewriting is much faster without intrel_iff, etc.*) +by (ASM_SIMP_TAC (intrel_ss0 addrews [zadd,add_assoc]) 1); +val zadd_assoc = result(); + +val prems = goalw Integ.thy [znat_def] + "[| m: nat; n: nat |] ==> $# (m #+ n) = ($#m) $+ ($#n)"; +by (ASM_SIMP_TAC (arith_ss addrews (zadd::prems)) 1); +val znat_add = result(); + +goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> z $+ ($~ z) = $#0"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (ASM_SIMP_TAC (intrel_ss addrews [zminus,zadd,add_0_right]) 1); +by (REPEAT (ares_tac [add_commute] 1)); +val zadd_zminus_inverse = result(); + +val prems = goal Integ.thy + "z : integ ==> ($~ z) $+ z = $#0"; +by (rtac (zadd_commute RS trans) 1); +by (REPEAT (resolve_tac (prems@[zminus_type, zadd_zminus_inverse]) 1)); +val zadd_zminus_inverse2 = result(); + +val prems = goal Integ.thy "z:integ ==> z $+ $#0 = z"; +by (rtac (zadd_commute RS trans) 1); +by (REPEAT (resolve_tac (prems@[znat_type,nat_0I,zadd_0]) 1)); +val zadd_0_right = result(); + + +(*Need properties of $- ??? Or use $- just as an abbreviation? + [| m: nat; n: nat; m>=n |] ==> $# (m #- n) = ($#m) $- ($#n) +*) + +(**** zmult: multiplication on integ ****) + +(** Congruence property for multiplication **) + +val prems = goalw Integ.thy [znat_def] + "[| k: nat; l: nat; m: nat; n: nat |] ==> \ +\ (k #+ l) #+ (m #+ n) = (k #+ m) #+ (n #+ l)"; +val add_commute' = read_instantiate [("m","l")] add_commute; +by (SIMP_TAC (arith_ss addrews ([add_commute',add_assoc]@prems)) 1); +val zmult_congruent_lemma = result(); + +goal Integ.thy + "congruent2(intrel, %p1 p2. \ +\ split(%x1 y1. split(%x2 y2. \ +\ intrel``{}, p2), p1))"; +by (rtac (equiv_intrel RS congruent2_commuteI) 1); +by (safe_tac intrel_cs); +by (ALLGOALS (INTEG_SIMP_TAC [])); +(*Proof that zmult is congruent in one argument*) +by (rtac (zmult_congruent_lemma RS trans) 2); +by (rtac (zmult_congruent_lemma RS trans RS sym) 6); +by (typechk_tac [mult_type]); +by (ASM_SIMP_TAC (arith_ss addrews [add_mult_distrib_left RS sym]) 2); +(*Proof that zmult is commutative on representatives*) +by (rtac add_cong 1); +by (rtac (add_commute RS trans) 2); +by (REPEAT (ares_tac [mult_commute,add_type,mult_type,add_cong] 1)); +val zmult_congruent2 = result(); + +(*Resolve th against the corresponding facts for zmult*) +val zmult_ize = RSLIST [equiv_intrel, zmult_congruent2]; + +val prems = goalw Integ.thy [integ_def,zmult_def] + "[| z: integ; w: integ |] ==> z $* w : integ"; +by (REPEAT (ares_tac (prems@[zmult_ize UN_equiv_class_type2, + split_type, add_type, mult_type, + quotientI, SigmaI]) 1)); +val zmult_type = result(); + + +val prems = goalw Integ.thy [zmult_def] + "[| x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ +\ (intrel``{}) $* (intrel``{}) = \ +\ intrel `` {}"; +by (ASM_SIMP_TAC (ZF_ss addrews + (prems @ [zmult_ize UN_equiv_class2, SigmaI])) 1); +val zmult = result(); + +goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $* z = $#0"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (ASM_SIMP_TAC (arith_ss addrews [zmult]) 1); +val zmult_0 = result(); + +goalw Integ.thy [integ_def,znat_def,one_def] + "!!z. z : integ ==> $#1 $* z = z"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (ASM_SIMP_TAC (arith_ss addrews [zmult,add_0_right]) 1); +val zmult_1 = result(); + +goalw Integ.thy [integ_def] + "!!z w. [| z: integ; w: integ |] ==> ($~ z) $* w = $~ (z $* w)"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (INTEG_SIMP_TAC [zminus,zmult] 1); +by (REPEAT (ares_tac [mult_type,add_commute,add_cong] 1)); +val zmult_zminus = result(); + +goalw Integ.thy [integ_def] + "!!z w. [| z: integ; w: integ |] ==> ($~ z) $* ($~ w) = (z $* w)"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (INTEG_SIMP_TAC [zminus,zmult] 1); +by (REPEAT (ares_tac [mult_type,add_commute,add_cong] 1)); +val zmult_zminus_zminus = result(); + +goalw Integ.thy [integ_def] + "!!z w. [| z: integ; w: integ |] ==> z $* w = w $* z"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (INTEG_SIMP_TAC [zmult] 1); +by (res_inst_tac [("m1","xc #* y")] (add_commute RS ssubst) 1); +by (REPEAT (ares_tac [mult_type,mult_commute,add_cong] 1)); +val zmult_commute = result(); + +goalw Integ.thy [integ_def] + "!!z1 z2 z3. [| z1: integ; z2: integ; z3: integ |] ==> \ +\ (z1 $* z2) $* z3 = z1 $* (z2 $* z3)"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (INTEG_SIMP_TAC [zmult, add_mult_distrib_left, + add_mult_distrib, add_assoc, mult_assoc] 1); +(*takes 54 seconds due to wasteful type-checking*) +by (REPEAT (ares_tac [add_type, mult_type, add_commute, add_kill, + add_assoc_swap_kill, add_assoc_swap_kill RS sym] 1)); +val zmult_assoc = result(); + +goalw Integ.thy [integ_def] + "!!z1 z2 z3. [| z1: integ; z2: integ; w: integ |] ==> \ +\ (z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (INTEG_SIMP_TAC [zadd, zmult, add_mult_distrib, add_assoc] 1); +(*takes 30 seconds due to wasteful type-checking*) +by (REPEAT (ares_tac [add_type, mult_type, refl, add_commute, add_kill, + add_assoc_swap_kill, add_assoc_swap_kill RS sym] 1)); +val zadd_zmult_distrib = result(); + +val integ_typechecks = + [znat_type, zminus_type, zmagnitude_type, zadd_type, zmult_type]; + +val integ_ss = + arith_ss addcongs integ_congs + addrews ([zminus_zminus,zmagnitude_znat,zmagnitude_zminus_znat, + zadd_0] @ integ_typechecks); diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Integ.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Integ.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,54 @@ +(* Title: ZF/ex/integ.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +The integers as equivalence classes over nat*nat. +*) + +Integ = Equiv + Arith + +consts + intrel,integ:: "i" + znat :: "i=>i" ("$# _" [80] 80) + zminus :: "i=>i" ("$~ _" [80] 80) + znegative :: "i=>o" + zmagnitude :: "i=>i" + "$*" :: "[i,i]=>i" (infixl 70) + "$'/" :: "[i,i]=>i" (infixl 70) + "$'/'/" :: "[i,i]=>i" (infixl 70) + "$+" :: "[i,i]=>i" (infixl 65) + "$-" :: "[i,i]=>i" (infixl 65) + "$<" :: "[i,i]=>o" (infixl 50) + +rules + + intrel_def + "intrel == {p:(nat*nat)*(nat*nat). \ +\ EX x1 y1 x2 y2. p=<,> & x1#+y2 = x2#+y1}" + + integ_def "integ == (nat*nat)/intrel" + + znat_def "$# m == intrel `` {}" + + zminus_def "$~ Z == UN p:Z. split(%x y. intrel``{}, p)" + + znegative_def + "znegative(Z) == EX x y. x:y & y:nat & :Z" + + zmagnitude_def + "zmagnitude(Z) == UN p:Z. split(%x y. (y#-x) #+ (x#-y), p)" + + zadd_def + "Z1 $+ Z2 == \ +\ UN p1:Z1. UN p2:Z2. split(%x1 y1. split(%x2 y2. \ +\ intrel``{}, p2), p1)" + + zdiff_def "Z1 $- Z2 == Z1 $+ zminus(Z2)" + zless_def "Z1 $< Z2 == znegative(Z1 $- Z2)" + + zmult_def + "Z1 $* Z2 == \ +\ UN p1:Z1. UN p2:Z2. split(%x1 y1. split(%x2 y2. \ +\ intrel``{}, p2), p1)" + + end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/LList.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/LList.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,151 @@ +(* Title: ZF/ex/llist.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Co-Datatype definition of Lazy Lists + +Needs a "take-lemma" to prove llist_subset_quniv and to justify co-induction +for proving equality +*) + +structure LList = Co_Datatype_Fun + (val thy = QUniv.thy; + val thy = QUniv.thy; + val rec_specs = + [("llist", "quniv(A)", + [(["LNil"], "i"), + (["LCons"], "[i,i]=>i")])]; + val rec_styp = "i=>i"; + val ext = None + val sintrs = + ["LNil : llist(A)", + "[| a: A; l: llist(A) |] ==> LCons(a,l) : llist(A)"]; + val monos = []; + val type_intrs = co_data_typechecks + val type_elims = []); + +val [LNilI, LConsI] = LList.intrs; + +(*An elimination rule, for type-checking*) +val LConsE = LList.mk_cases LList.con_defs "LCons(a,l) : llist(A)"; + +(*Proving freeness results*) +val LCons_iff = LList.mk_free "LCons(a,l)=LCons(a',l') <-> a=a' & l=l'"; +val LNil_LCons_iff = LList.mk_free "~ LNil=LCons(a,l)"; + +(*** Lemmas to justify using "llist" in other recursive type definitions ***) + +goalw LList.thy LList.defs "!!A B. A<=B ==> llist(A) <= llist(B)"; +by (rtac gfp_mono 1); +by (REPEAT (rtac LList.bnd_mono 1)); +by (REPEAT (ares_tac (quniv_mono::basic_monos) 1)); +val llist_mono = result(); + +(** Closure of quniv(A) under llist -- why so complex? Its a gfp... **) + +val in_quniv_rls = + [Transset_quniv, QPair_Int_quniv_in_quniv, Int_Vfrom_0_in_quniv, + zero_Int_in_quniv, one_Int_in_quniv, + QPair_Int_Vfrom_succ_in_quniv, QPair_Int_Vfrom_in_quniv]; + +val quniv_cs = ZF_cs addSIs in_quniv_rls + addIs (Int_quniv_in_quniv::co_data_typechecks); + +(*Keep unfolding the lazy list until the induction hypothesis applies*) +goal LList.thy + "!!i. i : nat ==> \ +\ ALL l: llist(quniv(A)). l Int Vfrom(quniv(A), i) : quniv(A)"; +be complete_induct 1; +br ballI 1; +be LList.elim 1; +bws ([QInl_def,QInr_def]@LList.con_defs); +by (fast_tac quniv_cs 1); +by (etac natE 1 THEN REPEAT_FIRST hyp_subst_tac); +by (fast_tac quniv_cs 1); +by (fast_tac quniv_cs 1); +val llist_quniv_lemma = result(); + +goal LList.thy "llist(quniv(A)) <= quniv(A)"; +br subsetI 1; +br quniv_Int_Vfrom 1; +be (LList.dom_subset RS subsetD) 1; +by (REPEAT (ares_tac [llist_quniv_lemma RS bspec] 1)); +val llist_quniv = result(); + +val llist_subset_quniv = standard + (llist_mono RS (llist_quniv RSN (2,subset_trans))); + +(*** Equality for llist(A) as a greatest fixed point ***) + +structure LList_Eq = Co_Inductive_Fun + (val thy = LList.thy addconsts [(["lleq"],"i=>i")]; + val rec_doms = [("lleq","llist(A) <*> llist(A)")]; + val sintrs = + [" : lleq(A)", + "[| a:A; : lleq(A) |] ==> : lleq(A)"]; + val monos = []; + val con_defs = []; + val type_intrs = LList.intrs@[QSigmaI]; + val type_elims = [QSigmaE2]); + +(** Alternatives for above: + val con_defs = LList.con_defs + val type_intrs = co_data_typechecks + val type_elims = [quniv_QPair_E] +**) + +val lleq_cs = subset_cs + addSIs [succI1, Int_Vset_0_subset, + QPair_Int_Vset_succ_subset_trans, + QPair_Int_Vset_subset_trans]; + +(*Keep unfolding the lazy list until the induction hypothesis applies*) +goal LList_Eq.thy + "!!i. Ord(i) ==> ALL l l'. : lleq(A) --> l Int Vset(i) <= l'"; +be trans_induct 1; +by (safe_tac subset_cs); +be LList_Eq.elim 1; +by (safe_tac (subset_cs addSEs [QPair_inject])); +bws LList.con_defs; +by (etac Ord_cases 1 THEN REPEAT_FIRST hyp_subst_tac); +(*0 case*) +by (fast_tac lleq_cs 1); +(*succ(j) case*) +bw QInr_def; +by (fast_tac lleq_cs 1); +(*Limit(i) case*) +be (Limit_Vfrom_eq RS ssubst) 1; +br (Int_UN_distrib RS ssubst) 1; +by (fast_tac lleq_cs 1); +val lleq_Int_Vset_subset_lemma = result(); + +val lleq_Int_Vset_subset = standard + (lleq_Int_Vset_subset_lemma RS spec RS spec RS mp); + + +(*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*) +val [prem] = goal LList_Eq.thy " : lleq(A) ==> : lleq(A)"; +br (prem RS qconverseI RS LList_Eq.co_induct) 1; +br (LList_Eq.dom_subset RS qconverse_type) 1; +by (safe_tac qconverse_cs); +be LList_Eq.elim 1; +by (ALLGOALS (fast_tac qconverse_cs)); +val lleq_symmetric = result(); + +goal LList_Eq.thy "!!l l'. : lleq(A) ==> l=l'"; +br equalityI 1; +by (REPEAT (ares_tac [lleq_Int_Vset_subset RS Int_Vset_subset] 1 + ORELSE etac lleq_symmetric 1)); +val lleq_implies_equal = result(); + +val [eqprem,lprem] = goal LList_Eq.thy + "[| l=l'; l: llist(A) |] ==> : lleq(A)"; +by (res_inst_tac [("X", "{. l: llist(A)}")] LList_Eq.co_induct 1); +br (lprem RS RepFunI RS (eqprem RS subst)) 1; +by (safe_tac qpair_cs); +be LList.elim 1; +by (ALLGOALS (fast_tac qpair_cs)); +val equal_llist_implies_leq = result(); + + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/LListFn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/LListFn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,37 @@ +(* Title: ZF/ex/llist-fn.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Functions for Lazy Lists in Zermelo-Fraenkel Set Theory +*) + +open LListFn; + +(*** lconst -- defined directly using lfp, but equivalent to a LList_corec ***) + +goalw LListFn.thy LList.con_defs "bnd_mono(univ(a), %l. LCons(a,l))"; +by (rtac bnd_monoI 1); +by (REPEAT (ares_tac [subset_refl, QInr_mono, QPair_mono] 2)); +by (REPEAT (ares_tac [subset_refl, A_subset_univ, + QInr_subset_univ, QPair_subset_univ] 1)); +val lconst_fun_bnd_mono = result(); + +(* lconst(a) = LCons(a,lconst(a)) *) +val lconst = standard + ([lconst_def, lconst_fun_bnd_mono] MRS def_lfp_Tarski); + +val lconst_subset = lconst_def RS def_lfp_subset; + +val member_subset_Union_eclose = standard (arg_into_eclose RS Union_upper); + +goal LListFn.thy "!!a A. a : A ==> lconst(a) : quniv(A)"; +by (rtac (lconst_subset RS subset_trans RS qunivI) 1); +by (etac (arg_into_eclose RS eclose_subset RS univ_mono) 1); +val lconst_in_quniv = result(); + +goal LListFn.thy "!!a A. a:A ==> lconst(a): llist(A)"; +by (rtac (singletonI RS LList.co_induct) 1); +by (fast_tac (ZF_cs addSIs [lconst_in_quniv]) 1); +by (fast_tac (ZF_cs addSIs [lconst]) 1); +val lconst_type = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/LListFn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/LListFn.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,16 @@ +(* Title: ZF/ex/llist-fn.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Functions for Lazy Lists in Zermelo-Fraenkel Set Theory +*) + +LListFn = LList + +consts + lconst :: "i => i" + +rules + lconst_def "lconst(a) == lfp(univ(a), %l. LCons(a,l))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/ListN.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/ListN.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,47 @@ +(* Title: ZF/ex/listn + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Inductive definition of lists of n elements + +See Ch. Paulin-Mohring, Inductive Definitions in the System Coq. +Research Report 92-49, LIP, ENS Lyon. Dec 1992. +*) + +structure ListN = Inductive_Fun + (val thy = ListFn.thy addconsts [(["listn"],"i=>i")]; + val rec_doms = [("listn", "nat*list(A)")]; + val sintrs = + ["<0,Nil> : listn(A)", + "[| a: A; : listn(A) |] ==> : listn(A)"]; + val monos = []; + val con_defs = []; + val type_intrs = nat_typechecks@List.intrs@[SigmaI] + val type_elims = [SigmaE2]); + +(*Could be generated automatically; requires use of fsplitD*) +val listn_induct = standard + (fsplitI RSN + (2, fsplitI RSN + (3, (read_instantiate [("P", "fsplit(R)")] ListN.induct) RS fsplitD))); + +val [major] = goal ListN.thy "l:list(A) ==> : listn(A)"; +by (rtac (major RS List.induct) 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +by (REPEAT (ares_tac ListN.intrs 1)); +val list_into_listn = result(); + +goal ListN.thy " : listn(A) <-> l:list(A) & length(l)=n"; +by (rtac iffI 1); +by (etac listn_induct 1); +by (dtac fsplitD 2); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [Pair_iff]))); +by (fast_tac (ZF_cs addSIs [list_into_listn]) 1); +val listn_iff = result(); + +goal ListN.thy "listn(A)``{n} = {l:list(A). length(l)=n}"; +by (rtac equality_iffI 1); +by (SIMP_TAC (list_ss addrews [listn_iff,separation,image_singleton_iff]) 1); +val listn_image_eq = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/ParContract.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/ParContract.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,144 @@ +(* Title: ZF/ex/parcontract.ML + ID: $Id$ + Author: Lawrence C Paulson + Copyright 1993 University of Cambridge + +Parallel contraction + +HOL system proofs may be found in +/usr/groups/theory/hvg-aftp/contrib/rule-induction/cl.ml +*) + +structure ParContract = Inductive_Fun + (val thy = Contract.thy; + val rec_doms = [("parcontract","comb*comb")]; + val sintrs = + ["[| p:comb |] ==> p =1=> p", + "[| p:comb; q:comb |] ==> K#p#q =1=> p", + "[| p:comb; q:comb; r:comb |] ==> S#p#q#r =1=> (p#r)#(q#r)", + "[| p=1=>q; r=1=>s |] ==> p#r =1=> q#s"]; + val monos = []; + val con_defs = []; + val type_intrs = Comb.intrs@[SigmaI]; + val type_elims = [SigmaE2]); + +val [parcontract_refl,K_parcontract,S_parcontract,Ap_parcontract] = ParContract.intrs; + +val parcontract_induct = standard + (ParContract.mutual_induct RS spec RS spec RSN (2,rev_mp)); + +(*For type checking: replaces a=1=>b by a,b:comb *) +val parcontract_combE2 = ParContract.dom_subset RS subsetD RS SigmaE2; +val parcontract_combD1 = ParContract.dom_subset RS subsetD RS SigmaD1; +val parcontract_combD2 = ParContract.dom_subset RS subsetD RS SigmaD2; + +goal ParContract.thy "field(parcontract) = comb"; +by (fast_tac (ZF_cs addIs [equalityI,K_parcontract] + addSEs [parcontract_combE2]) 1); +val field_parcontract_eq = result(); + +val parcontract_caseE = standard + (ParContract.unfold RS equalityD1 RS subsetD RS CollectE); + +(*Derive a case for each combinator constructor*) +val K_parcontract_case = ParContract.mk_cases Comb.con_defs "K =1=> r"; +val S_parcontract_case = ParContract.mk_cases Comb.con_defs "S =1=> r"; +val Ap_parcontract_case = ParContract.mk_cases Comb.con_defs "p#q =1=> r"; + +val parcontract_cs = + ZF_cs addSIs Comb.intrs + addIs ParContract.intrs + addSEs [Ap_E, K_parcontract_case, S_parcontract_case, + Ap_parcontract_case] + addSEs [parcontract_combD1, parcontract_combD2] (*type checking*) + addSEs Comb.free_SEs; + +(*** Basic properties of parallel contraction ***) + +goal ParContract.thy "!!p r. K#p =1=> r ==> (EX p'. r = K#p' & p =1=> p')"; +by (fast_tac parcontract_cs 1); +val K1_parcontractD = result(); + +goal ParContract.thy "!!p r. S#p =1=> r ==> (EX p'. r = S#p' & p =1=> p')"; +by (fast_tac parcontract_cs 1); +val S1_parcontractD = result(); + +goal ParContract.thy + "!!p q r. S#p#q =1=> r ==> (EX p' q'. r = S#p'#q' & p =1=> p' & q =1=> q')"; +by (fast_tac (parcontract_cs addSDs [S1_parcontractD]) 1); +val S2_parcontractD = result(); + +(*Church-Rosser property for parallel contraction*) +goalw ParContract.thy [diamond_def] "diamond(parcontract)"; +by (rtac (impI RS allI RS allI) 1); +by (etac parcontract_induct 1); +by (ALLGOALS + (fast_tac (parcontract_cs addSDs [K1_parcontractD,S2_parcontractD]))); +val diamond_parcontract = result(); + +(*** Transitive closure preserves the Church-Rosser property ***) + +goalw ParContract.thy [diamond_def] + "!!x y r. [| diamond(r); :r^+ |] ==> \ +\ ALL y'. :r --> (EX z. : r^+ & : r)"; +by (etac trancl_induct 1); +by (fast_tac (ZF_cs addIs [r_into_trancl]) 1); +by (slow_best_tac (ZF_cs addSDs [spec RS mp] + addIs [r_into_trancl, trans_trancl RS transD]) 1); +val diamond_trancl_lemma = result(); + +val diamond_lemmaE = diamond_trancl_lemma RS spec RS mp RS exE; + +val [major] = goal ParContract.thy "diamond(r) ==> diamond(r^+)"; +bw diamond_def; (*unfold only in goal, not in premise!*) +by (rtac (impI RS allI RS allI) 1); +by (etac trancl_induct 1); +by (ALLGOALS + (slow_best_tac (ZF_cs addIs [r_into_trancl, trans_trancl RS transD] + addEs [major RS diamond_lemmaE]))); +val diamond_trancl = result(); + + +(*** Equivalence of p--->q and p===>q ***) + +goal ParContract.thy "!!p q. p-1->q ==> p=1=>q"; +by (etac contract_induct 1); +by (ALLGOALS (fast_tac (parcontract_cs))); +val contract_imp_parcontract = result(); + +goal ParContract.thy "!!p q. p--->q ==> p===>q"; +by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1); +by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1); +by (etac rtrancl_induct 1); +by (fast_tac (parcontract_cs addIs [r_into_trancl]) 1); +by (fast_tac (ZF_cs addIs [contract_imp_parcontract, + r_into_trancl, trans_trancl RS transD]) 1); +val reduce_imp_parreduce = result(); + + +goal ParContract.thy "!!p q. p=1=>q ==> p--->q"; +by (etac parcontract_induct 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +by (rtac (trans_rtrancl RS transD) 1); +by (ALLGOALS + (fast_tac + (contract_cs addIs [Ap_reduce1, Ap_reduce2] + addSEs [parcontract_combD1,parcontract_combD2]))); +val parcontract_imp_reduce = result(); + +goal ParContract.thy "!!p q. p===>q ==> p--->q"; +by (forward_tac [trancl_type RS subsetD RS SigmaD1] 1); +by (dtac (field_parcontract_eq RS equalityD1 RS subsetD) 1); +by (etac trancl_induct 1); +by (etac parcontract_imp_reduce 1); +by (etac (trans_rtrancl RS transD) 1); +by (etac parcontract_imp_reduce 1); +val parreduce_imp_reduce = result(); + +goal ParContract.thy "p===>q <-> p--->q"; +by (REPEAT (ares_tac [iffI, parreduce_imp_reduce, reduce_imp_parreduce] 1)); +val parreduce_iff_reduce = result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Prop.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Prop.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,47 @@ +(* Title: ZF/ex/prop.ML + ID: $Id$ + Author: Lawrence C Paulson + Copyright 1993 University of Cambridge + +Datatype definition of propositional logic formulae and inductive definition +of the propositional tautologies. +*) + +(*Example of a datatype with mixfix syntax for some constructors*) +structure Prop = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("prop", "univ(0)", + [(["Fls"], "i"), + (["Var"], "i=>i"), + (["op =>"], "[i,i]=>i")])]; + val rec_styp = "i"; + val ext = Some (NewSext { + mixfix = + [Mixfix("#_", "i => i", "Var", [100], 100), + Infixr("=>", "[i,i] => i", 90)], + xrules = [], + parse_ast_translation = [], + parse_preproc = None, + parse_postproc = None, + parse_translation = [], + print_translation = [], + print_preproc = None, + print_postproc = None, + print_ast_translation = []}); + val sintrs = + ["Fls : prop", + "n: nat ==> #n : prop", + "[| p: prop; q: prop |] ==> p=>q : prop"]; + val monos = []; + val type_intrs = data_typechecks; + val type_elims = []); + +val [FlsI,VarI,ImpI] = Prop.intrs; + + +(** Type-checking rules **) + +val ImpE = Prop.mk_cases Prop.con_defs "p=>q : prop"; + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/PropLog.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/PropLog.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,333 @@ +(* Title: ZF/ex/prop-log.ML + ID: $Id$ + Author: Tobias Nipkow & Lawrence C Paulson + Copyright 1992 University of Cambridge + +For ex/prop-log.thy. Inductive definition of propositional logic. +Soundness and completeness w.r.t. truth-tables. + +Prove: If H|=p then G|=p where G:Fin(H) +*) + +open PropLog; + +(*** prop_rec -- by Vset recursion ***) + +val prop_congs = mk_typed_congs Prop.thy + [("b", "[i,i,i]=>i"), ("d", "[i,i,i,i]=>i")]; + +(** conversion rules **) + +goal PropLog.thy "prop_rec(Fls,b,c,d) = b"; +by (rtac (prop_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Prop.con_defs); +by (SIMP_TAC rank_ss 1); +val prop_rec_Fls = result(); + +goal PropLog.thy "prop_rec(#v,b,c,d) = c(v)"; +by (rtac (prop_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Prop.con_defs); +by (SIMP_TAC (rank_ss addcongs prop_congs) 1); +val prop_rec_Var = result(); + +goal PropLog.thy "prop_rec(p=>q,b,c,d) = \ +\ d(p, q, prop_rec(p,b,c,d), prop_rec(q,b,c,d))"; +by (rtac (prop_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Prop.con_defs); +by (SIMP_TAC (rank_ss addcongs prop_congs) 1); +val prop_rec_Imp = result(); + +val prop_rec_ss = + arith_ss addrews [prop_rec_Fls, prop_rec_Var, prop_rec_Imp]; + +(*** Semantics of propositional logic ***) + +(** The function is_true **) + +goalw PropLog.thy [is_true_def] "is_true(Fls,t) <-> False"; +by (SIMP_TAC (prop_rec_ss addrews [one_not_0 RS not_sym]) 1); +val is_true_Fls = result(); + +goalw PropLog.thy [is_true_def] "is_true(#v,t) <-> v:t"; +by (SIMP_TAC (prop_rec_ss addrews [one_not_0 RS not_sym] + addsplits [expand_if]) 1); +val is_true_Var = result(); + +goalw PropLog.thy [is_true_def] + "is_true(p=>q,t) <-> (is_true(p,t)-->is_true(q,t))"; +by (SIMP_TAC (prop_rec_ss addsplits [expand_if]) 1); +val is_true_Imp = result(); + +(** The function hyps **) + +goalw PropLog.thy [hyps_def] "hyps(Fls,t) = 0"; +by (SIMP_TAC prop_rec_ss 1); +val hyps_Fls = result(); + +goalw PropLog.thy [hyps_def] "hyps(#v,t) = {if(v:t, #v, #v=>Fls)}"; +by (SIMP_TAC prop_rec_ss 1); +val hyps_Var = result(); + +goalw PropLog.thy [hyps_def] "hyps(p=>q,t) = hyps(p,t) Un hyps(q,t)"; +by (SIMP_TAC prop_rec_ss 1); +val hyps_Imp = result(); + +val prop_ss = prop_rec_ss + addcongs Prop.congs + addcongs (mk_congs PropLog.thy ["Fin", "thms", "op |=","is_true","hyps"]) + addrews Prop.intrs + addrews [is_true_Fls, is_true_Var, is_true_Imp, + hyps_Fls, hyps_Var, hyps_Imp]; + +(*** Proof theory of propositional logic ***) + +structure PropThms = Inductive_Fun + (val thy = PropLog.thy; + val rec_doms = [("thms","prop")]; + val sintrs = + ["[| p:H; p:prop |] ==> H |- p", + "[| p:prop; q:prop |] ==> H |- p=>q=>p", + "[| p:prop; q:prop; r:prop |] ==> H |- (p=>q=>r) => (p=>q) => p=>r", + "p:prop ==> H |- ((p=>Fls) => Fls) => p", + "[| H |- p=>q; H |- p; p:prop; q:prop |] ==> H |- q"]; + val monos = []; + val con_defs = []; + val type_intrs = Prop.intrs; + val type_elims = []); + +goalw PropThms.thy PropThms.defs "!!G H. G<=H ==> thms(G) <= thms(H)"; +by (rtac lfp_mono 1); +by (REPEAT (rtac PropThms.bnd_mono 1)); +by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); +val thms_mono = result(); + +val thms_in_pl = PropThms.dom_subset RS subsetD; + +val [thms_H, thms_K, thms_S, thms_DN, weak_thms_MP] = PropThms.intrs; + +(*Modus Ponens rule -- this stronger version avoids typecheck*) +goal PropThms.thy "!!p q H. [| H |- p=>q; H |- p |] ==> H |- q"; +by (rtac weak_thms_MP 1); +by (REPEAT (eresolve_tac [asm_rl, thms_in_pl, thms_in_pl RS ImpE] 1)); +val thms_MP = result(); + +(*Rule is called I for Identity Combinator, not for Introduction*) +goal PropThms.thy "!!p H. p:prop ==> H |- p=>p"; +by (rtac (thms_S RS thms_MP RS thms_MP) 1); +by (rtac thms_K 5); +by (rtac thms_K 4); +by (REPEAT (ares_tac [ImpI] 1)); +val thms_I = result(); + +(** Weakening, left and right **) + +(* [| G<=H; G|-p |] ==> H|-p Order of premises is convenient with RS*) +val weaken_left = standard (thms_mono RS subsetD); + +(* H |- p ==> cons(a,H) |- p *) +val weaken_left_cons = subset_consI RS weaken_left; + +val weaken_left_Un1 = Un_upper1 RS weaken_left; +val weaken_left_Un2 = Un_upper2 RS weaken_left; + +goal PropThms.thy "!!H p q. [| H |- q; p:prop |] ==> H |- p=>q"; +by (rtac (thms_K RS thms_MP) 1); +by (REPEAT (ares_tac [thms_in_pl] 1)); +val weaken_right = result(); + +(*The deduction theorem*) +goal PropThms.thy "!!p q H. [| cons(p,H) |- q; p:prop |] ==> H |- p=>q"; +by (etac PropThms.induct 1); +by (fast_tac (ZF_cs addIs [thms_I, thms_H RS weaken_right]) 1); +by (fast_tac (ZF_cs addIs [thms_K RS weaken_right]) 1); +by (fast_tac (ZF_cs addIs [thms_S RS weaken_right]) 1); +by (fast_tac (ZF_cs addIs [thms_DN RS weaken_right]) 1); +by (fast_tac (ZF_cs addIs [thms_S RS thms_MP RS thms_MP]) 1); +val deduction = result(); + + +(*The cut rule*) +goal PropThms.thy "!!H p q. [| H|-p; cons(p,H) |- q |] ==> H |- q"; +by (rtac (deduction RS thms_MP) 1); +by (REPEAT (ares_tac [thms_in_pl] 1)); +val cut = result(); + +goal PropThms.thy "!!H p. [| H |- Fls; p:prop |] ==> H |- p"; +by (rtac (thms_DN RS thms_MP) 1); +by (rtac weaken_right 2); +by (REPEAT (ares_tac (Prop.intrs@[consI1]) 1)); +val thms_FlsE = result(); + +(* [| H |- p=>Fls; H |- p; q: prop |] ==> H |- q *) +val thms_notE = standard (thms_MP RS thms_FlsE); + +(*Soundness of the rules wrt truth-table semantics*) +val [major] = goalw PropThms.thy [sat_def] "H |- p ==> H |= p"; +by (rtac (major RS PropThms.induct) 1); +by (fast_tac (ZF_cs addSDs [is_true_Imp RS iffD1 RS mp]) 5); +by (ALLGOALS (SIMP_TAC prop_ss)); +val soundness = result(); + +(*** Towards the completeness proof ***) + +val [premf,premq] = goal PropThms.thy + "[| H |- p=>Fls; q: prop |] ==> H |- p=>q"; +by (rtac (premf RS thms_in_pl RS ImpE) 1); +by (rtac deduction 1); +by (rtac (premf RS weaken_left_cons RS thms_notE) 1); +by (REPEAT (ares_tac [premq, consI1, thms_H] 1)); +val Fls_Imp = result(); + +val [premp,premq] = goal PropThms.thy + "[| H |- p; H |- q=>Fls |] ==> H |- (p=>q)=>Fls"; +by (cut_facts_tac ([premp,premq] RL [thms_in_pl]) 1); +by (etac ImpE 1); +by (rtac deduction 1); +by (rtac (premq RS weaken_left_cons RS thms_MP) 1); +by (rtac (consI1 RS thms_H RS thms_MP) 1); +by (rtac (premp RS weaken_left_cons) 2); +by (REPEAT (ares_tac Prop.intrs 1)); +val Imp_Fls = result(); + +(*Typical example of strengthening the induction formula*) +val [major] = goal PropThms.thy + "p: prop ==> hyps(p,t) |- if(is_true(p,t), p, p=>Fls)"; +by (rtac (expand_if RS iffD2) 1); +by (rtac (major RS Prop.induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (prop_ss addrews [thms_I, thms_H]))); +by (fast_tac (ZF_cs addIs [weaken_left_Un1, weaken_left_Un2, + weaken_right, Imp_Fls] + addSEs [Fls_Imp]) 1); +val hyps_thms_if = result(); + +(*Key lemma for completeness; yields a set of assumptions satisfying p*) +val [premp,sat] = goalw PropThms.thy [sat_def] + "[| p: prop; 0 |= p |] ==> hyps(p,t) |- p"; +by (rtac (sat RS spec RS mp RS if_P RS subst) 1 THEN + rtac (premp RS hyps_thms_if) 2); +by (fast_tac ZF_cs 1); +val sat_thms_p = result(); + +(*For proving certain theorems in our new propositional logic*) +val thms_cs = + ZF_cs addSIs [FlsI, VarI, ImpI, deduction] + addIs [thms_in_pl, thms_H, thms_H RS thms_MP]; + +(*The excluded middle in the form of an elimination rule*) +val prems = goal PropThms.thy + "[| p: prop; q: prop |] ==> H |- (p=>q) => ((p=>Fls)=>q) => q"; +by (rtac (deduction RS deduction) 1); +by (rtac (thms_DN RS thms_MP) 1); +by (ALLGOALS (best_tac (thms_cs addSIs prems))); +val thms_excluded_middle = result(); + +(*Hard to prove directly because it requires cuts*) +val prems = goal PropThms.thy + "[| cons(p,H) |- q; cons(p=>Fls,H) |- q; p: prop |] ==> H |- q"; +by (rtac (thms_excluded_middle RS thms_MP RS thms_MP) 1); +by (REPEAT (resolve_tac (prems@Prop.intrs@[deduction,thms_in_pl]) 1)); +val thms_excluded_middle_rule = result(); + +(*** Completeness -- lemmas for reducing the set of assumptions ***) + +(*For the case hyps(p,t)-cons(#v,Y) |- p; + we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *) +val [major] = goal PropThms.thy + "p: prop ==> hyps(p, t-{v}) <= cons(#v=>Fls, hyps(p,t)-{#v})"; +by (rtac (major RS Prop.induct) 1); +by (SIMP_TAC prop_ss 1); +by (ASM_SIMP_TAC (prop_ss addsplits [expand_if]) 1); +by (fast_tac (ZF_cs addSEs Prop.free_SEs) 1); +by (ASM_SIMP_TAC prop_ss 1); +by (fast_tac ZF_cs 1); +val hyps_Diff = result(); + +(*For the case hyps(p,t)-cons(#v => Fls,Y) |- p; + we also have hyps(p,t)-{#v=>Fls} <= hyps(p, cons(v,t)) *) +val [major] = goal PropThms.thy + "p: prop ==> hyps(p, cons(v,t)) <= cons(#v, hyps(p,t)-{#v=>Fls})"; +by (rtac (major RS Prop.induct) 1); +by (SIMP_TAC prop_ss 1); +by (ASM_SIMP_TAC (prop_ss addsplits [expand_if]) 1); +by (fast_tac (ZF_cs addSEs Prop.free_SEs) 1); +by (ASM_SIMP_TAC prop_ss 1); +by (fast_tac ZF_cs 1); +val hyps_cons = result(); + +(** Two lemmas for use with weaken_left **) + +goal ZF.thy "B-C <= cons(a, B-cons(a,C))"; +by (fast_tac ZF_cs 1); +val cons_Diff_same = result(); + +goal ZF.thy "cons(a, B-{c}) - D <= cons(a, B-cons(c,D))"; +by (fast_tac ZF_cs 1); +val cons_Diff_subset2 = result(); + +(*The set hyps(p,t) is finite, and elements have the form #v or #v=>Fls; + could probably prove the stronger hyps(p,t) : Fin(hyps(p,0) Un hyps(p,nat))*) +val [major] = goal PropThms.thy + "p: prop ==> hyps(p,t) : Fin(UN v:nat. {#v, #v=>Fls})"; +by (rtac (major RS Prop.induct) 1); +by (ASM_SIMP_TAC (prop_ss addrews [Fin_0I, Fin_consI, UN_I] + addsplits [expand_if]) 2); +by (ALLGOALS (ASM_SIMP_TAC (prop_ss addrews [Un_0, Fin_0I, Fin_UnI]))); +val hyps_finite = result(); + +val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left; + +(*Induction on the finite set of assumptions hyps(p,t0). + We may repeatedly subtract assumptions until none are left!*) +val [premp,sat] = goal PropThms.thy + "[| p: prop; 0 |= p |] ==> ALL t. hyps(p,t) - hyps(p,t0) |- p"; +by (rtac (premp RS hyps_finite RS Fin_induct) 1); +by (SIMP_TAC (prop_ss addrews [premp, sat, sat_thms_p, Diff_0]) 1); +by (safe_tac ZF_cs); +(*Case hyps(p,t)-cons(#v,Y) |- p *) +by (rtac thms_excluded_middle_rule 1); +by (etac VarI 3); +by (rtac (cons_Diff_same RS weaken_left) 1); +by (etac spec 1); +by (rtac (cons_Diff_subset2 RS weaken_left) 1); +by (rtac (premp RS hyps_Diff RS Diff_weaken_left) 1); +by (etac spec 1); +(*Case hyps(p,t)-cons(#v => Fls,Y) |- p *) +by (rtac thms_excluded_middle_rule 1); +by (etac VarI 3); +by (rtac (cons_Diff_same RS weaken_left) 2); +by (etac spec 2); +by (rtac (cons_Diff_subset2 RS weaken_left) 1); +by (rtac (premp RS hyps_cons RS Diff_weaken_left) 1); +by (etac spec 1); +val completeness_0_lemma = result(); + +(*The base case for completeness*) +val [premp,sat] = goal PropThms.thy "[| p: prop; 0 |= p |] ==> 0 |- p"; +by (rtac (Diff_cancel RS subst) 1); +by (rtac (sat RS (premp RS completeness_0_lemma RS spec)) 1); +val completeness_0 = result(); + +(*A semantic analogue of the Deduction Theorem*) +goalw PropThms.thy [sat_def] "!!H p q. [| cons(p,H) |= q |] ==> H |= p=>q"; +by (SIMP_TAC prop_ss 1); +by (fast_tac ZF_cs 1); +val sat_Imp = result(); + +goal PropThms.thy "!!H. H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p"; +by (etac Fin_induct 1); +by (safe_tac (ZF_cs addSIs [completeness_0])); +by (rtac (weaken_left_cons RS thms_MP) 1); +by (fast_tac (ZF_cs addSIs [sat_Imp,ImpI]) 1); +by (fast_tac thms_cs 1); +val completeness_lemma = result(); + +val completeness = completeness_lemma RS bspec RS mp; + +val [finite] = goal PropThms.thy "H: Fin(prop) ==> H |- p <-> H |= p & p:prop"; +by (fast_tac (ZF_cs addSEs [soundness, finite RS completeness, + thms_in_pl]) 1); +val thms_iff = result(); + +writeln"Reached end of file."; + + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/PropLog.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/PropLog.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,43 @@ +(* Title: ZF/ex/prop-log.thy + ID: $Id$ + Author: Tobias Nipkow & Lawrence C Paulson + Copyright 1993 University of Cambridge + +Inductive definition of propositional logic. +*) + +PropLog = Prop + Fin + +consts + (*semantics*) + prop_rec :: "[i, i, i=>i, [i,i,i,i]=>i] => i" + is_true :: "[i,i] => o" + "|=" :: "[i,i] => o" (infixl 50) + hyps :: "[i,i] => i" + + (*proof theory*) + thms :: "i => i" + "|-" :: "[i,i] => o" (infixl 50) + +translations + "H |- p" == "p : thms(H)" + +rules + + prop_rec_def + "prop_rec(p,b,c,h) == \ +\ Vrec(p, %p g.prop_case(b, c, %x y. h(x, y, g`x, g`y), p))" + + (** Semantics of propositional logic **) + is_true_def + "is_true(p,t) == prop_rec(p, 0, %v. if(v:t, 1, 0), \ +\ %p q tp tq. if(tp=1,tq,1)) = 1" + + (*For every valuation, if all elements of H are true then so is p*) + sat_def "H |= p == ALL t. (ALL q:H. is_true(q,t)) --> is_true(p,t)" + + (** A finite set of hypotheses from t and the Vars in p **) + hyps_def + "hyps(p,t) == prop_rec(p, 0, %v. {if(v:t, #v, #v=>Fls)}, \ +\ %p q Hp Hq. Hp Un Hq)" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/ROOT.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,55 @@ +(* Title: ZF/ex/ROOT + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Executes all examples for Zermelo-Fraenkel Set Theory +*) + +ZF_build_completed; (*Cause examples to fail if ZF did*) + +writeln"Root file for ZF Set Theory examples"; +proof_timing := true; + +time_use "ex/misc.ML"; +time_use_thy "ex/ramsey"; + +(*Equivalence classes and integers*) +time_use_thy "ex/equiv"; +time_use_thy "ex/integ"; +(*Binary integer arithmetic*) +use "ex/twos-compl.ML"; +time_use "ex/bin.ML"; +time_use_thy "ex/bin-fn"; + +(** Datatypes **) +(*binary trees*) +time_use "ex/bt.ML"; +time_use_thy "ex/bt-fn"; +(*terms: recursion over the list functor*) +time_use "ex/term.ML"; +time_use_thy "ex/term-fn"; +(*trees/forests: mutual recursion*) +time_use "ex/tf.ML"; +time_use_thy "ex/tf-fn"; +(*Enormous enumeration type -- could be even bigger?*) +time_use "ex/enum.ML"; + +(** Inductive definitions **) +(*completeness of propositional logic*) +time_use "ex/prop.ML"; +time_use_thy "ex/prop-log"; +(*two Coq examples by Ch. Paulin-Mohring*) +time_use "ex/listn.ML"; +time_use "ex/acc.ML"; +(*Diamond property for combinatory logic*) +time_use "ex/comb.ML"; +time_use_thy "ex/contract"; +time_use "ex/parcontract.ML"; + +(** Co-Datatypes **) +time_use "ex/llist.ML"; +time_use_thy "ex/llist-fn"; + + +maketest"END: Root file for ZF Set Theory examples"; diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Ramsey.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Ramsey.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,236 @@ +(* Title: ZF/ex/ramsey.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Ramsey's Theorem (finite exponent 2 version) + +Based upon the article + D Basin and M Kaufmann, + The Boyer-Moore Prover and Nuprl: An Experimental Comparison. + In G Huet and G Plotkin, editors, Logical Frameworks. + (CUP, 1991), pages 89--119 + +See also + M Kaufmann, + An example in NQTHM: Ramsey's Theorem + Internal Note, Computational Logic, Inc., Austin, Texas 78703 + Available from the author: kaufmann@cli.com +*) + +open Ramsey; + +val ramsey_congs = mk_congs Ramsey.thy ["Atleast"]; +val ramsey_ss = arith_ss addcongs ramsey_congs; + +(*** Cliques and Independent sets ***) + +goalw Ramsey.thy [Clique_def] "Clique(0,V,E)"; +by (fast_tac ZF_cs 1); +val Clique0 = result(); + +goalw Ramsey.thy [Clique_def] + "!!C V E. [| Clique(C,V',E); V'<=V |] ==> Clique(C,V,E)"; +by (fast_tac ZF_cs 1); +val Clique_superset = result(); + +goalw Ramsey.thy [Indept_def] "Indept(0,V,E)"; +by (fast_tac ZF_cs 1); +val Indept0 = result(); + +val prems = goalw Ramsey.thy [Indept_def] + "!!I V E. [| Indept(I,V',E); V'<=V |] ==> Indept(I,V,E)"; +by (fast_tac ZF_cs 1); +val Indept_superset = result(); + +(*** Atleast ***) + +goalw Ramsey.thy [Atleast_def,inj_def] "Atleast(0,A)"; +by (fast_tac (ZF_cs addIs [PiI]) 1); +val Atleast0 = result(); + +val [major] = goalw Ramsey.thy [Atleast_def] + "Atleast(succ(m),A) ==> EX x:A. Atleast(m, A-{x})"; +by (rtac (major RS exE) 1); +by (rtac bexI 1); +by (etac (inj_is_fun RS apply_type) 2); +by (rtac succI1 2); +by (rtac exI 1); +by (etac inj_succ_restrict 1); +val Atleast_succD = result(); + +val major::prems = goalw Ramsey.thy [Atleast_def] + "[| Atleast(n,A); A<=B |] ==> Atleast(n,B)"; +by (rtac (major RS exE) 1); +by (rtac exI 1); +by (etac inj_weaken_type 1); +by (resolve_tac prems 1); +val Atleast_superset = result(); + +val prems = goalw Ramsey.thy [Atleast_def,succ_def] + "[| Atleast(m,B); ~ b: B |] ==> Atleast(succ(m), cons(b,B))"; +by (cut_facts_tac prems 1); +by (etac exE 1); +by (rtac exI 1); +by (etac inj_extend 1); +by (rtac mem_not_refl 1); +by (assume_tac 1); +val Atleast_succI = result(); + +val prems = goal Ramsey.thy + "[| Atleast(m, B-{x}); x: B |] ==> Atleast(succ(m), B)"; +by (cut_facts_tac prems 1); +by (etac (Atleast_succI RS Atleast_superset) 1); +by (fast_tac ZF_cs 1); +by (fast_tac ZF_cs 1); +val Atleast_Diff_succI = result(); + +(*** Main Cardinality Lemma ***) + +(*The #-succ(0) strengthens the original theorem statement, but precisely + the same proof could be used!!*) +val prems = goal Ramsey.thy + "m: nat ==> \ +\ ALL n: nat. ALL A B. Atleast((m#+n) #- succ(0), A Un B) --> \ +\ Atleast(m,A) | Atleast(n,B)"; +by (nat_ind_tac "m" prems 1); +by (fast_tac (ZF_cs addSIs [Atleast0]) 1); +by (ASM_SIMP_TAC ramsey_ss 1); +by (rtac ballI 1); +by (nat_ind_tac "n" [] 1); +by (fast_tac (ZF_cs addSIs [Atleast0]) 1); +by (ASM_SIMP_TAC (ramsey_ss addrews [add_succ_right]) 1); +by (safe_tac ZF_cs); +by (etac (Atleast_succD RS bexE) 1); +by (etac UnE 1); +(**case x:B. Instantiate the 'ALL A B' induction hypothesis. **) +by (dres_inst_tac [("x1","A"), ("x","B-{x}")] (spec RS spec) 2); +by (etac (mp RS disjE) 2); +(*cases Atleast(succ(m1),A) and Atleast(succ(n1),B)*) +by (REPEAT (eresolve_tac [asm_rl, notE, Atleast_Diff_succI] 3)); +(*proving the condition*) +by (etac Atleast_superset 2 THEN fast_tac ZF_cs 2); +(**case x:A. Instantiate the 'ALL n:nat. ALL A B' induction hypothesis. **) +by (dres_inst_tac [("x2","succ(n1)"), ("x1","A-{x}"), ("x","B")] + (bspec RS spec RS spec) 1); +by (etac nat_succI 1); +by (etac (mp RS disjE) 1); +(*cases Atleast(succ(m1),A) and Atleast(succ(n1),B)*) +by (REPEAT (eresolve_tac [asm_rl, Atleast_Diff_succI, notE] 2)); +(*proving the condition*) +by (ASM_SIMP_TAC (ramsey_ss addrews [add_succ_right]) 1); +by (etac Atleast_superset 1 THEN fast_tac ZF_cs 1); +val pigeon2_lemma = result(); + +(* [| m:nat; n:nat; Atleast(m #+ n #- succ(0), A Un B) |] ==> + Atleast(m,A) | Atleast(n,B) *) +val pigeon2 = standard (pigeon2_lemma RS bspec RS spec RS spec RS mp); + + +(**** Ramsey's Theorem ****) + +(** Base cases of induction; they now admit ANY Ramsey number **) + +goalw Ramsey.thy [Ramsey_def] "Ramsey(n,0,j)"; +by (fast_tac (ZF_cs addIs [Clique0,Atleast0]) 1); +val Ramsey0j = result(); + +goalw Ramsey.thy [Ramsey_def] "Ramsey(n,i,0)"; +by (fast_tac (ZF_cs addIs [Indept0,Atleast0]) 1); +val Ramseyi0 = result(); + +(** Lemmas for induction step **) + +(*The use of succ(m) here, rather than #-succ(0), simplifies the proof of + Ramsey_step_lemma.*) +val prems = goal Ramsey.thy + "[| Atleast(m #+ n, A); m: nat; n: nat |] ==> \ +\ Atleast(succ(m), {x:A. ~P(x)}) | Atleast(n, {x:A. P(x)})"; +by (rtac (nat_succI RS pigeon2) 1); +by (SIMP_TAC (ramsey_ss addrews prems) 3); +by (rtac Atleast_superset 3); +by (REPEAT (resolve_tac prems 1)); +by (fast_tac ZF_cs 1); +val Atleast_partition = result(); + +(*For the Atleast part, proves ~(a:I) from the second premise!*) +val prems = goalw Ramsey.thy [Symmetric_def,Indept_def] + "[| Symmetric(E); Indept(I, {z: V-{a}. ~ :E}, E); a: V; \ +\ Atleast(j,I) |] ==> \ +\ Indept(cons(a,I), V, E) & Atleast(succ(j), cons(a,I))"; +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addSEs [Atleast_succI]) 1); (*34 secs*) +val Indept_succ = result(); + +val prems = goalw Ramsey.thy [Symmetric_def,Clique_def] + "[| Symmetric(E); Clique(C, {z: V-{a}. :E}, E); a: V; \ +\ Atleast(j,C) |] ==> \ +\ Clique(cons(a,C), V, E) & Atleast(succ(j), cons(a,C))"; +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addSEs [Atleast_succI]) 1); (*41 secs*) +val Clique_succ = result(); + +(** Induction step **) + +(*Published proofs gloss over the need for Ramsey numbers to be POSITIVE.*) +val ram1::ram2::prems = goalw Ramsey.thy [Ramsey_def] + "[| Ramsey(succ(m), succ(i), j); Ramsey(n, i, succ(j)); \ +\ m: nat; n: nat |] ==> \ +\ Ramsey(succ(m#+n), succ(i), succ(j))"; +by (safe_tac ZF_cs); +by (etac (Atleast_succD RS bexE) 1); +by (eres_inst_tac [("P1","%z.:E")] (Atleast_partition RS disjE) 1); +by (REPEAT (resolve_tac prems 1)); +(*case m*) +by (rtac (ram1 RS spec RS spec RS mp RS disjE) 1); +by (fast_tac ZF_cs 1); +by (fast_tac (ZF_cs addEs [Clique_superset]) 1); (*easy -- given a Clique*) +by (safe_tac ZF_cs); +by (eresolve_tac (swapify [exI]) 1); (*ignore main EX quantifier*) +by (REPEAT (ares_tac [Indept_succ] 1)); (*make a bigger Indept*) +(*case n*) +by (rtac (ram2 RS spec RS spec RS mp RS disjE) 1); +by (fast_tac ZF_cs 1); +by (safe_tac ZF_cs); +by (rtac exI 1); +by (REPEAT (ares_tac [Clique_succ] 1)); (*make a bigger Clique*) +by (fast_tac (ZF_cs addEs [Indept_superset]) 1); (*easy -- given an Indept*) +val Ramsey_step_lemma = result(); + + +(** The actual proof **) + +(*Again, the induction requires Ramsey numbers to be positive.*) +val prems = goal Ramsey.thy + "i: nat ==> ALL j: nat. EX n:nat. Ramsey(succ(n), i, j)"; +by (nat_ind_tac "i" prems 1); +by (fast_tac (ZF_cs addSIs [nat_0I,Ramsey0j]) 1); +by (rtac ballI 1); +by (nat_ind_tac "j" [] 1); +by (fast_tac (ZF_cs addSIs [nat_0I,Ramseyi0]) 1); +by (dres_inst_tac [("x","succ(j1)")] bspec 1); +by (REPEAT (eresolve_tac [nat_succI,bexE] 1)); +by (rtac bexI 1); +by (rtac Ramsey_step_lemma 1); +by (REPEAT (ares_tac [nat_succI,add_type] 1)); +val ramsey_lemma = result(); + +(*Final statement in a tidy form, without succ(...) *) +val prems = goal Ramsey.thy + "[| i: nat; j: nat |] ==> EX n:nat. Ramsey(n,i,j)"; +by (rtac (ramsey_lemma RS bspec RS bexE) 1); +by (etac bexI 3); +by (REPEAT (ares_tac (prems@[nat_succI]) 1)); +val ramsey = result(); + +(*Computer Ramsey numbers according to proof above -- which, actually, + does not constrain the base case values at all!*) +fun ram 0 j = 1 + | ram i 0 = 1 + | ram i j = ram (i-1) j + ram i (j-1); + +(*Previous proof gave the following Ramsey numbers, which are smaller than + those above by one!*) +fun ram' 0 j = 0 + | ram' i 0 = 0 + | ram' i j = ram' (i-1) j + ram' i (j-1) + 1; diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Ramsey.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Ramsey.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,46 @@ +(* Title: ZF/ex/ramsey.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Ramsey's Theorem (finite exponent 2 version) + +Based upon the article + D Basin and M Kaufmann, + The Boyer-Moore Prover and Nuprl: An Experimental Comparison. + In G Huet and G Plotkin, editors, Logical Frameworks. + (CUP, 1991), pages 89--119 + +See also + M Kaufmann, + An example in NQTHM: Ramsey's Theorem + Internal Note, Computational Logic, Inc., Austin, Texas 78703 + Available from the author: kaufmann@cli.com +*) + +Ramsey = Arith + +consts + Symmetric :: "i=>o" + Atleast :: "[i,i]=>o" + Clique,Indept,Ramsey :: "[i,i,i]=>o" + +rules + + Symmetric_def + "Symmetric(E) == (ALL x y. :E --> :E)" + + Clique_def + "Clique(C,V,E) == (C<=V) & (ALL x:C. ALL y:C. ~ x=y --> :E)" + + Indept_def + "Indept(I,V,E) == (I<=V) & (ALL x:I. ALL y:I. ~ x=y --> ~ :E)" + + Atleast_def + "Atleast(n,S) == (EX f. f: inj(n,S))" + + Ramsey_def + "Ramsey(n,i,j) == ALL V E. Symmetric(E) & Atleast(n,V) --> \ +\ (EX C. Clique(C,V,E) & Atleast(i,C)) | \ +\ (EX I. Indept(I,V,E) & Atleast(j,I))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/TF.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/TF.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,65 @@ +(* Title: ZF/ex/tf.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Trees & forests, a mutually recursive type definition. +*) + +structure TF = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("tree", "univ(A)", + [(["Tcons"], "[i,i]=>i")]), + ("forest", "univ(A)", + [(["Fnil"], "i"), + (["Fcons"], "[i,i]=>i")])]; + val rec_styp = "i=>i"; + val ext = None + val sintrs = + ["[| a:A; tf: forest(A) |] ==> Tcons(a,tf) : tree(A)", + "Fnil : forest(A)", + "[| t: tree(A); tf: forest(A) |] ==> Fcons(t,tf) : forest(A)"]; + val monos = []; + val type_intrs = data_typechecks + val type_elims = []); + +val [TconsI, FnilI, FconsI] = TF.intrs; + +(** tree_forest(A) as the union of tree(A) and forest(A) **) + +goalw TF.thy (tl TF.defs) "tree(A) <= tree_forest(A)"; +by (rtac Part_subset 1); +val tree_subset_TF = result(); + +goalw TF.thy (tl TF.defs) "forest(A) <= tree_forest(A)"; +by (rtac Part_subset 1); +val forest_subset_TF = result(); + +goalw TF.thy (tl TF.defs) "tree(A) Un forest(A) = tree_forest(A)"; +by (rtac (TF.dom_subset RS Part_sum_equality) 1); +val TF_equals_Un = result(); + +(** NOT useful, but interesting... **) + +goalw TF.thy (tl TF.defs) "tree(A) = {Inl(x). x: A*forest(A)}"; +by (res_inst_tac [("P", "%x.?t(x) = ?u::i")] (TF.unfold RS ssubst) 1); +by (rewrite_goals_tac TF.con_defs); +by (rtac equalityI 1); +by (fast_tac (ZF_cs addIs [PartI] addSEs (PartE::TF.free_SEs)) 1); +by (fast_tac (ZF_cs addIs ([PartI,InlI,InrI] @ [TF.dom_subset RS subsetD] + @ data_typechecks) + addSEs (PartE::TF.free_SEs)) 1); +val tree_unfold = result(); + +goalw TF.thy (tl TF.defs) "forest(A) = {Inr(x). x: {0} + tree(A)*forest(A)}"; +by (res_inst_tac [("P", "%x.?t(x) = ?u::i")] (TF.unfold RS ssubst) 1); +by (rewrite_goals_tac TF.con_defs); +by (rtac equalityI 1); +by (fast_tac (ZF_cs addIs [PartI,InlI,InrI] + addSEs (PartE::TF.free_SEs)) 1); +by (fast_tac (ZF_cs addIs ([PartI,InlI,InrI] @ [TF.dom_subset RS subsetD] + @ data_typechecks) + addSEs ([PartE, sumE] @ TF.free_SEs)) 1); +val forest_unfold = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/TF_Fn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/TF_Fn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,233 @@ +(* Title: ZF/ex/tf.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For tf.thy. Trees & forests, a mutually recursive type definition. + +Still needs + +"TF_reflect == (%z. TF_rec(z, %x ts r. Tcons(x,r), 0, + %t ts r1 r2. TF_of_list(list_of_TF(r2) @ )))" +*) + +open TF_Fn; + + +(*** TF_rec -- by Vset recursion ***) + +(*Used only to verify TF_rec*) +val TF_congs = mk_typed_congs TF.thy + [("b", "[i,i,i]=>i"), ("d", "[i,i,i,i]=>i")]; + +(** conversion rules **) + +goal TF_Fn.thy "TF_rec(Tcons(a,tf), b, c, d) = b(a, tf, TF_rec(tf,b,c,d))"; +by (rtac (TF_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac TF.con_defs); +by (SIMP_TAC (rank_ss addcongs TF_congs) 1); +val TF_rec_Tcons = result(); + +goal TF_Fn.thy "TF_rec(Fnil, b, c, d) = c"; +by (rtac (TF_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac TF.con_defs); +by (SIMP_TAC rank_ss 1); +val TF_rec_Fnil = result(); + +goal TF_Fn.thy "TF_rec(Fcons(t,tf), b, c, d) = \ +\ d(t, tf, TF_rec(t, b, c, d), TF_rec(tf, b, c, d))"; +by (rtac (TF_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac TF.con_defs); +by (SIMP_TAC (rank_ss addcongs TF_congs) 1); +val TF_rec_Fcons = result(); + +(*list_ss includes list operations as well as arith_ss*) +val TF_rec_ss = list_ss addrews + [TF_rec_Tcons, TF_rec_Fnil, TF_rec_Fcons, TconsI, FnilI, FconsI]; + +(** Type checking **) + +val major::prems = goal TF_Fn.thy + "[| z: tree_forest(A); \ +\ !!x tf r. [| x: A; tf: forest(A); r: C(tf) \ +\ |] ==> b(x,tf,r): C(Tcons(x,tf)); \ +\ c : C(Fnil); \ +\ !!t tf r1 r2. [| t: tree(A); tf: forest(A); r1: C(t); r2: C(tf) \ +\ |] ==> d(t,tf,r1,r2): C(Fcons(t,tf)) \ +\ |] ==> TF_rec(z,b,c,d) : C(z)"; +by (rtac (major RS TF.induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (TF_rec_ss addrews prems))); +val TF_rec_type = result(); + +(*Mutually recursive version*) +val prems = goal TF_Fn.thy + "[| !!x tf r. [| x: A; tf: forest(A); r: D(tf) \ +\ |] ==> b(x,tf,r): C(Tcons(x,tf)); \ +\ c : D(Fnil); \ +\ !!t tf r1 r2. [| t: tree(A); tf: forest(A); r1: C(t); r2: D(tf) \ +\ |] ==> d(t,tf,r1,r2): D(Fcons(t,tf)) \ +\ |] ==> (ALL t:tree(A). TF_rec(t,b,c,d) : C(t)) & \ +\ (ALL tf: forest(A). TF_rec(tf,b,c,d) : D(tf))"; +by (rewtac Ball_def); +by (rtac TF.mutual_induct 1); +by (ALLGOALS (ASM_SIMP_TAC (TF_rec_ss addrews prems))); +val tree_forest_rec_type = result(); + + +(** Versions for use with definitions **) + +val [rew] = goal TF_Fn.thy + "[| !!z. j(z)==TF_rec(z,b,c,d) |] ==> j(Tcons(a,tf)) = b(a,tf,j(tf))"; +by (rewtac rew); +by (rtac TF_rec_Tcons 1); +val def_TF_rec_Tcons = result(); + +val [rew] = goal TF_Fn.thy + "[| !!z. j(z)==TF_rec(z,b,c,d) |] ==> j(Fnil) = c"; +by (rewtac rew); +by (rtac TF_rec_Fnil 1); +val def_TF_rec_Fnil = result(); + +val [rew] = goal TF_Fn.thy + "[| !!z. j(z)==TF_rec(z,b,c,d) |] ==> j(Fcons(t,tf)) = d(t,tf,j(t),j(tf))"; +by (rewtac rew); +by (rtac TF_rec_Fcons 1); +val def_TF_rec_Fcons = result(); + +fun TF_recs def = map standard + ([def] RL [def_TF_rec_Tcons, def_TF_rec_Fnil, def_TF_rec_Fcons]); + + +(** list_of_TF and TF_of_list **) + +val [list_of_TF_Tcons, list_of_TF_Fnil, list_of_TF_Fcons] = + TF_recs list_of_TF_def; + +goalw TF_Fn.thy [list_of_TF_def] + "!!z A. z: tree_forest(A) ==> list_of_TF(z) : list(tree(A))"; +by (REPEAT (ares_tac ([TF_rec_type,TconsI] @ list_typechecks) 1)); +val list_of_TF_type = result(); + +val [TF_of_list_Nil,TF_of_list_Cons] = list_recs TF_of_list_def; + +goalw TF_Fn.thy [TF_of_list_def] + "!!l A. l: list(tree(A)) ==> TF_of_list(l) : forest(A)"; +by (REPEAT (ares_tac [list_rec_type, FnilI, FconsI] 1)); +val TF_of_list_type = result(); + + +(** TF_map **) + +val [TF_map_Tcons, TF_map_Fnil, TF_map_Fcons] = TF_recs TF_map_def; + +val prems = goalw TF_Fn.thy [TF_map_def] + "[| !!x. x: A ==> h(x): B |] ==> \ +\ (ALL t:tree(A). TF_map(h,t) : tree(B)) & \ +\ (ALL tf: forest(A). TF_map(h,tf) : forest(B))"; +by (REPEAT + (ares_tac ([tree_forest_rec_type, TconsI, FnilI, FconsI] @ prems) 1)); +val TF_map_type = result(); + + +(** TF_size **) + +val [TF_size_Tcons, TF_size_Fnil, TF_size_Fcons] = TF_recs TF_size_def; + +goalw TF_Fn.thy [TF_size_def] + "!!z A. z: tree_forest(A) ==> TF_size(z) : nat"; +by (REPEAT (ares_tac [TF_rec_type, add_type, nat_0I, nat_succI] 1)); +val TF_size_type = result(); + + +(** TF_preorder **) + +val [TF_preorder_Tcons, TF_preorder_Fnil, TF_preorder_Fcons] = + TF_recs TF_preorder_def; + +goalw TF_Fn.thy [TF_preorder_def] + "!!z A. z: tree_forest(A) ==> TF_preorder(z) : list(A)"; +by (REPEAT (ares_tac [TF_rec_type, app_type,NilI, ConsI] 1)); +val TF_preorder_type = result(); + + +(** Term simplification **) + +val treeI = tree_subset_TF RS subsetD +and forestI = forest_subset_TF RS subsetD; + +val TF_typechecks = + [TconsI, FnilI, FconsI, treeI, forestI, + list_of_TF_type, TF_map_type, TF_size_type, TF_preorder_type]; + +val TF_congs = TF.congs @ + mk_congs TF_Fn.thy + ["TF_rec", "list_of_TF", "TF_of_list", "TF_map", "TF_size", "TF_preorder"]; + +val TF_rewrites = + [TF_rec_Tcons, TF_rec_Fnil, TF_rec_Fcons, + list_of_TF_Tcons, list_of_TF_Fnil, list_of_TF_Fcons, + TF_of_list_Nil,TF_of_list_Cons, + TF_map_Tcons, TF_map_Fnil, TF_map_Fcons, + TF_size_Tcons, TF_size_Fnil, TF_size_Fcons, + TF_preorder_Tcons, TF_preorder_Fnil, TF_preorder_Fcons]; + +val TF_ss = list_ss addcongs TF_congs + addrews (TF_rewrites@TF_typechecks); + +(** theorems about list_of_TF and TF_of_list **) + +(*essentially the same as list induction*) +val major::prems = goal TF_Fn.thy + "[| tf: forest(A); \ +\ R(Fnil); \ +\ !!t tf. [| t: tree(A); tf: forest(A); R(tf) |] ==> R(Fcons(t,tf)) \ +\ |] ==> R(tf)"; +by (rtac (major RS (TF.mutual_induct RS conjunct2 RS spec RSN (2,rev_mp))) 1); +by (REPEAT (ares_tac (TrueI::prems) 1)); +val forest_induct = result(); + +goal TF_Fn.thy "!!tf A. tf: forest(A) ==> TF_of_list(list_of_TF(tf)) = tf"; +by (etac forest_induct 1); +by (ALLGOALS (ASM_SIMP_TAC TF_ss)); +val forest_iso = result(); + +goal TF_Fn.thy + "!!ts. ts: list(tree(A)) ==> list_of_TF(TF_of_list(ts)) = ts"; +by (etac List.induct 1); +by (ALLGOALS (ASM_SIMP_TAC TF_ss)); +val tree_list_iso = result(); + +(** theorems about TF_map **) + +goal TF_Fn.thy "!!z A. z: tree_forest(A) ==> TF_map(%u.u, z) = z"; +by (etac TF.induct 1); +by (ALLGOALS (ASM_SIMP_TAC TF_ss)); +val TF_map_ident = result(); + +goal TF_Fn.thy + "!!z A. z: tree_forest(A) ==> TF_map(h, TF_map(j,z)) = TF_map(%u.h(j(u)), z)"; +by (etac TF.induct 1); +by (ALLGOALS (ASM_SIMP_TAC TF_ss)); +val TF_map_compose = result(); + +(** theorems about TF_size **) + +goal TF_Fn.thy + "!!z A. z: tree_forest(A) ==> TF_size(TF_map(h,z)) = TF_size(z)"; +by (etac TF.induct 1); +by (ALLGOALS (ASM_SIMP_TAC TF_ss)); +val TF_size_TF_map = result(); + +goal TF_Fn.thy + "!!z A. z: tree_forest(A) ==> TF_size(z) = length(TF_preorder(z))"; +by (etac TF.induct 1); +by (ALLGOALS (ASM_SIMP_TAC (TF_ss addrews [length_app]))); +val TF_size_length = result(); + +(** theorems about TF_preorder **) + +goal TF_Fn.thy "!!z A. z: tree_forest(A) ==> \ +\ TF_preorder(TF_map(h,z)) = map(h, TF_preorder(z))"; +by (etac TF.induct 1); +by (ALLGOALS (ASM_SIMP_TAC (TF_ss addrews [map_app_distrib]))); +val TF_preorder_TF_map = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/TF_Fn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/TF_Fn.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,42 @@ +(* Title: ZF/ex/TF.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Trees & forests, a mutually recursive type definition. +*) + +TF_Fn = TF + ListFn + +consts + TF_rec :: "[i, [i,i,i]=>i, i, [i,i,i,i]=>i] => i" + TF_map :: "[i=>i, i] => i" + TF_size :: "i=>i" + TF_preorder :: "i=>i" + list_of_TF :: "i=>i" + TF_of_list :: "i=>i" + +rules + TF_rec_def + "TF_rec(z,b,c,d) == Vrec(z, \ +\ %z r. tree_forest_case(%x tf. b(x, tf, r`tf), \ +\ c, \ +\ %t tf. d(t, tf, r`t, r`tf), z))" + + list_of_TF_def + "list_of_TF(z) == TF_rec(z, %x tf r. [Tcons(x,tf)], [], \ +\ %t tf r1 r2. Cons(t, r2))" + + TF_of_list_def + "TF_of_list(tf) == list_rec(tf, Fnil, %t tf r. Fcons(t,r))" + + TF_map_def + "TF_map(h,z) == TF_rec(z, %x tf r.Tcons(h(x),r), Fnil, \ +\ %t tf r1 r2. Fcons(r1,r2))" + + TF_size_def + "TF_size(z) == TF_rec(z, %x tf r.succ(r), 0, %t tf r1 r2. r1#+r2)" + + TF_preorder_def + "TF_preorder(z) == TF_rec(z, %x tf r.Cons(x,r), Nil, %t tf r1 r2. r1@r2)" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/Term.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/Term.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,66 @@ +(* Title: ZF/ex/term.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Datatype definition of terms over an alphabet. +Illustrates the list functor (essentially the same type as in Trees & Forests) +*) + +structure Term = Datatype_Fun + (val thy = List.thy; + val rec_specs = + [("term", "univ(A)", + [(["Apply"], "[i,i]=>i")])]; + val rec_styp = "i=>i"; + val ext = None + val sintrs = ["[| a: A; l: list(term(A)) |] ==> Apply(a,l) : term(A)"]; + val monos = [list_mono]; + val type_intrs = [SigmaI,Pair_in_univ, list_univ RS subsetD, A_into_univ]; + val type_elims = []); + +val [ApplyI] = Term.intrs; + +(*Induction on term(A) followed by induction on List *) +val major::prems = goal Term.thy + "[| t: term(A); \ +\ !!x. [| x: A |] ==> P(Apply(x,Nil)); \ +\ !!x z zs. [| x: A; z: term(A); zs: list(term(A)); P(Apply(x,zs)) \ +\ |] ==> P(Apply(x, Cons(z,zs))) \ +\ |] ==> P(t)"; +by (rtac (major RS Term.induct) 1); +by (etac List.induct 1); +by (etac CollectE 2); +by (REPEAT (ares_tac (prems@[list_CollectD]) 1)); +val term_induct2 = result(); + +(*Induction on term(A) to prove an equation*) +val major::prems = goal (merge_theories(Term.thy,ListFn.thy)) + "[| t: term(A); \ +\ !!x zs. [| x: A; zs: list(term(A)); map(f,zs) = map(g,zs) |] ==> \ +\ f(Apply(x,zs)) = g(Apply(x,zs)) \ +\ |] ==> f(t)=g(t)"; +by (rtac (major RS Term.induct) 1); +by (resolve_tac prems 1); +by (REPEAT (eresolve_tac [asm_rl, map_list_Collect, list_CollectD] 1)); +val term_induct_eqn = result(); + +(** Lemmas to justify using "term" in other recursive type definitions **) + +goalw Term.thy Term.defs "!!A B. A<=B ==> term(A) <= term(B)"; +by (rtac lfp_mono 1); +by (REPEAT (rtac Term.bnd_mono 1)); +by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); +val term_mono = result(); + +(*Easily provable by induction also*) +goalw Term.thy (Term.defs@Term.con_defs) "term(univ(A)) <= univ(A)"; +by (rtac lfp_lowerbound 1); +by (rtac (A_subset_univ RS univ_mono) 2); +by (safe_tac ZF_cs); +by (REPEAT (ares_tac [Pair_in_univ, list_univ RS subsetD] 1)); +val term_univ = result(); + +val term_subset_univ = standard + (term_mono RS (term_univ RSN (2,subset_trans))); + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/TermFn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/TermFn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,192 @@ +(* Title: ZF/ex/term + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Terms over a given alphabet -- function applications; illustrates list functor + (essentially the same type as in Trees & Forests) +*) + +writeln"File ZF/ex/term-fn."; + +open TermFn; + +(*** term_rec -- by Vset recursion ***) + +(*Lemma: map works correctly on the underlying list of terms*) +val [major,ordi] = goal ListFn.thy + "[| l: list(A); Ord(i) |] ==> \ +\ rank(l): i --> map(%z. (lam x:Vset(i).h(x)) ` z, l) = map(h,l)"; +by (rtac (major RS List.induct) 1); +by (SIMP_TAC list_ss 1); +by (rtac impI 1); +by (forward_tac [rank_Cons1 RS Ord_trans] 1); +by (dtac (rank_Cons2 RS Ord_trans) 2); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [ordi, VsetI]))); +val map_lemma = result(); + +(*Typing premise is necessary to invoke map_lemma*) +val [prem] = goal TermFn.thy + "ts: list(A) ==> \ +\ term_rec(Apply(a,ts), d) = d(a, ts, map (%z. term_rec(z,d), ts))"; +by (rtac (term_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Term.con_defs); +val term_rec_ss = ZF_ss + addcongs (mk_typed_congs TermFn.thy [("d", "[i,i,i]=>i")]) + addrews [Ord_rank, rank_pair2, prem RS map_lemma]; +by (SIMP_TAC term_rec_ss 1); +val term_rec = result(); + +(*Slightly odd typing condition on r in the second premise!*) +val major::prems = goal TermFn.thy + "[| t: term(A); \ +\ !!x zs r. [| x: A; zs: list(term(A)); \ +\ r: list(UN t:term(A). C(t)) |] \ +\ ==> d(x, zs, r): C(Apply(x,zs)) \ +\ |] ==> term_rec(t,d) : C(t)"; +by (rtac (major RS Term.induct) 1); +by (forward_tac [list_CollectD] 1); +by (rtac (term_rec RS ssubst) 1); +by (REPEAT (ares_tac prems 1)); +by (etac List.induct 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [term_rec]))); +by (etac CollectE 1); +by (REPEAT (ares_tac [ConsI, UN_I] 1)); +val term_rec_type = result(); + +val [rew,tslist] = goal TermFn.thy + "[| !!t. j(t)==term_rec(t,d); ts: list(A) |] ==> \ +\ j(Apply(a,ts)) = d(a, ts, map(%Z.j(Z), ts))"; +by (rewtac rew); +by (rtac (tslist RS term_rec) 1); +val def_term_rec = result(); + +val prems = goal TermFn.thy + "[| t: term(A); \ +\ !!x zs r. [| x: A; zs: list(term(A)); r: list(C) |] \ +\ ==> d(x, zs, r): C \ +\ |] ==> term_rec(t,d) : C"; +by (REPEAT (ares_tac (term_rec_type::prems) 1)); +by (etac (subset_refl RS UN_least RS list_mono RS subsetD) 1); +val term_rec_simple_type = result(); + + +(** term_map **) + +val term_map = standard (term_map_def RS def_term_rec); + +val prems = goalw TermFn.thy [term_map_def] + "[| t: term(A); !!x. x: A ==> f(x): B |] ==> term_map(f,t) : term(B)"; +by (REPEAT (ares_tac ([term_rec_simple_type, ApplyI] @ prems) 1)); +val term_map_type = result(); + +val [major] = goal TermFn.thy + "t: term(A) ==> term_map(f,t) : term({f(u). u:A})"; +by (rtac (major RS term_map_type) 1); +by (etac RepFunI 1); +val term_map_type2 = result(); + + +(** term_size **) + +val term_size = standard (term_size_def RS def_term_rec); + +goalw TermFn.thy [term_size_def] "!!t A. t: term(A) ==> term_size(t) : nat"; +by (REPEAT (ares_tac [term_rec_simple_type, list_add_type, nat_succI] 1)); +val term_size_type = result(); + + +(** reflect **) + +val reflect = standard (reflect_def RS def_term_rec); + +goalw TermFn.thy [reflect_def] "!!t A. t: term(A) ==> reflect(t) : term(A)"; +by (REPEAT (ares_tac [term_rec_simple_type, rev_type, ApplyI] 1)); +val reflect_type = result(); + +(** preorder **) + +val preorder = standard (preorder_def RS def_term_rec); + +goalw TermFn.thy [preorder_def] + "!!t A. t: term(A) ==> preorder(t) : list(A)"; +by (REPEAT (ares_tac [term_rec_simple_type, ConsI, flat_type] 1)); +val preorder_type = result(); + + +(** Term simplification **) + +val term_typechecks = + [ApplyI, term_map_type, term_map_type2, term_size_type, reflect_type, + preorder_type]; + +(*map_type2 and term_map_type2 instantiate variables*) +val term_ss = list_ss + addcongs (Term.congs @ + mk_congs TermFn.thy ["term_rec","term_map","term_size", + "reflect","preorder"]) + addrews [term_rec, term_map, term_size, reflect, + preorder] + setauto type_auto_tac (list_typechecks@term_typechecks); + + +(** theorems about term_map **) + +goal TermFn.thy "!!t A. t: term(A) ==> term_map(%u.u, t) = t"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [map_ident]) 1); +val term_map_ident = result(); + +goal TermFn.thy + "!!t A. t: term(A) ==> term_map(f, term_map(g,t)) = term_map(%u.f(g(u)), t)"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [map_compose]) 1); +val term_map_compose = result(); + +goal TermFn.thy + "!!t A. t: term(A) ==> term_map(f, reflect(t)) = reflect(term_map(f,t))"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [rev_map_distrib RS sym, map_compose]) 1); +val term_map_reflect = result(); + + +(** theorems about term_size **) + +goal TermFn.thy + "!!t A. t: term(A) ==> term_size(term_map(f,t)) = term_size(t)"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [map_compose]) 1); +val term_size_term_map = result(); + +goal TermFn.thy "!!t A. t: term(A) ==> term_size(reflect(t)) = term_size(t)"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [rev_map_distrib RS sym, map_compose, + list_add_rev]) 1); +val term_size_reflect = result(); + +goal TermFn.thy "!!t A. t: term(A) ==> term_size(t) = length(preorder(t))"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [length_flat, map_compose]) 1); +val term_size_length = result(); + + +(** theorems about reflect **) + +goal TermFn.thy "!!t A. t: term(A) ==> reflect(reflect(t)) = t"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [rev_map_distrib, map_compose, + map_ident, rev_rev_ident]) 1); +val reflect_reflect_ident = result(); + + +(** theorems about preorder **) + +goal TermFn.thy + "!!t A. t: term(A) ==> preorder(term_map(f,t)) = map(f, preorder(t))"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [map_compose, map_flat]) 1); +val preorder_term_map = result(); + +(** preorder(reflect(t)) = rev(postorder(t)) **) + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/TermFn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/TermFn.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,31 @@ +(* Title: ZF/ex/term-fn.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Terms over an alphabet. +Illustrates the list functor (essentially the same type as in Trees & Forests) +*) + +TermFn = Term + ListFn + +consts + term_rec :: "[i, [i,i,i]=>i] => i" + term_map :: "[i=>i, i] => i" + term_size :: "i=>i" + reflect :: "i=>i" + preorder :: "i=>i" + +rules + term_rec_def + "term_rec(t,d) == \ +\ Vrec(t, %t g. term_case(%x zs. d(x, zs, map(%z.g`z, zs)), t))" + + term_map_def "term_map(f,t) == term_rec(t, %x zs rs. Apply(f(x), rs))" + + term_size_def "term_size(t) == term_rec(t, %x zs rs. succ(list_add(rs)))" + + reflect_def "reflect(t) == term_rec(t, %x zs rs. Apply(x, rev(rs)))" + + preorder_def "preorder(t) == term_rec(t, %x zs rs. Cons(x, flat(rs)))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/acc.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/acc.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,65 @@ +(* Title: ZF/ex/acc + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Inductive definition of acc(r) + +See Ch. Paulin-Mohring, Inductive Definitions in the System Coq. +Research Report 92-49, LIP, ENS Lyon. Dec 1992. +*) + +structure Acc = Inductive_Fun + (val thy = WF.thy addconsts [(["acc"],"i=>i")]; + val rec_doms = [("acc", "field(r)")]; + val sintrs = + ["[| r-``{b} : Pow(acc(r)); b : field(r) |] ==> b : acc(r)"]; + val monos = [Pow_mono]; + val con_defs = []; + val type_intrs = []; + val type_elims = []); + +goal Acc.thy "!!a b r. [| b: acc(r); : r |] ==> a: acc(r)"; +by (etac Acc.elim 1); +by (fast_tac ZF_cs 1); +val acc_downward = result(); + +val [major] = goal Acc.thy "field(r) <= acc(r) ==> wf(r)"; +by (rtac (major RS wfI2) 1); +by (rtac subsetI 1); +by (etac Acc.induct 1); +by (etac (bspec RS mp) 1); +by (resolve_tac Acc.intrs 1); +by (assume_tac 2); +by (ALLGOALS (fast_tac ZF_cs)); +val acc_wfI = result(); + +goal ZF.thy "!!r A. field(r Int A*A) <= field(r) Int A"; +by (fast_tac ZF_cs 1); +val field_Int_prodself = result(); + +goal Acc.thy "wf(r Int (acc(r)*acc(r)))"; +by (rtac (field_Int_prodself RS wfI2) 1); +by (rtac subsetI 1); +by (etac IntE 1); +by (etac Acc.induct 1); +by (etac (bspec RS mp) 1); +by (rtac IntI 1); +by (assume_tac 1); +by (resolve_tac Acc.intrs 1); +by (assume_tac 2); +by (ALLGOALS (fast_tac ZF_cs)); +val wf_acc_Int = result(); + +val [major] = goal Acc.thy "wf(r) ==> field(r) <= acc(r)"; +by (rtac subsetI 1); +by (etac (major RS wf_induct2) 1); +by (rtac subset_refl 1); +by (resolve_tac Acc.intrs 1); +by (assume_tac 2); +by (fast_tac ZF_cs 1); +val acc_wfD = result(); + +goal Acc.thy "wf(r) <-> field(r) <= acc(r)"; +by (EVERY1 [rtac iffI, etac acc_wfD, etac acc_wfI]); +val wf_acc_iff = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/bin.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/bin.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,42 @@ +(* Title: ZF/ex/bin.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Datatype of binary integers +*) + +(*Example of a datatype with an infix constructor*) +structure Bin = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("bin", "univ(0)", + [(["Plus", "Minus"], "i"), + (["op $$"], "[i,i]=>i")])]; + val rec_styp = "i"; + val ext = Some (NewSext { + mixfix = + [Infixl("$$", "[i,i] => i", 60)], + xrules = [], + parse_ast_translation = [], + parse_preproc = None, + parse_postproc = None, + parse_translation = [], + print_translation = [], + print_preproc = None, + print_postproc = None, + print_ast_translation = []}); + val sintrs = + ["Plus : bin", + "Minus : bin", + "[| w: bin; b: bool |] ==> w$$b : bin"]; + val monos = []; + val type_intrs = bool_into_univ::data_typechecks; + val type_elims = []); + +(*Perform induction on l, then prove the major premise using prems. *) +fun bin_ind_tac a prems i = + EVERY [res_inst_tac [("x",a)] Bin.induct i, + rename_last_tac a ["1"] (i+3), + ares_tac prems i]; + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/binfn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/binfn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,463 @@ +(* Title: ZF/ex/bin.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For bin.thy. Arithmetic on binary integers. +*) + +open BinFn; + + +(** bin_rec -- by Vset recursion **) + +goal BinFn.thy "bin_rec(Plus,a,b,h) = a"; +by (rtac (bin_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Bin.con_defs); +by (SIMP_TAC rank_ss 1); +val bin_rec_Plus = result(); + +goal BinFn.thy "bin_rec(Minus,a,b,h) = b"; +by (rtac (bin_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Bin.con_defs); +by (SIMP_TAC rank_ss 1); +val bin_rec_Minus = result(); + +goal BinFn.thy "bin_rec(w$$x,a,b,h) = h(w, x, bin_rec(w,a,b,h))"; +by (rtac (bin_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Bin.con_defs); +by (SIMP_TAC (rank_ss addcongs + (mk_typed_congs BinFn.thy [("h", "[i,i,i]=>i")])) 1); +val bin_rec_Bcons = result(); + +(*Type checking*) +val prems = goal BinFn.thy + "[| w: bin; \ +\ a: C(Plus); b: C(Minus); \ +\ !!w x r. [| w: bin; x: bool; r: C(w) |] ==> h(w,x,r): C(w$$x) \ +\ |] ==> bin_rec(w,a,b,h) : C(w)"; +by (bin_ind_tac "w" prems 1); +by (ALLGOALS + (ASM_SIMP_TAC (ZF_ss addrews (prems@[bin_rec_Plus,bin_rec_Minus, + bin_rec_Bcons])))); +val bin_rec_type = result(); + +(** Versions for use with definitions **) + +val [rew] = goal BinFn.thy + "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(Plus) = a"; +by (rewtac rew); +by (rtac bin_rec_Plus 1); +val def_bin_rec_Plus = result(); + +val [rew] = goal BinFn.thy + "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(Minus) = b"; +by (rewtac rew); +by (rtac bin_rec_Minus 1); +val def_bin_rec_Minus = result(); + +val [rew] = goal BinFn.thy + "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(w$$x) = h(w,x,j(w))"; +by (rewtac rew); +by (rtac bin_rec_Bcons 1); +val def_bin_rec_Bcons = result(); + +fun bin_recs def = map standard + ([def] RL [def_bin_rec_Plus, def_bin_rec_Minus, def_bin_rec_Bcons]); + +(** Type checking **) + +val bin_typechecks0 = bin_rec_type :: Bin.intrs; + +goalw BinFn.thy [integ_of_bin_def] + "!!w. w: bin ==> integ_of_bin(w) : integ"; +by (typechk_tac (bin_typechecks0@integ_typechecks@ + nat_typechecks@[bool_into_nat])); +val integ_of_bin_type = result(); + +goalw BinFn.thy [bin_succ_def] + "!!w. w: bin ==> bin_succ(w) : bin"; +by (typechk_tac (bin_typechecks0@bool_typechecks)); +val bin_succ_type = result(); + +goalw BinFn.thy [bin_pred_def] + "!!w. w: bin ==> bin_pred(w) : bin"; +by (typechk_tac (bin_typechecks0@bool_typechecks)); +val bin_pred_type = result(); + +goalw BinFn.thy [bin_minus_def] + "!!w. w: bin ==> bin_minus(w) : bin"; +by (typechk_tac ([bin_pred_type]@bin_typechecks0@bool_typechecks)); +val bin_minus_type = result(); + +goalw BinFn.thy [bin_add_def] + "!!v w. [| v: bin; w: bin |] ==> bin_add(v,w) : bin"; +by (typechk_tac ([bin_succ_type,bin_pred_type]@bin_typechecks0@ + bool_typechecks@ZF_typechecks)); +val bin_add_type = result(); + +goalw BinFn.thy [bin_mult_def] + "!!v w. [| v: bin; w: bin |] ==> bin_mult(v,w) : bin"; +by (typechk_tac ([bin_minus_type,bin_add_type]@bin_typechecks0@ + bool_typechecks)); +val bin_mult_type = result(); + +val bin_typechecks = bin_typechecks0 @ + [integ_of_bin_type, bin_succ_type, bin_pred_type, + bin_minus_type, bin_add_type, bin_mult_type]; + +val bin_congs = mk_congs BinFn.thy + ["bin_rec","op $$","integ_of_bin","bin_succ","bin_pred", + "bin_minus","bin_add","bin_mult"]; + +val bin_ss = integ_ss + addcongs (bin_congs@bool_congs) + addrews([bool_1I, bool_0I, + bin_rec_Plus, bin_rec_Minus, bin_rec_Bcons] @ + bin_recs integ_of_bin_def @ bool_rews @ bin_typechecks); + +val typechecks = bin_typechecks @ integ_typechecks @ nat_typechecks @ + [bool_subset_nat RS subsetD]; + +(**** The carry/borrow functions, bin_succ and bin_pred ****) + +(** Lemmas **) + +goal Integ.thy + "!!z v. [| z $+ v = z' $+ v'; \ +\ z: integ; z': integ; v: integ; v': integ; w: integ |] \ +\ ==> z $+ (v $+ w) = z' $+ (v' $+ w)"; +by (ASM_SIMP_TAC (integ_ss addrews ([zadd_assoc RS sym])) 1); +val zadd_assoc_cong = result(); + +goal Integ.thy + "!!z v w. [| z: integ; v: integ; w: integ |] \ +\ ==> z $+ (v $+ w) = v $+ (z $+ w)"; +by (REPEAT (ares_tac [zadd_commute RS zadd_assoc_cong] 1)); +val zadd_assoc_swap = result(); + +val [zadd_cong] = mk_congs Integ.thy ["op $+"]; + +val zadd_kill = (refl RS zadd_cong); +val zadd_assoc_swap_kill = zadd_kill RSN (4, zadd_assoc_swap RS trans); + +(*Pushes 'constants' of the form $#m to the right -- LOOPS if two!*) +val zadd_assoc_znat = standard (znat_type RS zadd_assoc_swap); + +goal Integ.thy + "!!z w. [| z: integ; w: integ |] \ +\ ==> w $+ (z $+ (w $+ z)) = w $+ (w $+ (z $+ z))"; +by (REPEAT (ares_tac [zadd_kill, zadd_assoc_swap] 1)); +val zadd_swap_pairs = result(); + + +val carry_ss = bin_ss addrews + (bin_recs bin_succ_def @ bin_recs bin_pred_def); + +goal BinFn.thy + "!!w. w: bin ==> integ_of_bin(bin_succ(w)) = $#1 $+ integ_of_bin(w)"; +by (etac Bin.induct 1); +by (SIMP_TAC (carry_ss addrews [zadd_0_right]) 1); +by (SIMP_TAC (carry_ss addrews [zadd_zminus_inverse]) 1); +by (etac boolE 1); +by (ALLGOALS (ASM_SIMP_TAC (carry_ss addrews [zadd_assoc]))); +by (REPEAT (ares_tac (zadd_swap_pairs::typechecks) 1)); +val integ_of_bin_succ = result(); + +goal BinFn.thy + "!!w. w: bin ==> integ_of_bin(bin_pred(w)) = $~ ($#1) $+ integ_of_bin(w)"; +by (etac Bin.induct 1); +by (SIMP_TAC (carry_ss addrews [zadd_0_right]) 1); +by (SIMP_TAC (carry_ss addrews [zadd_zminus_inverse]) 1); +by (etac boolE 1); +by (ALLGOALS + (ASM_SIMP_TAC + (carry_ss addrews [zadd_assoc RS sym, + zadd_zminus_inverse, zadd_zminus_inverse2]))); +by (REPEAT (ares_tac ([zadd_commute, zadd_cong, refl]@typechecks) 1)); +val integ_of_bin_pred = result(); + +(*These two results replace the definitions of bin_succ and bin_pred*) + + +(*** bin_minus: (unary!) negation of binary integers ***) + +val bin_minus_ss = + bin_ss addrews (bin_recs bin_minus_def @ + [integ_of_bin_succ, integ_of_bin_pred]); + +goal BinFn.thy + "!!w. w: bin ==> integ_of_bin(bin_minus(w)) = $~ integ_of_bin(w)"; +by (etac Bin.induct 1); +by (SIMP_TAC (bin_minus_ss addrews [zminus_0]) 1); +by (SIMP_TAC (bin_minus_ss addrews [zadd_0_right]) 1); +by (etac boolE 1); +by (ALLGOALS + (ASM_SIMP_TAC (bin_minus_ss addrews [zminus_zadd_distrib, zadd_assoc]))); +val integ_of_bin_minus = result(); + + +(*** bin_add: binary addition ***) + +goalw BinFn.thy [bin_add_def] "!!w. w: bin ==> bin_add(Plus,w) = w"; +by (ASM_SIMP_TAC bin_ss 1); +val bin_add_Plus = result(); + +goalw BinFn.thy [bin_add_def] "!!w. w: bin ==> bin_add(Minus,w) = bin_pred(w)"; +by (ASM_SIMP_TAC bin_ss 1); +val bin_add_Minus = result(); + +goalw BinFn.thy [bin_add_def] "bin_add(v$$x,Plus) = v$$x"; +by (SIMP_TAC bin_ss 1); +val bin_add_Bcons_Plus = result(); + +goalw BinFn.thy [bin_add_def] "bin_add(v$$x,Minus) = bin_pred(v$$x)"; +by (SIMP_TAC bin_ss 1); +val bin_add_Bcons_Minus = result(); + +goalw BinFn.thy [bin_add_def] + "!!w y. [| w: bin; y: bool |] ==> \ +\ bin_add(v$$x, w$$y) = \ +\ bin_add(v, cond(x and y, bin_succ(w), w)) $$ (x xor y)"; +by (ASM_SIMP_TAC bin_ss 1); +val bin_add_Bcons_Bcons = result(); + +val bin_add_rews = [bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, + bin_add_Bcons_Minus, bin_add_Bcons_Bcons, + integ_of_bin_succ, integ_of_bin_pred]; + +val bin_add_ss = bin_ss addrews ([bool_subset_nat RS subsetD] @ bin_add_rews); + +goal BinFn.thy + "!!v. v: bin ==> \ +\ ALL w: bin. integ_of_bin(bin_add(v,w)) = \ +\ integ_of_bin(v) $+ integ_of_bin(w)"; +by (etac Bin.induct 1); +by (SIMP_TAC bin_add_ss 1); +by (SIMP_TAC bin_add_ss 1); +by (rtac ballI 1); +by (bin_ind_tac "wa" [] 1); +by (ASM_SIMP_TAC (bin_add_ss addrews [zadd_0_right]) 1); +by (ASM_SIMP_TAC bin_add_ss 1); +by (REPEAT (ares_tac (zadd_commute::typechecks) 1)); +by (etac boolE 1); +by (ASM_SIMP_TAC (bin_add_ss addrews [zadd_assoc, zadd_swap_pairs]) 2); +by (REPEAT (ares_tac ([refl, zadd_kill, zadd_assoc_swap_kill]@typechecks) 2)); +by (etac boolE 1); +by (ALLGOALS (ASM_SIMP_TAC (bin_add_ss addrews [zadd_assoc,zadd_swap_pairs]))); +by (REPEAT (ares_tac ([refl, zadd_kill, zadd_assoc_swap_kill RS sym]@ + typechecks) 1)); +val integ_of_bin_add_lemma = result(); + +val integ_of_bin_add = integ_of_bin_add_lemma RS bspec; + + +(*** bin_add: binary multiplication ***) + +val bin_mult_ss = + bin_ss addrews (bin_recs bin_mult_def @ + [integ_of_bin_minus, integ_of_bin_add]); + + +val major::prems = goal BinFn.thy + "[| v: bin; w: bin |] ==> \ +\ integ_of_bin(bin_mult(v,w)) = \ +\ integ_of_bin(v) $* integ_of_bin(w)"; +by (cut_facts_tac prems 1); +by (bin_ind_tac "v" [major] 1); +by (SIMP_TAC (bin_mult_ss addrews [zmult_0]) 1); +by (SIMP_TAC (bin_mult_ss addrews [zmult_1,zmult_zminus]) 1); +by (etac boolE 1); +by (ASM_SIMP_TAC (bin_mult_ss addrews [zadd_zmult_distrib]) 2); +by (ASM_SIMP_TAC + (bin_mult_ss addrews [zadd_zmult_distrib, zmult_1, zadd_assoc]) 1); +by (REPEAT (ares_tac ([zadd_commute, zadd_assoc_swap_kill RS sym]@ + typechecks) 1)); +val integ_of_bin_mult = result(); + +(**** Computations ****) + +(** extra rules for bin_succ, bin_pred **) + +val [bin_succ_Plus, bin_succ_Minus, _] = bin_recs bin_succ_def; +val [bin_pred_Plus, bin_pred_Minus, _] = bin_recs bin_pred_def; + +goal BinFn.thy "bin_succ(w$$1) = bin_succ(w) $$ 0"; +by (SIMP_TAC carry_ss 1); +val bin_succ_Bcons1 = result(); + +goal BinFn.thy "bin_succ(w$$0) = w$$1"; +by (SIMP_TAC carry_ss 1); +val bin_succ_Bcons0 = result(); + +goal BinFn.thy "bin_pred(w$$1) = w$$0"; +by (SIMP_TAC carry_ss 1); +val bin_pred_Bcons1 = result(); + +goal BinFn.thy "bin_pred(w$$0) = bin_pred(w) $$ 1"; +by (SIMP_TAC carry_ss 1); +val bin_pred_Bcons0 = result(); + +(** extra rules for bin_minus **) + +val [bin_minus_Plus, bin_minus_Minus, _] = bin_recs bin_minus_def; + +goal BinFn.thy "bin_minus(w$$1) = bin_pred(bin_minus(w) $$ 0)"; +by (SIMP_TAC bin_minus_ss 1); +val bin_minus_Bcons1 = result(); + +goal BinFn.thy "bin_minus(w$$0) = bin_minus(w) $$ 0"; +by (SIMP_TAC bin_minus_ss 1); +val bin_minus_Bcons0 = result(); + +(** extra rules for bin_add **) + +goal BinFn.thy + "!!w. w: bin ==> bin_add(v$$1, w$$1) = bin_add(v, bin_succ(w)) $$ 0"; +by (ASM_SIMP_TAC bin_add_ss 1); +val bin_add_Bcons_Bcons11 = result(); + +goal BinFn.thy + "!!w. w: bin ==> bin_add(v$$1, w$$0) = bin_add(v,w) $$ 1"; +by (ASM_SIMP_TAC bin_add_ss 1); +val bin_add_Bcons_Bcons10 = result(); + +goal BinFn.thy + "!!w y.[| w: bin; y: bool |] ==> bin_add(v$$0, w$$y) = bin_add(v,w) $$ y"; +by (ASM_SIMP_TAC bin_add_ss 1); +val bin_add_Bcons_Bcons0 = result(); + +(** extra rules for bin_mult **) + +val [bin_mult_Plus, bin_mult_Minus, _] = bin_recs bin_mult_def; + +goal BinFn.thy "bin_mult(v$$1, w) = bin_add(bin_mult(v,w)$$0, w)"; +by (SIMP_TAC bin_mult_ss 1); +val bin_mult_Bcons1 = result(); + +goal BinFn.thy "bin_mult(v$$0, w) = bin_mult(v,w)$$0"; +by (SIMP_TAC bin_mult_ss 1); +val bin_mult_Bcons0 = result(); + + +(*** The computation simpset ***) + +val bin_comp_ss = carry_ss addrews + [bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, bin_add_Bcons_Minus, + bin_add_Bcons_Bcons0, bin_add_Bcons_Bcons10, bin_add_Bcons_Bcons11] + setauto (type_auto_tac bin_typechecks0); + +val bin_comp_ss = integ_ss + addcongs bin_congs + addrews [bin_succ_Plus, bin_succ_Minus, + bin_succ_Bcons1, bin_succ_Bcons0, + bin_pred_Plus, bin_pred_Minus, + bin_pred_Bcons1, bin_pred_Bcons0, + bin_minus_Plus, bin_minus_Minus, + bin_minus_Bcons1, bin_minus_Bcons0, + bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, + bin_add_Bcons_Minus, bin_add_Bcons_Bcons0, + bin_add_Bcons_Bcons10, bin_add_Bcons_Bcons11, + bin_mult_Plus, bin_mult_Minus, + bin_mult_Bcons1, bin_mult_Bcons0] + setauto (type_auto_tac ([bool_1I, bool_0I] @ bin_typechecks0)); + +(*** Examples of performing binary arithmetic by simplification ***) + +proof_timing := true; +(*All runtimes below are on a SPARCserver 10*) + +(* 13+19 = 32 *) +goal BinFn.thy + "bin_add(Plus$$1$$1$$0$$1, Plus$$1$$0$$0$$1$$1) = Plus$$1$$0$$0$$0$$0$$0"; +by (SIMP_TAC bin_comp_ss 1); (*1.9 secs*) +result(); + +bin_add(binary_of_int 13, binary_of_int 19); + +(* 1234+5678 = 6912 *) +goal BinFn.thy + "bin_add(Plus$$1$$0$$0$$1$$1$$0$$1$$0$$0$$1$$0, \ +\ Plus$$1$$0$$1$$1$$0$$0$$0$$1$$0$$1$$1$$1$$0) = \ +\ Plus$$1$$1$$0$$1$$1$$0$$0$$0$$0$$0$$0$$0$$0"; +by (SIMP_TAC bin_comp_ss 1); (*8.9 secs*) +result(); + +bin_add(binary_of_int 1234, binary_of_int 5678); + +(* 1359-2468 = ~1109 *) +goal BinFn.thy + "bin_add(Plus$$1$$0$$1$$0$$1$$0$$0$$1$$1$$1$$1, \ +\ Minus$$0$$1$$1$$0$$0$$1$$0$$1$$1$$1$$0$$0) = \ +\ Minus$$1$$0$$1$$1$$1$$0$$1$$0$$1$$0$$1$$1"; +by (SIMP_TAC bin_comp_ss 1); (*7.4 secs*) +result(); + +bin_add(binary_of_int 1359, binary_of_int ~2468); + +(* 93746-46375 = 47371 *) +goal BinFn.thy + "bin_add(Plus$$1$$0$$1$$1$$0$$1$$1$$1$$0$$0$$0$$1$$1$$0$$0$$1$$0, \ +\ Minus$$0$$1$$0$$0$$1$$0$$1$$0$$1$$1$$0$$1$$1$$0$$0$$1) = \ +\ Plus$$0$$1$$0$$1$$1$$1$$0$$0$$1$$0$$0$$0$$0$$1$$0$$1$$1"; +by (SIMP_TAC bin_comp_ss 1); (*13.7 secs*) +result(); + +bin_add(binary_of_int 93746, binary_of_int ~46375); + +(* negation of 65745 *) +goal BinFn.thy + "bin_minus(Plus$$1$$0$$0$$0$$0$$0$$0$$0$$0$$1$$1$$0$$1$$0$$0$$0$$1) = \ +\ Minus$$0$$1$$1$$1$$1$$1$$1$$1$$1$$0$$0$$1$$0$$1$$1$$1$$1"; +by (SIMP_TAC bin_comp_ss 1); (*2.8 secs*) +result(); + +bin_minus(binary_of_int 65745); + +(* negation of ~54321 *) +goal BinFn.thy + "bin_minus(Minus$$0$$0$$1$$0$$1$$0$$1$$1$$1$$1$$0$$0$$1$$1$$1$$1) = \ +\ Plus$$0$$1$$1$$0$$1$$0$$1$$0$$0$$0$$0$$1$$1$$0$$0$$0$$1"; +by (SIMP_TAC bin_comp_ss 1); (*3.3 secs*) +result(); + +bin_minus(binary_of_int ~54321); + +(* 13*19 = 247 *) +goal BinFn.thy "bin_mult(Plus$$1$$1$$0$$1, Plus$$1$$0$$0$$1$$1) = \ +\ Plus$$1$$1$$1$$1$$0$$1$$1$$1"; +by (SIMP_TAC bin_comp_ss 1); (*4.4 secs*) +result(); + +bin_mult(binary_of_int 13, binary_of_int 19); + +(* ~84 * 51 = ~4284 *) +goal BinFn.thy + "bin_mult(Minus$$0$$1$$0$$1$$1$$0$$0, Plus$$1$$1$$0$$0$$1$$1) = \ +\ Minus$$0$$1$$1$$1$$1$$0$$1$$0$$0$$0$$1$$0$$0"; +by (SIMP_TAC bin_comp_ss 1); (*9.2 secs*) +result(); + +bin_mult(binary_of_int ~84, binary_of_int 51); + +(* 255*255 = 65025; the worst case for 8-bit operands *) +goal BinFn.thy + "bin_mult(Plus$$1$$1$$1$$1$$1$$1$$1$$1, \ +\ Plus$$1$$1$$1$$1$$1$$1$$1$$1) = \ +\ Plus$$1$$1$$1$$1$$1$$1$$1$$0$$0$$0$$0$$0$$0$$0$$0$$1"; +by (SIMP_TAC bin_comp_ss 1); (*38.4 secs*) +result(); + +bin_mult(binary_of_int 255, binary_of_int 255); + +(***************** TOO SLOW TO INCLUDE IN TEST RUNS +(* 1359 * ~2468 = ~3354012 *) +goal BinFn.thy + "bin_mult(Plus$$1$$0$$1$$0$$1$$0$$0$$1$$1$$1$$1, \ +\ Minus$$0$$1$$1$$0$$0$$1$$0$$1$$1$$1$$0$$0) = \ +\ Minus$$0$$0$$1$$1$$0$$0$$1$$1$$0$$1$$0$$0$$1$$0$$0$$1$$1$$0$$0$$1$$0$$0"; +by (SIMP_TAC bin_comp_ss 1); (*54.8 secs*) +result(); +****************) + +bin_mult(binary_of_int 1359, binary_of_int ~2468); diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/binfn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/binfn.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,52 @@ +(* Title: ZF/bin + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Arithmetic on binary integers. +*) + +BinFn = Integ + Bin + +consts + bin_rec :: "[i, i, i, [i,i,i]=>i] => i" + integ_of_bin :: "i=>i" + bin_succ :: "i=>i" + bin_pred :: "i=>i" + bin_minus :: "i=>i" + bin_add,bin_mult :: "[i,i]=>i" + +rules + + bin_rec_def + "bin_rec(z,a,b,h) == \ +\ Vrec(z, %z g. bin_case(a, b, %w x. h(w, x, g`w), z))" + + integ_of_bin_def + "integ_of_bin(w) == bin_rec(w, $#0, $~($#1), %w x r. $#x $+ r $+ r)" + + bin_succ_def + "bin_succ(w0) == bin_rec(w0, Plus$$1, Plus, %w x r. cond(x, r$$0, w$$1))" + + bin_pred_def + "bin_pred(w0) == \ +\ bin_rec(w0, Minus, Minus$$0, %w x r. cond(x, w$$0, r$$1))" + + bin_minus_def + "bin_minus(w0) == \ +\ bin_rec(w0, Plus, Plus$$1, %w x r. cond(x, bin_pred(r$$0), r$$0))" + + bin_add_def + "bin_add(v0,w0) == \ +\ bin_rec(v0, \ +\ lam w:bin. w, \ +\ lam w:bin. bin_pred(w), \ +\ %v x r. lam w1:bin. \ +\ bin_rec(w1, v$$x, bin_pred(v$$x), \ +\ %w y s. (r`cond(x and y, bin_succ(w), w)) \ +\ $$ (x xor y))) ` w0" + + bin_mult_def + "bin_mult(v0,w) == \ +\ bin_rec(v0, Plus, bin_minus(w), \ +\ %v x r. cond(x, bin_add(r$$0,w), r$$0))" +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/bt.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/bt.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,49 @@ +(* Title: ZF/ex/bt.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Datatype definition of binary trees +*) + +structure BT = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("bt", "univ(A)", + [(["Lf"],"i"), (["Br"],"[i,i,i]=>i")])]; + val rec_styp = "i=>i"; + val ext = None + val sintrs = + ["Lf : bt(A)", + "[| a: A; t1: bt(A); t2: bt(A) |] ==> Br(a,t1,t2) : bt(A)"]; + val monos = []; + val type_intrs = data_typechecks + val type_elims = []); + +val [LfI, BrI] = BT.intrs; + +(*Perform induction on l, then prove the major premise using prems. *) +fun bt_ind_tac a prems i = + EVERY [res_inst_tac [("x",a)] BT.induct i, + rename_last_tac a ["1","2"] (i+2), + ares_tac prems i]; + + +(** Lemmas to justify using "bt" in other recursive type definitions **) + +goalw BT.thy BT.defs "!!A B. A<=B ==> bt(A) <= bt(B)"; +by (rtac lfp_mono 1); +by (REPEAT (rtac BT.bnd_mono 1)); +by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); +val bt_mono = result(); + +goalw BT.thy (BT.defs@BT.con_defs) "bt(univ(A)) <= univ(A)"; +by (rtac lfp_lowerbound 1); +by (rtac (A_subset_univ RS univ_mono) 2); +by (fast_tac (ZF_cs addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ, + Pair_in_univ]) 1); +val bt_univ = result(); + +val bt_subset_univ = standard (bt_mono RS (bt_univ RSN (2,subset_trans))); + + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/bt_fn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/bt_fn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,138 @@ +(* Title: ZF/bt.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For bt.thy. Binary trees +*) + +open BT_Fn; + + + +(** bt_rec -- by Vset recursion **) + +(*Used to verify bt_rec*) +val bt_rec_ss = ZF_ss + addcongs (mk_typed_congs BT_Fn.thy [("h", "[i,i,i,i,i]=>i")]) + addrews BT.case_eqns; + +goalw BT.thy BT.con_defs "rank(l) : rank(Br(a,l,r))"; +by (SIMP_TAC rank_ss 1); +val rank_Br1 = result(); + +goalw BT.thy BT.con_defs "rank(r) : rank(Br(a,l,r))"; +by (SIMP_TAC rank_ss 1); +val rank_Br2 = result(); + +goal BT_Fn.thy "bt_rec(Lf,c,h) = c"; +by (rtac (bt_rec_def RS def_Vrec RS trans) 1); +by (SIMP_TAC bt_rec_ss 1); +val bt_rec_Lf = result(); + +goal BT_Fn.thy + "bt_rec(Br(a,l,r), c, h) = h(a, l, r, bt_rec(l,c,h), bt_rec(r,c,h))"; +by (rtac (bt_rec_def RS def_Vrec RS trans) 1); +by (SIMP_TAC (bt_rec_ss addrews [Vset_rankI, rank_Br1, rank_Br2]) 1); +val bt_rec_Br = result(); + +(*Type checking -- proved by induction, as usual*) +val prems = goal BT_Fn.thy + "[| t: bt(A); \ +\ c: C(Lf); \ +\ !!x y z r s. [| x:A; y:bt(A); z:bt(A); r:C(y); s:C(z) |] ==> \ +\ h(x,y,z,r,s): C(Br(x,y,z)) \ +\ |] ==> bt_rec(t,c,h) : C(t)"; +by (bt_ind_tac "t" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (ZF_ss addrews + (prems@[bt_rec_Lf,bt_rec_Br])))); +val bt_rec_type = result(); + +(** Versions for use with definitions **) + +val [rew] = goal BT_Fn.thy "[| !!t. j(t)==bt_rec(t, c, h) |] ==> j(Lf) = c"; +by (rewtac rew); +by (rtac bt_rec_Lf 1); +val def_bt_rec_Lf = result(); + +val [rew] = goal BT_Fn.thy + "[| !!t. j(t)==bt_rec(t, c, h) |] ==> j(Br(a,l,r)) = h(a,l,r,j(l),j(r))"; +by (rewtac rew); +by (rtac bt_rec_Br 1); +val def_bt_rec_Br = result(); + +fun bt_recs def = map standard ([def] RL [def_bt_rec_Lf, def_bt_rec_Br]); + +(** n_nodes **) + +val [n_nodes_Lf,n_nodes_Br] = bt_recs n_nodes_def; + +val prems = goalw BT_Fn.thy [n_nodes_def] + "xs: bt(A) ==> n_nodes(xs) : nat"; +by (REPEAT (ares_tac (prems @ [bt_rec_type, nat_0I, nat_succI, add_type]) 1)); +val n_nodes_type = result(); + + +(** n_leaves **) + +val [n_leaves_Lf,n_leaves_Br] = bt_recs n_leaves_def; + +val prems = goalw BT_Fn.thy [n_leaves_def] + "xs: bt(A) ==> n_leaves(xs) : nat"; +by (REPEAT (ares_tac (prems @ [bt_rec_type, nat_0I, nat_succI, add_type]) 1)); +val n_leaves_type = result(); + +(** bt_reflect **) + +val [bt_reflect_Lf, bt_reflect_Br] = bt_recs bt_reflect_def; + +val prems = goalw BT_Fn.thy [bt_reflect_def] + "xs: bt(A) ==> bt_reflect(xs) : bt(A)"; +by (REPEAT (ares_tac (prems @ [bt_rec_type, LfI, BrI]) 1)); +val bt_reflect_type = result(); + + +(** BT_Fn simplification **) + + +val bt_typechecks = + [LfI, BrI, bt_rec_type, n_nodes_type, n_leaves_type, bt_reflect_type]; + +val bt_congs = + BT.congs @ + mk_congs BT_Fn.thy ["bt_case","bt_rec","n_nodes","n_leaves","bt_reflect"]; + +val bt_ss = arith_ss + addcongs bt_congs + addrews BT.case_eqns + addrews bt_typechecks + addrews [bt_rec_Lf, bt_rec_Br, + n_nodes_Lf, n_nodes_Br, + n_leaves_Lf, n_leaves_Br, + bt_reflect_Lf, bt_reflect_Br]; + + +(*** theorems about n_leaves ***) + +val prems = goal BT_Fn.thy + "t: bt(A) ==> n_leaves(bt_reflect(t)) = n_leaves(t)"; +by (bt_ind_tac "t" prems 1); +by (ALLGOALS (ASM_SIMP_TAC bt_ss)); +by (REPEAT (ares_tac [add_commute, n_leaves_type] 1)); +val n_leaves_reflect = result(); + +val prems = goal BT_Fn.thy + "t: bt(A) ==> n_leaves(t) = succ(n_nodes(t))"; +by (bt_ind_tac "t" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (bt_ss addrews [add_succ_right]))); +val n_leaves_nodes = result(); + +(*** theorems about bt_reflect ***) + +val prems = goal BT_Fn.thy + "t: bt(A) ==> bt_reflect(bt_reflect(t))=t"; +by (bt_ind_tac "t" prems 1); +by (ALLGOALS (ASM_SIMP_TAC bt_ss)); +val bt_reflect_bt_reflect_ident = result(); + + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/bt_fn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/bt_fn.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,24 @@ +(* Title: ZF/ex/bt-fn.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Binary trees +*) + +BT_Fn = BT + +consts + bt_rec :: "[i, i, [i,i,i,i,i]=>i] => i" + n_nodes :: "i=>i" + n_leaves :: "i=>i" + bt_reflect :: "i=>i" + +rules + bt_rec_def + "bt_rec(t,c,h) == Vrec(t, %t g.bt_case(c, %x y z. h(x,y,z,g`y,g`z), t))" + + n_nodes_def "n_nodes(t) == bt_rec(t, 0, %x y z r s. succ(r#+s))" + n_leaves_def "n_leaves(t) == bt_rec(t, succ(0), %x y z r s. r#+s)" + bt_reflect_def "bt_reflect(t) == bt_rec(t, Lf, %x y z r s. Br(x,s,r))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/comb.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/comb.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,45 @@ +(* Title: ZF/ex/comb.ML + ID: $Id$ + Author: Lawrence C Paulson + Copyright 1993 University of Cambridge + +Datatype definition of combinators S and K + +J. Camilleri and T. F. Melham. +Reasoning with Inductively Defined Relations in the HOL Theorem Prover. +Report 265, University of Cambridge Computer Laboratory, 1992. +*) + + +(*Example of a datatype with mixfix syntax for some constructors*) +structure Comb = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("comb", "univ(0)", + [(["K","S"], "i"), + (["op #"], "[i,i]=>i")])]; + val rec_styp = "i"; + val ext = Some (NewSext { + mixfix = + [Infixl("#", "[i,i] => i", 90)], + xrules = [], + parse_ast_translation = [], + parse_preproc = None, + parse_postproc = None, + parse_translation = [], + print_translation = [], + print_preproc = None, + print_postproc = None, + print_ast_translation = []}); + val sintrs = + ["K : comb", + "S : comb", + "[| p: comb; q: comb |] ==> p#q : comb"]; + val monos = []; + val type_intrs = data_typechecks; + val type_elims = []); + +val [K_comb,S_comb,Ap_comb] = Comb.intrs; + +val Ap_E = Comb.mk_cases Comb.con_defs "p#q : comb"; + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/contract0.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/contract0.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,117 @@ +(* Title: ZF/ex/contract.ML + ID: $Id$ + Author: Tobias Nipkow & Lawrence C Paulson + Copyright 1992 University of Cambridge + +For ex/contract.thy. +*) + +open Contract0; + +structure Contract = Inductive_Fun + (val thy = Contract0.thy; + val rec_doms = [("contract","comb*comb")]; + val sintrs = + ["[| p:comb; q:comb |] ==> K#p#q -1-> p", + "[| p:comb; q:comb; r:comb |] ==> S#p#q#r -1-> (p#r)#(q#r)", + "[| p-1->q; r:comb |] ==> p#r -1-> q#r", + "[| p-1->q; r:comb |] ==> r#p -1-> r#q"]; + val monos = []; + val con_defs = []; + val type_intrs = Comb.intrs@[SigmaI]; + val type_elims = [SigmaE2]); + +val [K_contract,S_contract,Ap_contract1,Ap_contract2] = Contract.intrs; + +val contract_induct = standard + (Contract.mutual_induct RS spec RS spec RSN (2,rev_mp)); + +(*For type checking: replaces a-1->b by a,b:comb *) +val contract_combE2 = Contract.dom_subset RS subsetD RS SigmaE2; +val contract_combD1 = Contract.dom_subset RS subsetD RS SigmaD1; +val contract_combD2 = Contract.dom_subset RS subsetD RS SigmaD2; + +goal Contract.thy "field(contract) = comb"; +by (fast_tac (ZF_cs addIs [equalityI,K_contract] addSEs [contract_combE2]) 1); +val field_contract_eq = result(); + +val reduction_refl = standard + (field_contract_eq RS equalityD2 RS subsetD RS rtrancl_refl); + +val rtrancl_into_rtrancl2 = standard + (r_into_rtrancl RS (trans_rtrancl RS transD)); + +val reduction_rls = [reduction_refl, K_contract, S_contract, + K_contract RS rtrancl_into_rtrancl2, + S_contract RS rtrancl_into_rtrancl2, + Ap_contract1 RS rtrancl_into_rtrancl2, + Ap_contract2 RS rtrancl_into_rtrancl2]; + +goalw Contract.thy [I_def] "!!p. p:comb ==> I#p ---> p"; +by (REPEAT (ares_tac (Comb.intrs @ reduction_rls) 1)); +val I_reduce = result(); + +goalw Contract.thy [I_def] "I: comb"; +by (REPEAT (ares_tac Comb.intrs 1)); +val I_comb = result(); + +(** Non-contraction results **) + +(*Derive a case for each combinator constructor*) +val K_contract_case = Contract.mk_cases Comb.con_defs "K -1-> r"; +val S_contract_case = Contract.mk_cases Comb.con_defs "S -1-> r"; +val Ap_contract_case = Contract.mk_cases Comb.con_defs "p#q -1-> r"; + +val contract_cs = + ZF_cs addSIs Comb.intrs + addIs Contract.intrs + addSEs [contract_combD1,contract_combD2] (*type checking*) + addSEs [K_contract_case, S_contract_case, Ap_contract_case] + addSEs Comb.free_SEs; + +goalw Contract.thy [I_def] "!!r. I -1-> r ==> P"; +by (fast_tac contract_cs 1); +val I_contract_case = result(); + +goal Contract.thy "!!p r. K#p -1-> r ==> (EX q. r = K#q & p -1-> q)"; +by (fast_tac contract_cs 1); +val K1_contractD = result(); + +goal Contract.thy "!!p r. [| p ---> q; r: comb |] ==> p#r ---> q#r"; +by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1); +by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1); +by (etac rtrancl_induct 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +by (etac (trans_rtrancl RS transD) 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +val Ap_reduce1 = result(); + +goal Contract.thy "!!p r. [| p ---> q; r: comb |] ==> r#p ---> r#q"; +by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1); +by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1); +by (etac rtrancl_induct 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +by (etac (trans_rtrancl RS transD) 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +val Ap_reduce2 = result(); + +(** Counterexample to the diamond property for -1-> **) + +goal Contract.thy "K#I#(I#I) -1-> I"; +by (REPEAT (ares_tac [K_contract, I_comb, Ap_comb] 1)); +val KIII_contract1 = result(); + +goalw Contract.thy [I_def] "K#I#(I#I) -1-> K#I#((K#I)#(K#I))"; +by (DEPTH_SOLVE (resolve_tac (Comb.intrs @ Contract.intrs) 1)); +val KIII_contract2 = result(); + +goal Contract.thy "K#I#((K#I)#(K#I)) -1-> I"; +by (REPEAT (ares_tac (Comb.intrs @ [K_contract, I_comb]) 1)); +val KIII_contract3 = result(); + +goalw Contract.thy [diamond_def] "~ diamond(contract)"; +by (fast_tac (ZF_cs addIs [KIII_contract1,KIII_contract2,KIII_contract3] + addSEs [I_contract_case]) 1); +val not_diamond_contract = result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/contract0.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/contract0.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,36 @@ +(* Title: ZF/ex/contract.thy + ID: $Id$ + Author: Tobias Nipkow & Lawrence C Paulson + Copyright 1993 University of Cambridge + +Inductive definition of (1-step) contractions and (mult-step) reductions +*) + +Contract0 = Comb + +consts + diamond :: "i => o" + I :: "i" + + contract :: "i" + "-1->" :: "[i,i] => o" (infixl 50) + "--->" :: "[i,i] => o" (infixl 50) + + parcontract :: "i" + "=1=>" :: "[i,i] => o" (infixl 50) + "===>" :: "[i,i] => o" (infixl 50) + +translations + "p -1-> q" == " : contract" + "p ---> q" == " : contract^*" + "p =1=> q" == " : parcontract" + "p ===> q" == " : parcontract^+" + +rules + + diamond_def "diamond(r) == ALL x y. :r --> \ +\ (ALL y'. :r --> \ +\ (EX z. :r & : r))" + + I_def "I == S#K#K" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/enum.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/enum.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,33 @@ +(* Title: ZF/ex/enum + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Example of a BIG enumeration type + +Can go up to at least 100 constructors, but it takes over 10 minutes... +*) + + +(*An enumeration type with 60 contructors! -- takes about 214 seconds!*) +fun mk_ids a 0 = [] + | mk_ids a n = a :: mk_ids (bump_string a) (n-1); + +val consts = mk_ids "con1" 60; + +structure Enum = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("enum", "univ(0)", + [(consts, "i")])]; + val rec_styp = "i"; + val ext = None + val sintrs = map (fn const => const ^ " : enum") consts; + val monos = []; + val type_intrs = data_typechecks + val type_elims = []); + +goal Enum.thy "~ con59=con60"; +by (SIMP_TAC (ZF_ss addrews Enum.free_iffs) 1); (*2.3 secs*) +result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/equiv.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/equiv.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,268 @@ +(* Title: ZF/ex/equiv.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For equiv.thy. Equivalence relations in Zermelo-Fraenkel Set Theory +*) + +val RSLIST = curry (op MRS); + +open Equiv; + +(*** Suppes, Theorem 70: r is an equiv relation iff converse(r) O r = r ***) + +(** first half: equiv(A,r) ==> converse(r) O r = r **) + +goalw Equiv.thy [trans_def,sym_def] + "!!r. [| sym(r); trans(r) |] ==> converse(r) O r <= r"; +by (fast_tac (ZF_cs addSEs [converseD,compE]) 1); +val sym_trans_comp_subset = result(); + +goalw Equiv.thy [refl_def] + "!!A r. refl(A,r) ==> r <= converse(r) O r"; +by (fast_tac (ZF_cs addSIs [converseI] addIs [compI]) 1); +val refl_comp_subset = result(); + +goalw Equiv.thy [equiv_def] + "!!A r. equiv(A,r) ==> converse(r) O r = r"; +by (rtac equalityI 1); +by (REPEAT (ares_tac [sym_trans_comp_subset, refl_comp_subset] 1 + ORELSE etac conjE 1)); +val equiv_comp_eq = result(); + +(*second half*) +goalw Equiv.thy [equiv_def,refl_def,sym_def,trans_def] + "!!A r. [| converse(r) O r = r; domain(r) = A |] ==> equiv(A,r)"; +by (etac equalityE 1); +by (subgoal_tac "ALL x y. : r --> : r" 1); +by (safe_tac ZF_cs); +by (fast_tac (ZF_cs addSIs [converseI] addIs [compI]) 3); +by (ALLGOALS (fast_tac + (ZF_cs addSIs [converseI] addIs [compI] addSEs [compE]))); +by flexflex_tac; +val comp_equivI = result(); + +(** Equivalence classes **) + +(*Lemma for the next result*) +goalw Equiv.thy [equiv_def,trans_def,sym_def] + "!!A r. [| equiv(A,r); : r |] ==> r``{a} <= r``{b}"; +by (fast_tac ZF_cs 1); +val equiv_class_subset = result(); + +goal Equiv.thy "!!A r. [| equiv(A,r); : r |] ==> r``{a} = r``{b}"; +by (REPEAT (ares_tac [equalityI, equiv_class_subset] 1)); +by (rewrite_goals_tac [equiv_def,sym_def]); +by (fast_tac ZF_cs 1); +val equiv_class_eq = result(); + +val prems = goalw Equiv.thy [equiv_def,refl_def] + "[| equiv(A,r); a: A |] ==> a: r``{a}"; +by (cut_facts_tac prems 1); +by (fast_tac ZF_cs 1); +val equiv_class_self = result(); + +(*Lemma for the next result*) +goalw Equiv.thy [equiv_def,refl_def] + "!!A r. [| equiv(A,r); r``{b} <= r``{a}; b: A |] ==> : r"; +by (fast_tac ZF_cs 1); +val subset_equiv_class = result(); + +val prems = goal Equiv.thy + "[| r``{a} = r``{b}; equiv(A,r); b: A |] ==> : r"; +by (REPEAT (resolve_tac (prems @ [equalityD2, subset_equiv_class]) 1)); +val eq_equiv_class = result(); + +(*thus r``{a} = r``{b} as well*) +goalw Equiv.thy [equiv_def,trans_def,sym_def] + "!!A r. [| equiv(A,r); x: (r``{a} Int r``{b}) |] ==> : r"; +by (fast_tac ZF_cs 1); +val equiv_class_nondisjoint = result(); + +val [major] = goalw Equiv.thy [equiv_def,refl_def] + "equiv(A,r) ==> r <= A*A"; +by (rtac (major RS conjunct1 RS conjunct1) 1); +val equiv_type = result(); + +goal Equiv.thy + "!!A r. equiv(A,r) ==> : r <-> r``{x} = r``{y} & x:A & y:A"; +by (fast_tac (ZF_cs addIs [eq_equiv_class, equiv_class_eq] + addDs [equiv_type]) 1); +val equiv_class_eq_iff = result(); + +goal Equiv.thy + "!!A r. [| equiv(A,r); x: A; y: A |] ==> r``{x} = r``{y} <-> : r"; +by (fast_tac (ZF_cs addIs [eq_equiv_class, equiv_class_eq] + addDs [equiv_type]) 1); +val eq_equiv_class_iff = result(); + +(*** Quotients ***) + +(** Introduction/elimination rules -- needed? **) + +val prems = goalw Equiv.thy [quotient_def] "x:A ==> r``{x}: A/r"; +by (rtac RepFunI 1); +by (resolve_tac prems 1); +val quotientI = result(); + +val major::prems = goalw Equiv.thy [quotient_def] + "[| X: A/r; !!x. [| X = r``{x}; x:A |] ==> P |] \ +\ ==> P"; +by (rtac (major RS RepFunE) 1); +by (eresolve_tac prems 1); +by (assume_tac 1); +val quotientE = result(); + +goalw Equiv.thy [equiv_def,refl_def,quotient_def] + "!!A r. equiv(A,r) ==> Union(A/r) = A"; +by (fast_tac eq_cs 1); +val Union_quotient = result(); + +goalw Equiv.thy [quotient_def] + "!!A r. [| equiv(A,r); X: A/r; Y: A/r |] ==> X=Y | (X Int Y <= 0)"; +by (safe_tac (ZF_cs addSIs [equiv_class_eq])); +by (assume_tac 1); +by (rewrite_goals_tac [equiv_def,trans_def,sym_def]); +by (fast_tac ZF_cs 1); +val quotient_disj = result(); + +(**** Defining unary operations upon equivalence classes ****) + +(** These proofs really require as local premises + equiv(A,r); congruent(r,b) +**) + +(*Conversion rule*) +val prems as [equivA,bcong,_] = goal Equiv.thy + "[| equiv(A,r); congruent(r,b); a: A |] ==> (UN x:r``{a}. b(x)) = b(a)"; +by (cut_facts_tac prems 1); +by (rtac UN_singleton 1); +by (etac equiv_class_self 1); +by (assume_tac 1); +by (rewrite_goals_tac [equiv_def,sym_def,congruent_def]); +by (fast_tac ZF_cs 1); +val UN_equiv_class = result(); + +(*Resolve th against the "local" premises*) +val localize = RSLIST [equivA,bcong]; + +(*type checking of UN x:r``{a}. b(x) *) +val _::_::prems = goalw Equiv.thy [quotient_def] + "[| equiv(A,r); congruent(r,b); X: A/r; \ +\ !!x. x : A ==> b(x) : B |] \ +\ ==> (UN x:X. b(x)) : B"; +by (cut_facts_tac prems 1); +by (safe_tac ZF_cs); +by (rtac (localize UN_equiv_class RS ssubst) 1); +by (REPEAT (ares_tac prems 1)); +val UN_equiv_class_type = result(); + +(*Sufficient conditions for injectiveness. Could weaken premises! + major premise could be an inclusion; bcong could be !!y. y:A ==> b(y):B +*) +val _::_::prems = goalw Equiv.thy [quotient_def] + "[| equiv(A,r); congruent(r,b); \ +\ (UN x:X. b(x))=(UN y:Y. b(y)); X: A/r; Y: A/r; \ +\ !!x y. [| x:A; y:A; b(x)=b(y) |] ==> :r |] \ +\ ==> X=Y"; +by (cut_facts_tac prems 1); +by (safe_tac ZF_cs); +by (rtac (equivA RS equiv_class_eq) 1); +by (REPEAT (ares_tac prems 1)); +by (etac box_equals 1); +by (REPEAT (ares_tac [localize UN_equiv_class] 1)); +val UN_equiv_class_inject = result(); + + +(**** Defining binary operations upon equivalence classes ****) + + +goalw Equiv.thy [congruent_def,congruent2_def,equiv_def,refl_def] + "!!A r. [| equiv(A,r); congruent2(r,b); a: A |] ==> congruent(r,b(a))"; +by (fast_tac ZF_cs 1); +val congruent2_implies_congruent = result(); + +val equivA::prems = goalw Equiv.thy [congruent_def] + "[| equiv(A,r); congruent2(r,b); a: A |] ==> \ +\ congruent(r, %x1. UN x2:r``{a}. b(x1,x2))"; +by (cut_facts_tac (equivA::prems) 1); +by (safe_tac ZF_cs); +by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1); +by (assume_tac 1); +by (ASM_SIMP_TAC (ZF_ss addrews [equivA RS UN_equiv_class, + congruent2_implies_congruent]) 1); +by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]); +by (fast_tac ZF_cs 1); +val congruent2_implies_congruent_UN = result(); + +val prems as equivA::_ = goal Equiv.thy + "[| equiv(A,r); congruent2(r,b); a1: A; a2: A |] \ +\ ==> (UN x1:r``{a1}. UN x2:r``{a2}. b(x1,x2)) = b(a1,a2)"; +by (cut_facts_tac prems 1); +by (ASM_SIMP_TAC (ZF_ss addrews [equivA RS UN_equiv_class, + congruent2_implies_congruent, + congruent2_implies_congruent_UN]) 1); +val UN_equiv_class2 = result(); + +(*type checking*) +val prems = goalw Equiv.thy [quotient_def] + "[| equiv(A,r); congruent2(r,b); \ +\ X1: A/r; X2: A/r; \ +\ !!x1 x2. [| x1: A; x2: A |] ==> b(x1,x2) : B |] \ +\ ==> (UN x1:X1. UN x2:X2. b(x1,x2)) : B"; +by (cut_facts_tac prems 1); +by (safe_tac ZF_cs); +by (REPEAT (ares_tac (prems@[UN_equiv_class_type, + congruent2_implies_congruent_UN, + congruent2_implies_congruent, quotientI]) 1)); +val UN_equiv_class_type2 = result(); + + +(*Suggested by John Harrison -- the two subproofs may be MUCH simpler + than the direct proof*) +val prems = goalw Equiv.thy [congruent2_def,equiv_def,refl_def] + "[| equiv(A,r); \ +\ !! y z w. [| w: A; : r |] ==> b(y,w) = b(z,w); \ +\ !! y z w. [| w: A; : r |] ==> b(w,y) = b(w,z) \ +\ |] ==> congruent2(r,b)"; +by (cut_facts_tac prems 1); +by (safe_tac ZF_cs); +by (rtac trans 1); +by (REPEAT (ares_tac prems 1 + ORELSE etac (subsetD RS SigmaE2) 1 THEN assume_tac 2 THEN assume_tac 1)); +val congruent2I = result(); + +val [equivA,commute,congt] = goal Equiv.thy + "[| equiv(A,r); \ +\ !! y z w. [| y: A; z: A |] ==> b(y,z) = b(z,y); \ +\ !! y z w. [| w: A; : r |] ==> b(w,y) = b(w,z) \ +\ |] ==> congruent2(r,b)"; +by (resolve_tac [equivA RS congruent2I] 1); +by (rtac (commute RS trans) 1); +by (rtac (commute RS trans RS sym) 3); +by (rtac sym 5); +by (REPEAT (ares_tac [congt] 1 + ORELSE etac (equivA RS equiv_type RS subsetD RS SigmaE2) 1)); +val congruent2_commuteI = result(); + +(***OBSOLETE VERSION +(*Rules congruentI and congruentD would simplify use of rewriting below*) +val [equivA,ZinA,congt,commute] = goalw Equiv.thy [quotient_def] + "[| equiv(A,r); Z: A/r; \ +\ !!w. [| w: A |] ==> congruent(r, %z.b(w,z)); \ +\ !!x y. [| x: A; y: A |] ==> b(y,x) = b(x,y) \ +\ |] ==> congruent(r, %w. UN z: Z. b(w,z))"; +val congt' = rewrite_rule [congruent_def] congt; +by (cut_facts_tac [ZinA,congt] 1); +by (rewtac congruent_def); +by (safe_tac ZF_cs); +by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1); +by (assume_tac 1); +by (ASM_SIMP_TAC (ZF_ss addrews [congt RS (equivA RS UN_equiv_class)]) 1); +by (rtac (commute RS trans) 1); +by (rtac (commute RS trans RS sym) 3); +by (rtac sym 5); +by (REPEAT (ares_tac [congt' RS spec RS spec RS mp] 1)); +val congruent_commuteI = result(); +***) diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/equiv.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/equiv.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,28 @@ +(* Title: ZF/ex/equiv.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Equivalence relations in Zermelo-Fraenkel Set Theory +*) + +Equiv = Trancl + +consts + refl,equiv :: "[i,i]=>o" + sym :: "i=>o" + "'/" :: "[i,i]=>i" (infixl 90) (*set of equiv classes*) + congruent :: "[i,i=>i]=>o" + congruent2 :: "[i,[i,i]=>i]=>o" + +rules + refl_def "refl(A,r) == r <= (A*A) & (ALL x: A. : r)" + sym_def "sym(r) == ALL x y. : r --> : r" + equiv_def "equiv(A,r) == refl(A,r) & sym(r) & trans(r)" + quotient_def "A/r == {r``{x} . x:A}" + congruent_def "congruent(r,b) == ALL y z. :r --> b(y)=b(z)" + + congruent2_def + "congruent2(r,b) == ALL y1 z1 y2 z2. \ +\ :r --> :r --> b(y1,y2) = b(z1,z2)" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/integ.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/integ.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,437 @@ +(* Title: ZF/ex/integ.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For integ.thy. The integers as equivalence classes over nat*nat. + +Could also prove... +"znegative(z) ==> $# zmagnitude(z) = $~ z" +"~ znegative(z) ==> $# zmagnitude(z) = z" +$< is a linear ordering +$+ and $* are monotonic wrt $< +*) + +open Integ; + +val [add_cong] = mk_congs Arith.thy ["op #+"]; + +(*** Proving that intrel is an equivalence relation ***) + +val prems = goal Arith.thy + "[| m #+ n = m' #+ n'; m: nat; m': nat |] \ +\ ==> m #+ (n #+ k) = m' #+ (n' #+ k)"; +by (ASM_SIMP_TAC (arith_ss addrews ([add_assoc RS sym] @ prems)) 1); +val add_assoc_cong = result(); + +val prems = goal Arith.thy + "[| m: nat; n: nat |] \ +\ ==> m #+ (n #+ k) = n #+ (m #+ k)"; +by (REPEAT (resolve_tac ([add_commute RS add_assoc_cong] @ prems) 1)); +val add_assoc_swap = result(); + +val add_kill = (refl RS add_cong); +val add_assoc_swap_kill = add_kill RSN (3, add_assoc_swap RS trans); + +(*By luck, requires no typing premises for y1, y2,y3*) +val eqa::eqb::prems = goal Arith.thy + "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2; \ +\ x1: nat; x2: nat; x3: nat |] ==> x1 #+ y3 = x3 #+ y1"; +by (res_inst_tac [("k","x2")] add_left_cancel 1); +by (resolve_tac prems 1); +by (rtac (add_assoc_swap RS trans) 1 THEN typechk_tac prems); +by (rtac (eqb RS ssubst) 1); +by (rtac (add_assoc_swap RS trans) 1 THEN typechk_tac prems); +by (rtac (eqa RS ssubst) 1); +by (rtac (add_assoc_swap) 1 THEN typechk_tac prems); +val integ_trans_lemma = result(); + +(** Natural deduction for intrel **) + +val prems = goalw Integ.thy [intrel_def] + "[| x1#+y2 = x2#+y1; x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ +\ <,>: intrel"; +by (fast_tac (ZF_cs addIs prems) 1); +val intrelI = result(); + +(*intrelE is hard to derive because fast_tac tries hyp_subst_tac so soon*) +goalw Integ.thy [intrel_def] + "p: intrel --> (EX x1 y1 x2 y2. \ +\ p = <,> & x1#+y2 = x2#+y1 & \ +\ x1: nat & y1: nat & x2: nat & y2: nat)"; +by (fast_tac ZF_cs 1); +val intrelE_lemma = result(); + +val [major,minor] = goal Integ.thy + "[| p: intrel; \ +\ !!x1 y1 x2 y2. [| p = <,>; x1#+y2 = x2#+y1; \ +\ x1: nat; y1: nat; x2: nat; y2: nat |] ==> Q |] \ +\ ==> Q"; +by (cut_facts_tac [major RS (intrelE_lemma RS mp)] 1); +by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1)); +val intrelE = result(); + +val intrel_cs = ZF_cs addSIs [intrelI] addSEs [intrelE]; + +goal Integ.thy + "<,>: intrel <-> \ +\ x1#+y2 = x2#+y1 & x1: nat & y1: nat & x2: nat & y2: nat"; +by (fast_tac intrel_cs 1); +val intrel_iff = result(); + +val prems = goalw Integ.thy [equiv_def] "equiv(nat*nat, intrel)"; +by (safe_tac intrel_cs); +by (rewtac refl_def); +by (fast_tac intrel_cs 1); +by (rewtac sym_def); +by (fast_tac (intrel_cs addSEs [sym]) 1); +by (rewtac trans_def); +by (fast_tac (intrel_cs addSEs [integ_trans_lemma]) 1); +val equiv_intrel = result(); + + +val integ_congs = mk_congs Integ.thy + ["znat", "zminus", "znegative", "zmagnitude", "op $+", "op $-", "op $*"]; + +val intrel_ss0 = arith_ss addcongs integ_congs; + +val intrel_ss = + intrel_ss0 addrews [equiv_intrel RS eq_equiv_class_iff, intrel_iff]; + +(*More than twice as fast as simplifying with intrel_ss*) +fun INTEG_SIMP_TAC ths = + let val ss = intrel_ss0 addrews ths + in fn i => + EVERY [ASM_SIMP_TAC ss i, + rtac (intrelI RS (equiv_intrel RS equiv_class_eq)) i, + typechk_tac (ZF_typechecks@nat_typechecks@arith_typechecks), + ASM_SIMP_TAC ss i] + end; + + +(** znat: the injection from nat to integ **) + +val prems = goalw Integ.thy [integ_def,quotient_def,znat_def] + "m : nat ==> $#m : integ"; +by (fast_tac (ZF_cs addSIs (nat_0I::prems)) 1); +val znat_type = result(); + +val [major,nnat] = goalw Integ.thy [znat_def] + "[| $#m = $#n; n: nat |] ==> m=n"; +by (rtac (make_elim (major RS eq_equiv_class)) 1); +by (rtac equiv_intrel 1); +by (typechk_tac [nat_0I,nnat,SigmaI]); +by (safe_tac (intrel_cs addSEs [box_equals,add_0_right])); +val znat_inject = result(); + + +(**** zminus: unary negation on integ ****) + +goalw Integ.thy [congruent_def] + "congruent(intrel, split(%x y. intrel``{}))"; +by (safe_tac intrel_cs); +by (ALLGOALS (ASM_SIMP_TAC intrel_ss)); +by (etac (box_equals RS sym) 1); +by (REPEAT (ares_tac [add_commute] 1)); +val zminus_congruent = result(); + +(*Resolve th against the corresponding facts for zminus*) +val zminus_ize = RSLIST [equiv_intrel, zminus_congruent]; + +val [prem] = goalw Integ.thy [integ_def,zminus_def] + "z : integ ==> $~z : integ"; +by (typechk_tac [split_type, SigmaI, prem, zminus_ize UN_equiv_class_type, + quotientI]); +val zminus_type = result(); + +val major::prems = goalw Integ.thy [integ_def,zminus_def] + "[| $~z = $~w; z: integ; w: integ |] ==> z=w"; +by (rtac (major RS zminus_ize UN_equiv_class_inject) 1); +by (REPEAT (ares_tac prems 1)); +by (REPEAT (etac SigmaE 1)); +by (etac rev_mp 1); +by (ASM_SIMP_TAC ZF_ss 1); +by (fast_tac (intrel_cs addSIs [SigmaI, equiv_intrel] + addSEs [box_equals RS sym, add_commute, + make_elim eq_equiv_class]) 1); +val zminus_inject = result(); + +val prems = goalw Integ.thy [zminus_def] + "[| x: nat; y: nat |] ==> $~ (intrel``{}) = intrel `` {}"; +by (ASM_SIMP_TAC + (ZF_ss addrews (prems@[zminus_ize UN_equiv_class, SigmaI])) 1); +val zminus = result(); + +goalw Integ.thy [integ_def] "!!z. z : integ ==> $~ ($~ z) = z"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (ASM_SIMP_TAC (ZF_ss addrews [zminus] addcongs integ_congs) 1); +val zminus_zminus = result(); + +goalw Integ.thy [integ_def, znat_def] "$~ ($#0) = $#0"; +by (SIMP_TAC (arith_ss addrews [zminus]) 1); +val zminus_0 = result(); + + +(**** znegative: the test for negative integers ****) + +goalw Integ.thy [znegative_def, znat_def] + "~ znegative($# n)"; +by (safe_tac intrel_cs); +by (rtac (add_not_less_self RS notE) 1); +by (etac ssubst 3); +by (ASM_SIMP_TAC (arith_ss addrews [add_0_right]) 3); +by (REPEAT (assume_tac 1)); +val not_znegative_znat = result(); + +val [nnat] = goalw Integ.thy [znegative_def, znat_def] + "n: nat ==> znegative($~ $# succ(n))"; +by (SIMP_TAC (intrel_ss addrews [zminus,nnat]) 1); +by (REPEAT + (resolve_tac [refl, exI, conjI, naturals_are_ordinals RS Ord_0_mem_succ, + refl RS intrelI RS imageI, consI1, nnat, nat_0I, + nat_succI] 1)); +val znegative_zminus_znat = result(); + + +(**** zmagnitude: magnitide of an integer, as a natural number ****) + +goalw Integ.thy [congruent_def] + "congruent(intrel, split(%x y. (y#-x) #+ (x#-y)))"; +by (safe_tac intrel_cs); +by (ALLGOALS (ASM_SIMP_TAC intrel_ss)); +by (etac rev_mp 1); +by (res_inst_tac [("m","x1"),("n","y1")] diff_induct 1); +by (REPEAT (assume_tac 1)); +by (ASM_SIMP_TAC (arith_ss addrews [add_succ_right,succ_inject_iff]) 3); +by (ASM_SIMP_TAC + (arith_ss addrews [diff_add_inverse,diff_add_0,add_0_right]) 2); +by (ASM_SIMP_TAC (arith_ss addrews [add_0_right]) 1); +by (rtac impI 1); +by (etac subst 1); +by (res_inst_tac [("m1","x")] (add_commute RS ssubst) 1); +by (REPEAT (assume_tac 1)); +by (ASM_SIMP_TAC (arith_ss addrews [diff_add_inverse,diff_add_0]) 1); +val zmagnitude_congruent = result(); + +(*Resolve th against the corresponding facts for zmagnitude*) +val zmagnitude_ize = RSLIST [equiv_intrel, zmagnitude_congruent]; + +val [prem] = goalw Integ.thy [integ_def,zmagnitude_def] + "z : integ ==> zmagnitude(z) : nat"; +by (typechk_tac [split_type, prem, zmagnitude_ize UN_equiv_class_type, + add_type, diff_type]); +val zmagnitude_type = result(); + +val prems = goalw Integ.thy [zmagnitude_def] + "[| x: nat; y: nat |] ==> \ +\ zmagnitude (intrel``{}) = (y #- x) #+ (x #- y)"; +by (ASM_SIMP_TAC + (ZF_ss addrews (prems@[zmagnitude_ize UN_equiv_class, SigmaI])) 1); +val zmagnitude = result(); + +val [nnat] = goalw Integ.thy [znat_def] + "n: nat ==> zmagnitude($# n) = n"; +by (SIMP_TAC (intrel_ss addrews [zmagnitude,nnat]) 1); +val zmagnitude_znat = result(); + +val [nnat] = goalw Integ.thy [znat_def] + "n: nat ==> zmagnitude($~ $# n) = n"; +by (SIMP_TAC (intrel_ss addrews [zmagnitude,zminus,nnat,add_0_right]) 1); +val zmagnitude_zminus_znat = result(); + + +(**** zadd: addition on integ ****) + +(** Congruence property for addition **) + +goalw Integ.thy [congruent2_def] + "congruent2(intrel, %p1 p2. \ +\ split(%x1 y1. split(%x2 y2. intrel `` {}, p2), p1))"; +(*Proof via congruent2_commuteI seems longer*) +by (safe_tac intrel_cs); +by (INTEG_SIMP_TAC [add_assoc] 1); +(*The rest should be trivial, but rearranging terms is hard*) +by (res_inst_tac [("m1","x1a")] (add_assoc_swap RS ssubst) 1); +by (res_inst_tac [("m1","x2a")] (add_assoc_swap RS ssubst) 3); +by (typechk_tac [add_type]); +by (ASM_SIMP_TAC (arith_ss addrews [add_assoc RS sym]) 1); +val zadd_congruent2 = result(); + +(*Resolve th against the corresponding facts for zadd*) +val zadd_ize = RSLIST [equiv_intrel, zadd_congruent2]; + +val prems = goalw Integ.thy [integ_def,zadd_def] + "[| z: integ; w: integ |] ==> z $+ w : integ"; +by (REPEAT (ares_tac (prems@[zadd_ize UN_equiv_class_type2, + split_type, add_type, quotientI, SigmaI]) 1)); +val zadd_type = result(); + +val prems = goalw Integ.thy [zadd_def] + "[| x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ +\ (intrel``{}) $+ (intrel``{}) = intrel `` {}"; +by (ASM_SIMP_TAC (ZF_ss addrews + (prems @ [zadd_ize UN_equiv_class2, SigmaI])) 1); +val zadd = result(); + +goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $+ z = z"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (ASM_SIMP_TAC (arith_ss addrews [zadd]) 1); +val zadd_0 = result(); + +goalw Integ.thy [integ_def] + "!!z w. [| z: integ; w: integ |] ==> $~ (z $+ w) = $~ z $+ $~ w"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (ASM_SIMP_TAC (arith_ss addrews [zminus,zadd] addcongs integ_congs) 1); +val zminus_zadd_distrib = result(); + +goalw Integ.thy [integ_def] + "!!z w. [| z: integ; w: integ |] ==> z $+ w = w $+ z"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (INTEG_SIMP_TAC [zadd] 1); +by (REPEAT (ares_tac [add_commute,add_cong] 1)); +val zadd_commute = result(); + +goalw Integ.thy [integ_def] + "!!z1 z2 z3. [| z1: integ; z2: integ; z3: integ |] ==> \ +\ (z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +(*rewriting is much faster without intrel_iff, etc.*) +by (ASM_SIMP_TAC (intrel_ss0 addrews [zadd,add_assoc]) 1); +val zadd_assoc = result(); + +val prems = goalw Integ.thy [znat_def] + "[| m: nat; n: nat |] ==> $# (m #+ n) = ($#m) $+ ($#n)"; +by (ASM_SIMP_TAC (arith_ss addrews (zadd::prems)) 1); +val znat_add = result(); + +goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> z $+ ($~ z) = $#0"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (ASM_SIMP_TAC (intrel_ss addrews [zminus,zadd,add_0_right]) 1); +by (REPEAT (ares_tac [add_commute] 1)); +val zadd_zminus_inverse = result(); + +val prems = goal Integ.thy + "z : integ ==> ($~ z) $+ z = $#0"; +by (rtac (zadd_commute RS trans) 1); +by (REPEAT (resolve_tac (prems@[zminus_type, zadd_zminus_inverse]) 1)); +val zadd_zminus_inverse2 = result(); + +val prems = goal Integ.thy "z:integ ==> z $+ $#0 = z"; +by (rtac (zadd_commute RS trans) 1); +by (REPEAT (resolve_tac (prems@[znat_type,nat_0I,zadd_0]) 1)); +val zadd_0_right = result(); + + +(*Need properties of $- ??? Or use $- just as an abbreviation? + [| m: nat; n: nat; m>=n |] ==> $# (m #- n) = ($#m) $- ($#n) +*) + +(**** zmult: multiplication on integ ****) + +(** Congruence property for multiplication **) + +val prems = goalw Integ.thy [znat_def] + "[| k: nat; l: nat; m: nat; n: nat |] ==> \ +\ (k #+ l) #+ (m #+ n) = (k #+ m) #+ (n #+ l)"; +val add_commute' = read_instantiate [("m","l")] add_commute; +by (SIMP_TAC (arith_ss addrews ([add_commute',add_assoc]@prems)) 1); +val zmult_congruent_lemma = result(); + +goal Integ.thy + "congruent2(intrel, %p1 p2. \ +\ split(%x1 y1. split(%x2 y2. \ +\ intrel``{}, p2), p1))"; +by (rtac (equiv_intrel RS congruent2_commuteI) 1); +by (safe_tac intrel_cs); +by (ALLGOALS (INTEG_SIMP_TAC [])); +(*Proof that zmult is congruent in one argument*) +by (rtac (zmult_congruent_lemma RS trans) 2); +by (rtac (zmult_congruent_lemma RS trans RS sym) 6); +by (typechk_tac [mult_type]); +by (ASM_SIMP_TAC (arith_ss addrews [add_mult_distrib_left RS sym]) 2); +(*Proof that zmult is commutative on representatives*) +by (rtac add_cong 1); +by (rtac (add_commute RS trans) 2); +by (REPEAT (ares_tac [mult_commute,add_type,mult_type,add_cong] 1)); +val zmult_congruent2 = result(); + +(*Resolve th against the corresponding facts for zmult*) +val zmult_ize = RSLIST [equiv_intrel, zmult_congruent2]; + +val prems = goalw Integ.thy [integ_def,zmult_def] + "[| z: integ; w: integ |] ==> z $* w : integ"; +by (REPEAT (ares_tac (prems@[zmult_ize UN_equiv_class_type2, + split_type, add_type, mult_type, + quotientI, SigmaI]) 1)); +val zmult_type = result(); + + +val prems = goalw Integ.thy [zmult_def] + "[| x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ +\ (intrel``{}) $* (intrel``{}) = \ +\ intrel `` {}"; +by (ASM_SIMP_TAC (ZF_ss addrews + (prems @ [zmult_ize UN_equiv_class2, SigmaI])) 1); +val zmult = result(); + +goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $* z = $#0"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (ASM_SIMP_TAC (arith_ss addrews [zmult]) 1); +val zmult_0 = result(); + +goalw Integ.thy [integ_def,znat_def,one_def] + "!!z. z : integ ==> $#1 $* z = z"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (ASM_SIMP_TAC (arith_ss addrews [zmult,add_0_right]) 1); +val zmult_1 = result(); + +goalw Integ.thy [integ_def] + "!!z w. [| z: integ; w: integ |] ==> ($~ z) $* w = $~ (z $* w)"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (INTEG_SIMP_TAC [zminus,zmult] 1); +by (REPEAT (ares_tac [mult_type,add_commute,add_cong] 1)); +val zmult_zminus = result(); + +goalw Integ.thy [integ_def] + "!!z w. [| z: integ; w: integ |] ==> ($~ z) $* ($~ w) = (z $* w)"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (INTEG_SIMP_TAC [zminus,zmult] 1); +by (REPEAT (ares_tac [mult_type,add_commute,add_cong] 1)); +val zmult_zminus_zminus = result(); + +goalw Integ.thy [integ_def] + "!!z w. [| z: integ; w: integ |] ==> z $* w = w $* z"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (INTEG_SIMP_TAC [zmult] 1); +by (res_inst_tac [("m1","xc #* y")] (add_commute RS ssubst) 1); +by (REPEAT (ares_tac [mult_type,mult_commute,add_cong] 1)); +val zmult_commute = result(); + +goalw Integ.thy [integ_def] + "!!z1 z2 z3. [| z1: integ; z2: integ; z3: integ |] ==> \ +\ (z1 $* z2) $* z3 = z1 $* (z2 $* z3)"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (INTEG_SIMP_TAC [zmult, add_mult_distrib_left, + add_mult_distrib, add_assoc, mult_assoc] 1); +(*takes 54 seconds due to wasteful type-checking*) +by (REPEAT (ares_tac [add_type, mult_type, add_commute, add_kill, + add_assoc_swap_kill, add_assoc_swap_kill RS sym] 1)); +val zmult_assoc = result(); + +goalw Integ.thy [integ_def] + "!!z1 z2 z3. [| z1: integ; z2: integ; w: integ |] ==> \ +\ (z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)"; +by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); +by (INTEG_SIMP_TAC [zadd, zmult, add_mult_distrib, add_assoc] 1); +(*takes 30 seconds due to wasteful type-checking*) +by (REPEAT (ares_tac [add_type, mult_type, refl, add_commute, add_kill, + add_assoc_swap_kill, add_assoc_swap_kill RS sym] 1)); +val zadd_zmult_distrib = result(); + +val integ_typechecks = + [znat_type, zminus_type, zmagnitude_type, zadd_type, zmult_type]; + +val integ_ss = + arith_ss addcongs integ_congs + addrews ([zminus_zminus,zmagnitude_znat,zmagnitude_zminus_znat, + zadd_0] @ integ_typechecks); diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/integ.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/integ.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,54 @@ +(* Title: ZF/ex/integ.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +The integers as equivalence classes over nat*nat. +*) + +Integ = Equiv + Arith + +consts + intrel,integ:: "i" + znat :: "i=>i" ("$# _" [80] 80) + zminus :: "i=>i" ("$~ _" [80] 80) + znegative :: "i=>o" + zmagnitude :: "i=>i" + "$*" :: "[i,i]=>i" (infixl 70) + "$'/" :: "[i,i]=>i" (infixl 70) + "$'/'/" :: "[i,i]=>i" (infixl 70) + "$+" :: "[i,i]=>i" (infixl 65) + "$-" :: "[i,i]=>i" (infixl 65) + "$<" :: "[i,i]=>o" (infixl 50) + +rules + + intrel_def + "intrel == {p:(nat*nat)*(nat*nat). \ +\ EX x1 y1 x2 y2. p=<,> & x1#+y2 = x2#+y1}" + + integ_def "integ == (nat*nat)/intrel" + + znat_def "$# m == intrel `` {}" + + zminus_def "$~ Z == UN p:Z. split(%x y. intrel``{}, p)" + + znegative_def + "znegative(Z) == EX x y. x:y & y:nat & :Z" + + zmagnitude_def + "zmagnitude(Z) == UN p:Z. split(%x y. (y#-x) #+ (x#-y), p)" + + zadd_def + "Z1 $+ Z2 == \ +\ UN p1:Z1. UN p2:Z2. split(%x1 y1. split(%x2 y2. \ +\ intrel``{}, p2), p1)" + + zdiff_def "Z1 $- Z2 == Z1 $+ zminus(Z2)" + zless_def "Z1 $< Z2 == znegative(Z1 $- Z2)" + + zmult_def + "Z1 $* Z2 == \ +\ UN p1:Z1. UN p2:Z2. split(%x1 y1. split(%x2 y2. \ +\ intrel``{}, p2), p1)" + + end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/listn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/listn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,47 @@ +(* Title: ZF/ex/listn + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Inductive definition of lists of n elements + +See Ch. Paulin-Mohring, Inductive Definitions in the System Coq. +Research Report 92-49, LIP, ENS Lyon. Dec 1992. +*) + +structure ListN = Inductive_Fun + (val thy = ListFn.thy addconsts [(["listn"],"i=>i")]; + val rec_doms = [("listn", "nat*list(A)")]; + val sintrs = + ["<0,Nil> : listn(A)", + "[| a: A; : listn(A) |] ==> : listn(A)"]; + val monos = []; + val con_defs = []; + val type_intrs = nat_typechecks@List.intrs@[SigmaI] + val type_elims = [SigmaE2]); + +(*Could be generated automatically; requires use of fsplitD*) +val listn_induct = standard + (fsplitI RSN + (2, fsplitI RSN + (3, (read_instantiate [("P", "fsplit(R)")] ListN.induct) RS fsplitD))); + +val [major] = goal ListN.thy "l:list(A) ==> : listn(A)"; +by (rtac (major RS List.induct) 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +by (REPEAT (ares_tac ListN.intrs 1)); +val list_into_listn = result(); + +goal ListN.thy " : listn(A) <-> l:list(A) & length(l)=n"; +by (rtac iffI 1); +by (etac listn_induct 1); +by (dtac fsplitD 2); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [Pair_iff]))); +by (fast_tac (ZF_cs addSIs [list_into_listn]) 1); +val listn_iff = result(); + +goal ListN.thy "listn(A)``{n} = {l:list(A). length(l)=n}"; +by (rtac equality_iffI 1); +by (SIMP_TAC (list_ss addrews [listn_iff,separation,image_singleton_iff]) 1); +val listn_image_eq = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/llist.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/llist.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,151 @@ +(* Title: ZF/ex/llist.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Co-Datatype definition of Lazy Lists + +Needs a "take-lemma" to prove llist_subset_quniv and to justify co-induction +for proving equality +*) + +structure LList = Co_Datatype_Fun + (val thy = QUniv.thy; + val thy = QUniv.thy; + val rec_specs = + [("llist", "quniv(A)", + [(["LNil"], "i"), + (["LCons"], "[i,i]=>i")])]; + val rec_styp = "i=>i"; + val ext = None + val sintrs = + ["LNil : llist(A)", + "[| a: A; l: llist(A) |] ==> LCons(a,l) : llist(A)"]; + val monos = []; + val type_intrs = co_data_typechecks + val type_elims = []); + +val [LNilI, LConsI] = LList.intrs; + +(*An elimination rule, for type-checking*) +val LConsE = LList.mk_cases LList.con_defs "LCons(a,l) : llist(A)"; + +(*Proving freeness results*) +val LCons_iff = LList.mk_free "LCons(a,l)=LCons(a',l') <-> a=a' & l=l'"; +val LNil_LCons_iff = LList.mk_free "~ LNil=LCons(a,l)"; + +(*** Lemmas to justify using "llist" in other recursive type definitions ***) + +goalw LList.thy LList.defs "!!A B. A<=B ==> llist(A) <= llist(B)"; +by (rtac gfp_mono 1); +by (REPEAT (rtac LList.bnd_mono 1)); +by (REPEAT (ares_tac (quniv_mono::basic_monos) 1)); +val llist_mono = result(); + +(** Closure of quniv(A) under llist -- why so complex? Its a gfp... **) + +val in_quniv_rls = + [Transset_quniv, QPair_Int_quniv_in_quniv, Int_Vfrom_0_in_quniv, + zero_Int_in_quniv, one_Int_in_quniv, + QPair_Int_Vfrom_succ_in_quniv, QPair_Int_Vfrom_in_quniv]; + +val quniv_cs = ZF_cs addSIs in_quniv_rls + addIs (Int_quniv_in_quniv::co_data_typechecks); + +(*Keep unfolding the lazy list until the induction hypothesis applies*) +goal LList.thy + "!!i. i : nat ==> \ +\ ALL l: llist(quniv(A)). l Int Vfrom(quniv(A), i) : quniv(A)"; +be complete_induct 1; +br ballI 1; +be LList.elim 1; +bws ([QInl_def,QInr_def]@LList.con_defs); +by (fast_tac quniv_cs 1); +by (etac natE 1 THEN REPEAT_FIRST hyp_subst_tac); +by (fast_tac quniv_cs 1); +by (fast_tac quniv_cs 1); +val llist_quniv_lemma = result(); + +goal LList.thy "llist(quniv(A)) <= quniv(A)"; +br subsetI 1; +br quniv_Int_Vfrom 1; +be (LList.dom_subset RS subsetD) 1; +by (REPEAT (ares_tac [llist_quniv_lemma RS bspec] 1)); +val llist_quniv = result(); + +val llist_subset_quniv = standard + (llist_mono RS (llist_quniv RSN (2,subset_trans))); + +(*** Equality for llist(A) as a greatest fixed point ***) + +structure LList_Eq = Co_Inductive_Fun + (val thy = LList.thy addconsts [(["lleq"],"i=>i")]; + val rec_doms = [("lleq","llist(A) <*> llist(A)")]; + val sintrs = + [" : lleq(A)", + "[| a:A; : lleq(A) |] ==> : lleq(A)"]; + val monos = []; + val con_defs = []; + val type_intrs = LList.intrs@[QSigmaI]; + val type_elims = [QSigmaE2]); + +(** Alternatives for above: + val con_defs = LList.con_defs + val type_intrs = co_data_typechecks + val type_elims = [quniv_QPair_E] +**) + +val lleq_cs = subset_cs + addSIs [succI1, Int_Vset_0_subset, + QPair_Int_Vset_succ_subset_trans, + QPair_Int_Vset_subset_trans]; + +(*Keep unfolding the lazy list until the induction hypothesis applies*) +goal LList_Eq.thy + "!!i. Ord(i) ==> ALL l l'. : lleq(A) --> l Int Vset(i) <= l'"; +be trans_induct 1; +by (safe_tac subset_cs); +be LList_Eq.elim 1; +by (safe_tac (subset_cs addSEs [QPair_inject])); +bws LList.con_defs; +by (etac Ord_cases 1 THEN REPEAT_FIRST hyp_subst_tac); +(*0 case*) +by (fast_tac lleq_cs 1); +(*succ(j) case*) +bw QInr_def; +by (fast_tac lleq_cs 1); +(*Limit(i) case*) +be (Limit_Vfrom_eq RS ssubst) 1; +br (Int_UN_distrib RS ssubst) 1; +by (fast_tac lleq_cs 1); +val lleq_Int_Vset_subset_lemma = result(); + +val lleq_Int_Vset_subset = standard + (lleq_Int_Vset_subset_lemma RS spec RS spec RS mp); + + +(*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*) +val [prem] = goal LList_Eq.thy " : lleq(A) ==> : lleq(A)"; +br (prem RS qconverseI RS LList_Eq.co_induct) 1; +br (LList_Eq.dom_subset RS qconverse_type) 1; +by (safe_tac qconverse_cs); +be LList_Eq.elim 1; +by (ALLGOALS (fast_tac qconverse_cs)); +val lleq_symmetric = result(); + +goal LList_Eq.thy "!!l l'. : lleq(A) ==> l=l'"; +br equalityI 1; +by (REPEAT (ares_tac [lleq_Int_Vset_subset RS Int_Vset_subset] 1 + ORELSE etac lleq_symmetric 1)); +val lleq_implies_equal = result(); + +val [eqprem,lprem] = goal LList_Eq.thy + "[| l=l'; l: llist(A) |] ==> : lleq(A)"; +by (res_inst_tac [("X", "{. l: llist(A)}")] LList_Eq.co_induct 1); +br (lprem RS RepFunI RS (eqprem RS subst)) 1; +by (safe_tac qpair_cs); +be LList.elim 1; +by (ALLGOALS (fast_tac qpair_cs)); +val equal_llist_implies_leq = result(); + + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/llistfn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/llistfn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,37 @@ +(* Title: ZF/ex/llist-fn.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Functions for Lazy Lists in Zermelo-Fraenkel Set Theory +*) + +open LListFn; + +(*** lconst -- defined directly using lfp, but equivalent to a LList_corec ***) + +goalw LListFn.thy LList.con_defs "bnd_mono(univ(a), %l. LCons(a,l))"; +by (rtac bnd_monoI 1); +by (REPEAT (ares_tac [subset_refl, QInr_mono, QPair_mono] 2)); +by (REPEAT (ares_tac [subset_refl, A_subset_univ, + QInr_subset_univ, QPair_subset_univ] 1)); +val lconst_fun_bnd_mono = result(); + +(* lconst(a) = LCons(a,lconst(a)) *) +val lconst = standard + ([lconst_def, lconst_fun_bnd_mono] MRS def_lfp_Tarski); + +val lconst_subset = lconst_def RS def_lfp_subset; + +val member_subset_Union_eclose = standard (arg_into_eclose RS Union_upper); + +goal LListFn.thy "!!a A. a : A ==> lconst(a) : quniv(A)"; +by (rtac (lconst_subset RS subset_trans RS qunivI) 1); +by (etac (arg_into_eclose RS eclose_subset RS univ_mono) 1); +val lconst_in_quniv = result(); + +goal LListFn.thy "!!a A. a:A ==> lconst(a): llist(A)"; +by (rtac (singletonI RS LList.co_induct) 1); +by (fast_tac (ZF_cs addSIs [lconst_in_quniv]) 1); +by (fast_tac (ZF_cs addSIs [lconst]) 1); +val lconst_type = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/llistfn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/llistfn.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,16 @@ +(* Title: ZF/ex/llist-fn.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Functions for Lazy Lists in Zermelo-Fraenkel Set Theory +*) + +LListFn = LList + +consts + lconst :: "i => i" + +rules + lconst_def "lconst(a) == lfp(univ(a), %l. LCons(a,l))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/misc.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/misc.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,213 @@ +(* Title: ZF/ex/misc + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Miscellaneous examples for Zermelo-Fraenkel Set Theory +Cantor's Theorem; Schroeder-Bernstein Theorem; Composition of homomorphisms... +*) + +writeln"ZF/ex/misc"; + + +(*Example 12 (credited to Peter Andrews) from + W. Bledsoe. A Maximal Method for Set Variables in Automatic Theorem-proving. + In: J. Hayes and D. Michie and L. Mikulich, eds. Machine Intelligence 9. + Ellis Horwood, 53-100 (1979). *) +goal ZF.thy "(ALL F. {x}: F --> {y}:F) --> (ALL A. x:A --> y:A)"; +by (best_tac ZF_cs 1); +result(); + + +(*** Cantor's Theorem: There is no surjection from a set to its powerset. ***) + +val cantor_cs = FOL_cs (*precisely the rules needed for the proof*) + addSIs [ballI, CollectI, PowI, subsetI] addIs [bexI] + addSEs [CollectE, equalityCE]; + +(*The search is undirected and similar proof attempts fail*) +goal ZF.thy "ALL f: A->Pow(A). EX S: Pow(A). ALL x:A. ~ f`x = S"; +by (best_tac cantor_cs 1); +result(); + +(*This form displays the diagonal term, {x: A . ~ x: f`x} *) +val [prem] = goal ZF.thy + "f: A->Pow(A) ==> (ALL x:A. ~ f`x = ?S) & ?S: Pow(A)"; +by (best_tac cantor_cs 1); +result(); + +(*yet another version...*) +goalw Perm.thy [surj_def] "~ f : surj(A,Pow(A))"; +by (safe_tac ZF_cs); +by (etac ballE 1); +by (best_tac (cantor_cs addSEs [bexE]) 1); +by (fast_tac ZF_cs 1); +result(); + + +(**** The Schroeder-Bernstein Theorem -- see Davey & Priestly, page 106 ****) + +val SB_thy = merge_theories (Fixedpt.thy, Perm.thy); + +(** Lemma: Banach's Decomposition Theorem **) + +goal SB_thy "bnd_mono(X, %W. X - g``(Y - f``W))"; +by (rtac bnd_monoI 1); +by (REPEAT (ares_tac [Diff_subset, subset_refl, Diff_mono, image_mono] 1)); +val decomp_bnd_mono = result(); + +val [gfun] = goal SB_thy + "g: Y->X ==> \ +\ g``(Y - f`` lfp(X, %W. X - g``(Y - f``W))) = \ +\ X - lfp(X, %W. X - g``(Y - f``W)) "; +by (res_inst_tac [("P", "%u. ?v = X-u")] + (decomp_bnd_mono RS lfp_Tarski RS ssubst) 1); +by (SIMP_TAC (ZF_ss addrews [subset_refl, double_complement, Diff_subset, + gfun RS fun_is_rel RS image_subset]) 1); +val Banach_last_equation = result(); + +val prems = goal SB_thy + "[| f: X->Y; g: Y->X |] ==> \ +\ EX XA XB YA YB. (XA Int XB = 0) & (XA Un XB = X) & \ +\ (YA Int YB = 0) & (YA Un YB = Y) & \ +\ f``XA=YA & g``YB=XB"; +by (REPEAT + (FIRSTGOAL + (resolve_tac [refl, exI, conjI, Diff_disjoint, Diff_partition]))); +by (rtac Banach_last_equation 3); +by (REPEAT (resolve_tac (prems@[fun_is_rel, image_subset, lfp_subset]) 1)); +val decomposition = result(); + +val prems = goal SB_thy + "[| f: inj(X,Y); g: inj(Y,X) |] ==> EX h. h: bij(X,Y)"; +by (cut_facts_tac prems 1); +by (cut_facts_tac [(prems RL [inj_is_fun]) MRS decomposition] 1); +by (fast_tac (ZF_cs addSIs [restrict_bij,bij_disjoint_Un] + addIs [bij_converse_bij]) 1); +(* The instantiation of exI to "restrict(f,XA) Un converse(restrict(g,YB))" + is forced by the context!! *) +val schroeder_bernstein = result(); + + +(*** Composition of homomorphisms is a homomorphism ***) + +(*Given as a challenge problem in + R. Boyer et al., + Set Theory in First-Order Logic: Clauses for G\"odel's Axioms, + JAR 2 (1986), 287-327 +*) + +val hom_ss = (*collecting the relevant lemmas*) + ZF_ss addrews [comp_func,comp_func_apply,SigmaI,apply_type] + addcongs (mk_congs Perm.thy ["op O"]); + +(*This version uses a super application of SIMP_TAC; it is SLOW + Expressing the goal by --> instead of ==> would make it slower still*) +val [hom_eq] = goal Perm.thy + "(ALL A f B g. hom(A,f,B,g) = \ +\ {H: A->B. f:A*A->A & g:B*B->B & \ +\ (ALL x:A. ALL y:A. H`(f`) = g`)}) ==> \ +\ J : hom(A,f,B,g) & K : hom(B,g,C,h) --> \ +\ (K O J) : hom(A,f,C,h)"; +by (SIMP_TAC (hom_ss setauto K(fast_tac prop_cs) addrews [hom_eq]) 1); +val comp_homs = result(); + +(*This version uses meta-level rewriting, safe_tac and ASM_SIMP_TAC*) +val [hom_def] = goal Perm.thy + "(!! A f B g. hom(A,f,B,g) == \ +\ {H: A->B. f:A*A->A & g:B*B->B & \ +\ (ALL x:A. ALL y:A. H`(f`) = g`)}) ==> \ +\ J : hom(A,f,B,g) & K : hom(B,g,C,h) --> \ +\ (K O J) : hom(A,f,C,h)"; +by (rewtac hom_def); +by (safe_tac ZF_cs); +by (ASM_SIMP_TAC hom_ss 1); +by (ASM_SIMP_TAC hom_ss 1); +val comp_homs = result(); + + +(** A characterization of functions, suggested by Tobias Nipkow **) + +goalw ZF.thy [Pi_def] + "r: domain(r)->B <-> r <= domain(r)*B & (ALL X. r `` (r -`` X) <= X)"; +by (safe_tac ZF_cs); +by (fast_tac (ZF_cs addSDs [bspec RS ex1_equalsE]) 1); +by (eres_inst_tac [("x", "{y}")] allE 1); +by (fast_tac ZF_cs 1); +result(); + + +(**** From D Pastre. Automatic theorem proving in set theory. + Artificial Intelligence, 10:1--27, 1978. + These examples require forward reasoning! ****) + +(*reduce the clauses to units by type checking -- beware of nontermination*) +fun forw_typechk tyrls [] = [] + | forw_typechk tyrls clauses = + let val (units, others) = partition (has_fewer_prems 1) clauses + in gen_union eq_thm (units, forw_typechk tyrls (tyrls RL others)) + end; + +(*A crude form of forward reasoning*) +fun forw_iterate tyrls rls facts 0 = facts + | forw_iterate tyrls rls facts n = + let val facts' = + gen_union eq_thm (forw_typechk (tyrls@facts) (facts RL rls), facts); + in forw_iterate tyrls rls facts' (n-1) end; + +val pastre_rls = + [comp_mem_injD1, comp_mem_surjD1, comp_mem_injD2, comp_mem_surjD2]; + +fun pastre_facts (fact1::fact2::fact3::prems) = + forw_iterate (prems @ [comp_surj, comp_inj, comp_func]) + pastre_rls [fact1,fact2,fact3] 4; + +val prems = goalw Perm.thy [bij_def] + "[| (h O g O f): inj(A,A); \ +\ (f O h O g): surj(B,B); \ +\ (g O f O h): surj(C,C); \ +\ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; +by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); +val pastre1 = result(); + +val prems = goalw Perm.thy [bij_def] + "[| (h O g O f): surj(A,A); \ +\ (f O h O g): inj(B,B); \ +\ (g O f O h): surj(C,C); \ +\ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; +by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); +val pastre2 = result(); + +val prems = goalw Perm.thy [bij_def] + "[| (h O g O f): surj(A,A); \ +\ (f O h O g): surj(B,B); \ +\ (g O f O h): inj(C,C); \ +\ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; +by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); +val pastre3 = result(); + +val prems = goalw Perm.thy [bij_def] + "[| (h O g O f): surj(A,A); \ +\ (f O h O g): inj(B,B); \ +\ (g O f O h): inj(C,C); \ +\ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; +by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); +val pastre4 = result(); + +val prems = goalw Perm.thy [bij_def] + "[| (h O g O f): inj(A,A); \ +\ (f O h O g): surj(B,B); \ +\ (g O f O h): inj(C,C); \ +\ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; +by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); +val pastre5 = result(); + +val prems = goalw Perm.thy [bij_def] + "[| (h O g O f): inj(A,A); \ +\ (f O h O g): inj(B,B); \ +\ (g O f O h): surj(C,C); \ +\ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; +by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); +val pastre6 = result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/parcontract.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/parcontract.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,144 @@ +(* Title: ZF/ex/parcontract.ML + ID: $Id$ + Author: Lawrence C Paulson + Copyright 1993 University of Cambridge + +Parallel contraction + +HOL system proofs may be found in +/usr/groups/theory/hvg-aftp/contrib/rule-induction/cl.ml +*) + +structure ParContract = Inductive_Fun + (val thy = Contract.thy; + val rec_doms = [("parcontract","comb*comb")]; + val sintrs = + ["[| p:comb |] ==> p =1=> p", + "[| p:comb; q:comb |] ==> K#p#q =1=> p", + "[| p:comb; q:comb; r:comb |] ==> S#p#q#r =1=> (p#r)#(q#r)", + "[| p=1=>q; r=1=>s |] ==> p#r =1=> q#s"]; + val monos = []; + val con_defs = []; + val type_intrs = Comb.intrs@[SigmaI]; + val type_elims = [SigmaE2]); + +val [parcontract_refl,K_parcontract,S_parcontract,Ap_parcontract] = ParContract.intrs; + +val parcontract_induct = standard + (ParContract.mutual_induct RS spec RS spec RSN (2,rev_mp)); + +(*For type checking: replaces a=1=>b by a,b:comb *) +val parcontract_combE2 = ParContract.dom_subset RS subsetD RS SigmaE2; +val parcontract_combD1 = ParContract.dom_subset RS subsetD RS SigmaD1; +val parcontract_combD2 = ParContract.dom_subset RS subsetD RS SigmaD2; + +goal ParContract.thy "field(parcontract) = comb"; +by (fast_tac (ZF_cs addIs [equalityI,K_parcontract] + addSEs [parcontract_combE2]) 1); +val field_parcontract_eq = result(); + +val parcontract_caseE = standard + (ParContract.unfold RS equalityD1 RS subsetD RS CollectE); + +(*Derive a case for each combinator constructor*) +val K_parcontract_case = ParContract.mk_cases Comb.con_defs "K =1=> r"; +val S_parcontract_case = ParContract.mk_cases Comb.con_defs "S =1=> r"; +val Ap_parcontract_case = ParContract.mk_cases Comb.con_defs "p#q =1=> r"; + +val parcontract_cs = + ZF_cs addSIs Comb.intrs + addIs ParContract.intrs + addSEs [Ap_E, K_parcontract_case, S_parcontract_case, + Ap_parcontract_case] + addSEs [parcontract_combD1, parcontract_combD2] (*type checking*) + addSEs Comb.free_SEs; + +(*** Basic properties of parallel contraction ***) + +goal ParContract.thy "!!p r. K#p =1=> r ==> (EX p'. r = K#p' & p =1=> p')"; +by (fast_tac parcontract_cs 1); +val K1_parcontractD = result(); + +goal ParContract.thy "!!p r. S#p =1=> r ==> (EX p'. r = S#p' & p =1=> p')"; +by (fast_tac parcontract_cs 1); +val S1_parcontractD = result(); + +goal ParContract.thy + "!!p q r. S#p#q =1=> r ==> (EX p' q'. r = S#p'#q' & p =1=> p' & q =1=> q')"; +by (fast_tac (parcontract_cs addSDs [S1_parcontractD]) 1); +val S2_parcontractD = result(); + +(*Church-Rosser property for parallel contraction*) +goalw ParContract.thy [diamond_def] "diamond(parcontract)"; +by (rtac (impI RS allI RS allI) 1); +by (etac parcontract_induct 1); +by (ALLGOALS + (fast_tac (parcontract_cs addSDs [K1_parcontractD,S2_parcontractD]))); +val diamond_parcontract = result(); + +(*** Transitive closure preserves the Church-Rosser property ***) + +goalw ParContract.thy [diamond_def] + "!!x y r. [| diamond(r); :r^+ |] ==> \ +\ ALL y'. :r --> (EX z. : r^+ & : r)"; +by (etac trancl_induct 1); +by (fast_tac (ZF_cs addIs [r_into_trancl]) 1); +by (slow_best_tac (ZF_cs addSDs [spec RS mp] + addIs [r_into_trancl, trans_trancl RS transD]) 1); +val diamond_trancl_lemma = result(); + +val diamond_lemmaE = diamond_trancl_lemma RS spec RS mp RS exE; + +val [major] = goal ParContract.thy "diamond(r) ==> diamond(r^+)"; +bw diamond_def; (*unfold only in goal, not in premise!*) +by (rtac (impI RS allI RS allI) 1); +by (etac trancl_induct 1); +by (ALLGOALS + (slow_best_tac (ZF_cs addIs [r_into_trancl, trans_trancl RS transD] + addEs [major RS diamond_lemmaE]))); +val diamond_trancl = result(); + + +(*** Equivalence of p--->q and p===>q ***) + +goal ParContract.thy "!!p q. p-1->q ==> p=1=>q"; +by (etac contract_induct 1); +by (ALLGOALS (fast_tac (parcontract_cs))); +val contract_imp_parcontract = result(); + +goal ParContract.thy "!!p q. p--->q ==> p===>q"; +by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1); +by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1); +by (etac rtrancl_induct 1); +by (fast_tac (parcontract_cs addIs [r_into_trancl]) 1); +by (fast_tac (ZF_cs addIs [contract_imp_parcontract, + r_into_trancl, trans_trancl RS transD]) 1); +val reduce_imp_parreduce = result(); + + +goal ParContract.thy "!!p q. p=1=>q ==> p--->q"; +by (etac parcontract_induct 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +by (fast_tac (contract_cs addIs reduction_rls) 1); +by (rtac (trans_rtrancl RS transD) 1); +by (ALLGOALS + (fast_tac + (contract_cs addIs [Ap_reduce1, Ap_reduce2] + addSEs [parcontract_combD1,parcontract_combD2]))); +val parcontract_imp_reduce = result(); + +goal ParContract.thy "!!p q. p===>q ==> p--->q"; +by (forward_tac [trancl_type RS subsetD RS SigmaD1] 1); +by (dtac (field_parcontract_eq RS equalityD1 RS subsetD) 1); +by (etac trancl_induct 1); +by (etac parcontract_imp_reduce 1); +by (etac (trans_rtrancl RS transD) 1); +by (etac parcontract_imp_reduce 1); +val parreduce_imp_reduce = result(); + +goal ParContract.thy "p===>q <-> p--->q"; +by (REPEAT (ares_tac [iffI, parreduce_imp_reduce, reduce_imp_parreduce] 1)); +val parreduce_iff_reduce = result(); + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/prop.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/prop.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,47 @@ +(* Title: ZF/ex/prop.ML + ID: $Id$ + Author: Lawrence C Paulson + Copyright 1993 University of Cambridge + +Datatype definition of propositional logic formulae and inductive definition +of the propositional tautologies. +*) + +(*Example of a datatype with mixfix syntax for some constructors*) +structure Prop = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("prop", "univ(0)", + [(["Fls"], "i"), + (["Var"], "i=>i"), + (["op =>"], "[i,i]=>i")])]; + val rec_styp = "i"; + val ext = Some (NewSext { + mixfix = + [Mixfix("#_", "i => i", "Var", [100], 100), + Infixr("=>", "[i,i] => i", 90)], + xrules = [], + parse_ast_translation = [], + parse_preproc = None, + parse_postproc = None, + parse_translation = [], + print_translation = [], + print_preproc = None, + print_postproc = None, + print_ast_translation = []}); + val sintrs = + ["Fls : prop", + "n: nat ==> #n : prop", + "[| p: prop; q: prop |] ==> p=>q : prop"]; + val monos = []; + val type_intrs = data_typechecks; + val type_elims = []); + +val [FlsI,VarI,ImpI] = Prop.intrs; + + +(** Type-checking rules **) + +val ImpE = Prop.mk_cases Prop.con_defs "p=>q : prop"; + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/proplog.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/proplog.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,333 @@ +(* Title: ZF/ex/prop-log.ML + ID: $Id$ + Author: Tobias Nipkow & Lawrence C Paulson + Copyright 1992 University of Cambridge + +For ex/prop-log.thy. Inductive definition of propositional logic. +Soundness and completeness w.r.t. truth-tables. + +Prove: If H|=p then G|=p where G:Fin(H) +*) + +open PropLog; + +(*** prop_rec -- by Vset recursion ***) + +val prop_congs = mk_typed_congs Prop.thy + [("b", "[i,i,i]=>i"), ("d", "[i,i,i,i]=>i")]; + +(** conversion rules **) + +goal PropLog.thy "prop_rec(Fls,b,c,d) = b"; +by (rtac (prop_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Prop.con_defs); +by (SIMP_TAC rank_ss 1); +val prop_rec_Fls = result(); + +goal PropLog.thy "prop_rec(#v,b,c,d) = c(v)"; +by (rtac (prop_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Prop.con_defs); +by (SIMP_TAC (rank_ss addcongs prop_congs) 1); +val prop_rec_Var = result(); + +goal PropLog.thy "prop_rec(p=>q,b,c,d) = \ +\ d(p, q, prop_rec(p,b,c,d), prop_rec(q,b,c,d))"; +by (rtac (prop_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Prop.con_defs); +by (SIMP_TAC (rank_ss addcongs prop_congs) 1); +val prop_rec_Imp = result(); + +val prop_rec_ss = + arith_ss addrews [prop_rec_Fls, prop_rec_Var, prop_rec_Imp]; + +(*** Semantics of propositional logic ***) + +(** The function is_true **) + +goalw PropLog.thy [is_true_def] "is_true(Fls,t) <-> False"; +by (SIMP_TAC (prop_rec_ss addrews [one_not_0 RS not_sym]) 1); +val is_true_Fls = result(); + +goalw PropLog.thy [is_true_def] "is_true(#v,t) <-> v:t"; +by (SIMP_TAC (prop_rec_ss addrews [one_not_0 RS not_sym] + addsplits [expand_if]) 1); +val is_true_Var = result(); + +goalw PropLog.thy [is_true_def] + "is_true(p=>q,t) <-> (is_true(p,t)-->is_true(q,t))"; +by (SIMP_TAC (prop_rec_ss addsplits [expand_if]) 1); +val is_true_Imp = result(); + +(** The function hyps **) + +goalw PropLog.thy [hyps_def] "hyps(Fls,t) = 0"; +by (SIMP_TAC prop_rec_ss 1); +val hyps_Fls = result(); + +goalw PropLog.thy [hyps_def] "hyps(#v,t) = {if(v:t, #v, #v=>Fls)}"; +by (SIMP_TAC prop_rec_ss 1); +val hyps_Var = result(); + +goalw PropLog.thy [hyps_def] "hyps(p=>q,t) = hyps(p,t) Un hyps(q,t)"; +by (SIMP_TAC prop_rec_ss 1); +val hyps_Imp = result(); + +val prop_ss = prop_rec_ss + addcongs Prop.congs + addcongs (mk_congs PropLog.thy ["Fin", "thms", "op |=","is_true","hyps"]) + addrews Prop.intrs + addrews [is_true_Fls, is_true_Var, is_true_Imp, + hyps_Fls, hyps_Var, hyps_Imp]; + +(*** Proof theory of propositional logic ***) + +structure PropThms = Inductive_Fun + (val thy = PropLog.thy; + val rec_doms = [("thms","prop")]; + val sintrs = + ["[| p:H; p:prop |] ==> H |- p", + "[| p:prop; q:prop |] ==> H |- p=>q=>p", + "[| p:prop; q:prop; r:prop |] ==> H |- (p=>q=>r) => (p=>q) => p=>r", + "p:prop ==> H |- ((p=>Fls) => Fls) => p", + "[| H |- p=>q; H |- p; p:prop; q:prop |] ==> H |- q"]; + val monos = []; + val con_defs = []; + val type_intrs = Prop.intrs; + val type_elims = []); + +goalw PropThms.thy PropThms.defs "!!G H. G<=H ==> thms(G) <= thms(H)"; +by (rtac lfp_mono 1); +by (REPEAT (rtac PropThms.bnd_mono 1)); +by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); +val thms_mono = result(); + +val thms_in_pl = PropThms.dom_subset RS subsetD; + +val [thms_H, thms_K, thms_S, thms_DN, weak_thms_MP] = PropThms.intrs; + +(*Modus Ponens rule -- this stronger version avoids typecheck*) +goal PropThms.thy "!!p q H. [| H |- p=>q; H |- p |] ==> H |- q"; +by (rtac weak_thms_MP 1); +by (REPEAT (eresolve_tac [asm_rl, thms_in_pl, thms_in_pl RS ImpE] 1)); +val thms_MP = result(); + +(*Rule is called I for Identity Combinator, not for Introduction*) +goal PropThms.thy "!!p H. p:prop ==> H |- p=>p"; +by (rtac (thms_S RS thms_MP RS thms_MP) 1); +by (rtac thms_K 5); +by (rtac thms_K 4); +by (REPEAT (ares_tac [ImpI] 1)); +val thms_I = result(); + +(** Weakening, left and right **) + +(* [| G<=H; G|-p |] ==> H|-p Order of premises is convenient with RS*) +val weaken_left = standard (thms_mono RS subsetD); + +(* H |- p ==> cons(a,H) |- p *) +val weaken_left_cons = subset_consI RS weaken_left; + +val weaken_left_Un1 = Un_upper1 RS weaken_left; +val weaken_left_Un2 = Un_upper2 RS weaken_left; + +goal PropThms.thy "!!H p q. [| H |- q; p:prop |] ==> H |- p=>q"; +by (rtac (thms_K RS thms_MP) 1); +by (REPEAT (ares_tac [thms_in_pl] 1)); +val weaken_right = result(); + +(*The deduction theorem*) +goal PropThms.thy "!!p q H. [| cons(p,H) |- q; p:prop |] ==> H |- p=>q"; +by (etac PropThms.induct 1); +by (fast_tac (ZF_cs addIs [thms_I, thms_H RS weaken_right]) 1); +by (fast_tac (ZF_cs addIs [thms_K RS weaken_right]) 1); +by (fast_tac (ZF_cs addIs [thms_S RS weaken_right]) 1); +by (fast_tac (ZF_cs addIs [thms_DN RS weaken_right]) 1); +by (fast_tac (ZF_cs addIs [thms_S RS thms_MP RS thms_MP]) 1); +val deduction = result(); + + +(*The cut rule*) +goal PropThms.thy "!!H p q. [| H|-p; cons(p,H) |- q |] ==> H |- q"; +by (rtac (deduction RS thms_MP) 1); +by (REPEAT (ares_tac [thms_in_pl] 1)); +val cut = result(); + +goal PropThms.thy "!!H p. [| H |- Fls; p:prop |] ==> H |- p"; +by (rtac (thms_DN RS thms_MP) 1); +by (rtac weaken_right 2); +by (REPEAT (ares_tac (Prop.intrs@[consI1]) 1)); +val thms_FlsE = result(); + +(* [| H |- p=>Fls; H |- p; q: prop |] ==> H |- q *) +val thms_notE = standard (thms_MP RS thms_FlsE); + +(*Soundness of the rules wrt truth-table semantics*) +val [major] = goalw PropThms.thy [sat_def] "H |- p ==> H |= p"; +by (rtac (major RS PropThms.induct) 1); +by (fast_tac (ZF_cs addSDs [is_true_Imp RS iffD1 RS mp]) 5); +by (ALLGOALS (SIMP_TAC prop_ss)); +val soundness = result(); + +(*** Towards the completeness proof ***) + +val [premf,premq] = goal PropThms.thy + "[| H |- p=>Fls; q: prop |] ==> H |- p=>q"; +by (rtac (premf RS thms_in_pl RS ImpE) 1); +by (rtac deduction 1); +by (rtac (premf RS weaken_left_cons RS thms_notE) 1); +by (REPEAT (ares_tac [premq, consI1, thms_H] 1)); +val Fls_Imp = result(); + +val [premp,premq] = goal PropThms.thy + "[| H |- p; H |- q=>Fls |] ==> H |- (p=>q)=>Fls"; +by (cut_facts_tac ([premp,premq] RL [thms_in_pl]) 1); +by (etac ImpE 1); +by (rtac deduction 1); +by (rtac (premq RS weaken_left_cons RS thms_MP) 1); +by (rtac (consI1 RS thms_H RS thms_MP) 1); +by (rtac (premp RS weaken_left_cons) 2); +by (REPEAT (ares_tac Prop.intrs 1)); +val Imp_Fls = result(); + +(*Typical example of strengthening the induction formula*) +val [major] = goal PropThms.thy + "p: prop ==> hyps(p,t) |- if(is_true(p,t), p, p=>Fls)"; +by (rtac (expand_if RS iffD2) 1); +by (rtac (major RS Prop.induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (prop_ss addrews [thms_I, thms_H]))); +by (fast_tac (ZF_cs addIs [weaken_left_Un1, weaken_left_Un2, + weaken_right, Imp_Fls] + addSEs [Fls_Imp]) 1); +val hyps_thms_if = result(); + +(*Key lemma for completeness; yields a set of assumptions satisfying p*) +val [premp,sat] = goalw PropThms.thy [sat_def] + "[| p: prop; 0 |= p |] ==> hyps(p,t) |- p"; +by (rtac (sat RS spec RS mp RS if_P RS subst) 1 THEN + rtac (premp RS hyps_thms_if) 2); +by (fast_tac ZF_cs 1); +val sat_thms_p = result(); + +(*For proving certain theorems in our new propositional logic*) +val thms_cs = + ZF_cs addSIs [FlsI, VarI, ImpI, deduction] + addIs [thms_in_pl, thms_H, thms_H RS thms_MP]; + +(*The excluded middle in the form of an elimination rule*) +val prems = goal PropThms.thy + "[| p: prop; q: prop |] ==> H |- (p=>q) => ((p=>Fls)=>q) => q"; +by (rtac (deduction RS deduction) 1); +by (rtac (thms_DN RS thms_MP) 1); +by (ALLGOALS (best_tac (thms_cs addSIs prems))); +val thms_excluded_middle = result(); + +(*Hard to prove directly because it requires cuts*) +val prems = goal PropThms.thy + "[| cons(p,H) |- q; cons(p=>Fls,H) |- q; p: prop |] ==> H |- q"; +by (rtac (thms_excluded_middle RS thms_MP RS thms_MP) 1); +by (REPEAT (resolve_tac (prems@Prop.intrs@[deduction,thms_in_pl]) 1)); +val thms_excluded_middle_rule = result(); + +(*** Completeness -- lemmas for reducing the set of assumptions ***) + +(*For the case hyps(p,t)-cons(#v,Y) |- p; + we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *) +val [major] = goal PropThms.thy + "p: prop ==> hyps(p, t-{v}) <= cons(#v=>Fls, hyps(p,t)-{#v})"; +by (rtac (major RS Prop.induct) 1); +by (SIMP_TAC prop_ss 1); +by (ASM_SIMP_TAC (prop_ss addsplits [expand_if]) 1); +by (fast_tac (ZF_cs addSEs Prop.free_SEs) 1); +by (ASM_SIMP_TAC prop_ss 1); +by (fast_tac ZF_cs 1); +val hyps_Diff = result(); + +(*For the case hyps(p,t)-cons(#v => Fls,Y) |- p; + we also have hyps(p,t)-{#v=>Fls} <= hyps(p, cons(v,t)) *) +val [major] = goal PropThms.thy + "p: prop ==> hyps(p, cons(v,t)) <= cons(#v, hyps(p,t)-{#v=>Fls})"; +by (rtac (major RS Prop.induct) 1); +by (SIMP_TAC prop_ss 1); +by (ASM_SIMP_TAC (prop_ss addsplits [expand_if]) 1); +by (fast_tac (ZF_cs addSEs Prop.free_SEs) 1); +by (ASM_SIMP_TAC prop_ss 1); +by (fast_tac ZF_cs 1); +val hyps_cons = result(); + +(** Two lemmas for use with weaken_left **) + +goal ZF.thy "B-C <= cons(a, B-cons(a,C))"; +by (fast_tac ZF_cs 1); +val cons_Diff_same = result(); + +goal ZF.thy "cons(a, B-{c}) - D <= cons(a, B-cons(c,D))"; +by (fast_tac ZF_cs 1); +val cons_Diff_subset2 = result(); + +(*The set hyps(p,t) is finite, and elements have the form #v or #v=>Fls; + could probably prove the stronger hyps(p,t) : Fin(hyps(p,0) Un hyps(p,nat))*) +val [major] = goal PropThms.thy + "p: prop ==> hyps(p,t) : Fin(UN v:nat. {#v, #v=>Fls})"; +by (rtac (major RS Prop.induct) 1); +by (ASM_SIMP_TAC (prop_ss addrews [Fin_0I, Fin_consI, UN_I] + addsplits [expand_if]) 2); +by (ALLGOALS (ASM_SIMP_TAC (prop_ss addrews [Un_0, Fin_0I, Fin_UnI]))); +val hyps_finite = result(); + +val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left; + +(*Induction on the finite set of assumptions hyps(p,t0). + We may repeatedly subtract assumptions until none are left!*) +val [premp,sat] = goal PropThms.thy + "[| p: prop; 0 |= p |] ==> ALL t. hyps(p,t) - hyps(p,t0) |- p"; +by (rtac (premp RS hyps_finite RS Fin_induct) 1); +by (SIMP_TAC (prop_ss addrews [premp, sat, sat_thms_p, Diff_0]) 1); +by (safe_tac ZF_cs); +(*Case hyps(p,t)-cons(#v,Y) |- p *) +by (rtac thms_excluded_middle_rule 1); +by (etac VarI 3); +by (rtac (cons_Diff_same RS weaken_left) 1); +by (etac spec 1); +by (rtac (cons_Diff_subset2 RS weaken_left) 1); +by (rtac (premp RS hyps_Diff RS Diff_weaken_left) 1); +by (etac spec 1); +(*Case hyps(p,t)-cons(#v => Fls,Y) |- p *) +by (rtac thms_excluded_middle_rule 1); +by (etac VarI 3); +by (rtac (cons_Diff_same RS weaken_left) 2); +by (etac spec 2); +by (rtac (cons_Diff_subset2 RS weaken_left) 1); +by (rtac (premp RS hyps_cons RS Diff_weaken_left) 1); +by (etac spec 1); +val completeness_0_lemma = result(); + +(*The base case for completeness*) +val [premp,sat] = goal PropThms.thy "[| p: prop; 0 |= p |] ==> 0 |- p"; +by (rtac (Diff_cancel RS subst) 1); +by (rtac (sat RS (premp RS completeness_0_lemma RS spec)) 1); +val completeness_0 = result(); + +(*A semantic analogue of the Deduction Theorem*) +goalw PropThms.thy [sat_def] "!!H p q. [| cons(p,H) |= q |] ==> H |= p=>q"; +by (SIMP_TAC prop_ss 1); +by (fast_tac ZF_cs 1); +val sat_Imp = result(); + +goal PropThms.thy "!!H. H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p"; +by (etac Fin_induct 1); +by (safe_tac (ZF_cs addSIs [completeness_0])); +by (rtac (weaken_left_cons RS thms_MP) 1); +by (fast_tac (ZF_cs addSIs [sat_Imp,ImpI]) 1); +by (fast_tac thms_cs 1); +val completeness_lemma = result(); + +val completeness = completeness_lemma RS bspec RS mp; + +val [finite] = goal PropThms.thy "H: Fin(prop) ==> H |- p <-> H |= p & p:prop"; +by (fast_tac (ZF_cs addSEs [soundness, finite RS completeness, + thms_in_pl]) 1); +val thms_iff = result(); + +writeln"Reached end of file."; + + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/proplog.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/proplog.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,43 @@ +(* Title: ZF/ex/prop-log.thy + ID: $Id$ + Author: Tobias Nipkow & Lawrence C Paulson + Copyright 1993 University of Cambridge + +Inductive definition of propositional logic. +*) + +PropLog = Prop + Fin + +consts + (*semantics*) + prop_rec :: "[i, i, i=>i, [i,i,i,i]=>i] => i" + is_true :: "[i,i] => o" + "|=" :: "[i,i] => o" (infixl 50) + hyps :: "[i,i] => i" + + (*proof theory*) + thms :: "i => i" + "|-" :: "[i,i] => o" (infixl 50) + +translations + "H |- p" == "p : thms(H)" + +rules + + prop_rec_def + "prop_rec(p,b,c,h) == \ +\ Vrec(p, %p g.prop_case(b, c, %x y. h(x, y, g`x, g`y), p))" + + (** Semantics of propositional logic **) + is_true_def + "is_true(p,t) == prop_rec(p, 0, %v. if(v:t, 1, 0), \ +\ %p q tp tq. if(tp=1,tq,1)) = 1" + + (*For every valuation, if all elements of H are true then so is p*) + sat_def "H |= p == ALL t. (ALL q:H. is_true(q,t)) --> is_true(p,t)" + + (** A finite set of hypotheses from t and the Vars in p **) + hyps_def + "hyps(p,t) == prop_rec(p, 0, %v. {if(v:t, #v, #v=>Fls)}, \ +\ %p q Hp Hq. Hp Un Hq)" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/ramsey.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/ramsey.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,236 @@ +(* Title: ZF/ex/ramsey.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Ramsey's Theorem (finite exponent 2 version) + +Based upon the article + D Basin and M Kaufmann, + The Boyer-Moore Prover and Nuprl: An Experimental Comparison. + In G Huet and G Plotkin, editors, Logical Frameworks. + (CUP, 1991), pages 89--119 + +See also + M Kaufmann, + An example in NQTHM: Ramsey's Theorem + Internal Note, Computational Logic, Inc., Austin, Texas 78703 + Available from the author: kaufmann@cli.com +*) + +open Ramsey; + +val ramsey_congs = mk_congs Ramsey.thy ["Atleast"]; +val ramsey_ss = arith_ss addcongs ramsey_congs; + +(*** Cliques and Independent sets ***) + +goalw Ramsey.thy [Clique_def] "Clique(0,V,E)"; +by (fast_tac ZF_cs 1); +val Clique0 = result(); + +goalw Ramsey.thy [Clique_def] + "!!C V E. [| Clique(C,V',E); V'<=V |] ==> Clique(C,V,E)"; +by (fast_tac ZF_cs 1); +val Clique_superset = result(); + +goalw Ramsey.thy [Indept_def] "Indept(0,V,E)"; +by (fast_tac ZF_cs 1); +val Indept0 = result(); + +val prems = goalw Ramsey.thy [Indept_def] + "!!I V E. [| Indept(I,V',E); V'<=V |] ==> Indept(I,V,E)"; +by (fast_tac ZF_cs 1); +val Indept_superset = result(); + +(*** Atleast ***) + +goalw Ramsey.thy [Atleast_def,inj_def] "Atleast(0,A)"; +by (fast_tac (ZF_cs addIs [PiI]) 1); +val Atleast0 = result(); + +val [major] = goalw Ramsey.thy [Atleast_def] + "Atleast(succ(m),A) ==> EX x:A. Atleast(m, A-{x})"; +by (rtac (major RS exE) 1); +by (rtac bexI 1); +by (etac (inj_is_fun RS apply_type) 2); +by (rtac succI1 2); +by (rtac exI 1); +by (etac inj_succ_restrict 1); +val Atleast_succD = result(); + +val major::prems = goalw Ramsey.thy [Atleast_def] + "[| Atleast(n,A); A<=B |] ==> Atleast(n,B)"; +by (rtac (major RS exE) 1); +by (rtac exI 1); +by (etac inj_weaken_type 1); +by (resolve_tac prems 1); +val Atleast_superset = result(); + +val prems = goalw Ramsey.thy [Atleast_def,succ_def] + "[| Atleast(m,B); ~ b: B |] ==> Atleast(succ(m), cons(b,B))"; +by (cut_facts_tac prems 1); +by (etac exE 1); +by (rtac exI 1); +by (etac inj_extend 1); +by (rtac mem_not_refl 1); +by (assume_tac 1); +val Atleast_succI = result(); + +val prems = goal Ramsey.thy + "[| Atleast(m, B-{x}); x: B |] ==> Atleast(succ(m), B)"; +by (cut_facts_tac prems 1); +by (etac (Atleast_succI RS Atleast_superset) 1); +by (fast_tac ZF_cs 1); +by (fast_tac ZF_cs 1); +val Atleast_Diff_succI = result(); + +(*** Main Cardinality Lemma ***) + +(*The #-succ(0) strengthens the original theorem statement, but precisely + the same proof could be used!!*) +val prems = goal Ramsey.thy + "m: nat ==> \ +\ ALL n: nat. ALL A B. Atleast((m#+n) #- succ(0), A Un B) --> \ +\ Atleast(m,A) | Atleast(n,B)"; +by (nat_ind_tac "m" prems 1); +by (fast_tac (ZF_cs addSIs [Atleast0]) 1); +by (ASM_SIMP_TAC ramsey_ss 1); +by (rtac ballI 1); +by (nat_ind_tac "n" [] 1); +by (fast_tac (ZF_cs addSIs [Atleast0]) 1); +by (ASM_SIMP_TAC (ramsey_ss addrews [add_succ_right]) 1); +by (safe_tac ZF_cs); +by (etac (Atleast_succD RS bexE) 1); +by (etac UnE 1); +(**case x:B. Instantiate the 'ALL A B' induction hypothesis. **) +by (dres_inst_tac [("x1","A"), ("x","B-{x}")] (spec RS spec) 2); +by (etac (mp RS disjE) 2); +(*cases Atleast(succ(m1),A) and Atleast(succ(n1),B)*) +by (REPEAT (eresolve_tac [asm_rl, notE, Atleast_Diff_succI] 3)); +(*proving the condition*) +by (etac Atleast_superset 2 THEN fast_tac ZF_cs 2); +(**case x:A. Instantiate the 'ALL n:nat. ALL A B' induction hypothesis. **) +by (dres_inst_tac [("x2","succ(n1)"), ("x1","A-{x}"), ("x","B")] + (bspec RS spec RS spec) 1); +by (etac nat_succI 1); +by (etac (mp RS disjE) 1); +(*cases Atleast(succ(m1),A) and Atleast(succ(n1),B)*) +by (REPEAT (eresolve_tac [asm_rl, Atleast_Diff_succI, notE] 2)); +(*proving the condition*) +by (ASM_SIMP_TAC (ramsey_ss addrews [add_succ_right]) 1); +by (etac Atleast_superset 1 THEN fast_tac ZF_cs 1); +val pigeon2_lemma = result(); + +(* [| m:nat; n:nat; Atleast(m #+ n #- succ(0), A Un B) |] ==> + Atleast(m,A) | Atleast(n,B) *) +val pigeon2 = standard (pigeon2_lemma RS bspec RS spec RS spec RS mp); + + +(**** Ramsey's Theorem ****) + +(** Base cases of induction; they now admit ANY Ramsey number **) + +goalw Ramsey.thy [Ramsey_def] "Ramsey(n,0,j)"; +by (fast_tac (ZF_cs addIs [Clique0,Atleast0]) 1); +val Ramsey0j = result(); + +goalw Ramsey.thy [Ramsey_def] "Ramsey(n,i,0)"; +by (fast_tac (ZF_cs addIs [Indept0,Atleast0]) 1); +val Ramseyi0 = result(); + +(** Lemmas for induction step **) + +(*The use of succ(m) here, rather than #-succ(0), simplifies the proof of + Ramsey_step_lemma.*) +val prems = goal Ramsey.thy + "[| Atleast(m #+ n, A); m: nat; n: nat |] ==> \ +\ Atleast(succ(m), {x:A. ~P(x)}) | Atleast(n, {x:A. P(x)})"; +by (rtac (nat_succI RS pigeon2) 1); +by (SIMP_TAC (ramsey_ss addrews prems) 3); +by (rtac Atleast_superset 3); +by (REPEAT (resolve_tac prems 1)); +by (fast_tac ZF_cs 1); +val Atleast_partition = result(); + +(*For the Atleast part, proves ~(a:I) from the second premise!*) +val prems = goalw Ramsey.thy [Symmetric_def,Indept_def] + "[| Symmetric(E); Indept(I, {z: V-{a}. ~ :E}, E); a: V; \ +\ Atleast(j,I) |] ==> \ +\ Indept(cons(a,I), V, E) & Atleast(succ(j), cons(a,I))"; +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addSEs [Atleast_succI]) 1); (*34 secs*) +val Indept_succ = result(); + +val prems = goalw Ramsey.thy [Symmetric_def,Clique_def] + "[| Symmetric(E); Clique(C, {z: V-{a}. :E}, E); a: V; \ +\ Atleast(j,C) |] ==> \ +\ Clique(cons(a,C), V, E) & Atleast(succ(j), cons(a,C))"; +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addSEs [Atleast_succI]) 1); (*41 secs*) +val Clique_succ = result(); + +(** Induction step **) + +(*Published proofs gloss over the need for Ramsey numbers to be POSITIVE.*) +val ram1::ram2::prems = goalw Ramsey.thy [Ramsey_def] + "[| Ramsey(succ(m), succ(i), j); Ramsey(n, i, succ(j)); \ +\ m: nat; n: nat |] ==> \ +\ Ramsey(succ(m#+n), succ(i), succ(j))"; +by (safe_tac ZF_cs); +by (etac (Atleast_succD RS bexE) 1); +by (eres_inst_tac [("P1","%z.:E")] (Atleast_partition RS disjE) 1); +by (REPEAT (resolve_tac prems 1)); +(*case m*) +by (rtac (ram1 RS spec RS spec RS mp RS disjE) 1); +by (fast_tac ZF_cs 1); +by (fast_tac (ZF_cs addEs [Clique_superset]) 1); (*easy -- given a Clique*) +by (safe_tac ZF_cs); +by (eresolve_tac (swapify [exI]) 1); (*ignore main EX quantifier*) +by (REPEAT (ares_tac [Indept_succ] 1)); (*make a bigger Indept*) +(*case n*) +by (rtac (ram2 RS spec RS spec RS mp RS disjE) 1); +by (fast_tac ZF_cs 1); +by (safe_tac ZF_cs); +by (rtac exI 1); +by (REPEAT (ares_tac [Clique_succ] 1)); (*make a bigger Clique*) +by (fast_tac (ZF_cs addEs [Indept_superset]) 1); (*easy -- given an Indept*) +val Ramsey_step_lemma = result(); + + +(** The actual proof **) + +(*Again, the induction requires Ramsey numbers to be positive.*) +val prems = goal Ramsey.thy + "i: nat ==> ALL j: nat. EX n:nat. Ramsey(succ(n), i, j)"; +by (nat_ind_tac "i" prems 1); +by (fast_tac (ZF_cs addSIs [nat_0I,Ramsey0j]) 1); +by (rtac ballI 1); +by (nat_ind_tac "j" [] 1); +by (fast_tac (ZF_cs addSIs [nat_0I,Ramseyi0]) 1); +by (dres_inst_tac [("x","succ(j1)")] bspec 1); +by (REPEAT (eresolve_tac [nat_succI,bexE] 1)); +by (rtac bexI 1); +by (rtac Ramsey_step_lemma 1); +by (REPEAT (ares_tac [nat_succI,add_type] 1)); +val ramsey_lemma = result(); + +(*Final statement in a tidy form, without succ(...) *) +val prems = goal Ramsey.thy + "[| i: nat; j: nat |] ==> EX n:nat. Ramsey(n,i,j)"; +by (rtac (ramsey_lemma RS bspec RS bexE) 1); +by (etac bexI 3); +by (REPEAT (ares_tac (prems@[nat_succI]) 1)); +val ramsey = result(); + +(*Computer Ramsey numbers according to proof above -- which, actually, + does not constrain the base case values at all!*) +fun ram 0 j = 1 + | ram i 0 = 1 + | ram i j = ram (i-1) j + ram i (j-1); + +(*Previous proof gave the following Ramsey numbers, which are smaller than + those above by one!*) +fun ram' 0 j = 0 + | ram' i 0 = 0 + | ram' i j = ram' (i-1) j + ram' i (j-1) + 1; diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/ramsey.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/ramsey.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,46 @@ +(* Title: ZF/ex/ramsey.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Ramsey's Theorem (finite exponent 2 version) + +Based upon the article + D Basin and M Kaufmann, + The Boyer-Moore Prover and Nuprl: An Experimental Comparison. + In G Huet and G Plotkin, editors, Logical Frameworks. + (CUP, 1991), pages 89--119 + +See also + M Kaufmann, + An example in NQTHM: Ramsey's Theorem + Internal Note, Computational Logic, Inc., Austin, Texas 78703 + Available from the author: kaufmann@cli.com +*) + +Ramsey = Arith + +consts + Symmetric :: "i=>o" + Atleast :: "[i,i]=>o" + Clique,Indept,Ramsey :: "[i,i,i]=>o" + +rules + + Symmetric_def + "Symmetric(E) == (ALL x y. :E --> :E)" + + Clique_def + "Clique(C,V,E) == (C<=V) & (ALL x:C. ALL y:C. ~ x=y --> :E)" + + Indept_def + "Indept(I,V,E) == (I<=V) & (ALL x:I. ALL y:I. ~ x=y --> ~ :E)" + + Atleast_def + "Atleast(n,S) == (EX f. f: inj(n,S))" + + Ramsey_def + "Ramsey(n,i,j) == ALL V E. Symmetric(E) & Atleast(n,V) --> \ +\ (EX C. Clique(C,V,E) & Atleast(i,C)) | \ +\ (EX I. Indept(I,V,E) & Atleast(j,I))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/term.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/term.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,66 @@ +(* Title: ZF/ex/term.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Datatype definition of terms over an alphabet. +Illustrates the list functor (essentially the same type as in Trees & Forests) +*) + +structure Term = Datatype_Fun + (val thy = List.thy; + val rec_specs = + [("term", "univ(A)", + [(["Apply"], "[i,i]=>i")])]; + val rec_styp = "i=>i"; + val ext = None + val sintrs = ["[| a: A; l: list(term(A)) |] ==> Apply(a,l) : term(A)"]; + val monos = [list_mono]; + val type_intrs = [SigmaI,Pair_in_univ, list_univ RS subsetD, A_into_univ]; + val type_elims = []); + +val [ApplyI] = Term.intrs; + +(*Induction on term(A) followed by induction on List *) +val major::prems = goal Term.thy + "[| t: term(A); \ +\ !!x. [| x: A |] ==> P(Apply(x,Nil)); \ +\ !!x z zs. [| x: A; z: term(A); zs: list(term(A)); P(Apply(x,zs)) \ +\ |] ==> P(Apply(x, Cons(z,zs))) \ +\ |] ==> P(t)"; +by (rtac (major RS Term.induct) 1); +by (etac List.induct 1); +by (etac CollectE 2); +by (REPEAT (ares_tac (prems@[list_CollectD]) 1)); +val term_induct2 = result(); + +(*Induction on term(A) to prove an equation*) +val major::prems = goal (merge_theories(Term.thy,ListFn.thy)) + "[| t: term(A); \ +\ !!x zs. [| x: A; zs: list(term(A)); map(f,zs) = map(g,zs) |] ==> \ +\ f(Apply(x,zs)) = g(Apply(x,zs)) \ +\ |] ==> f(t)=g(t)"; +by (rtac (major RS Term.induct) 1); +by (resolve_tac prems 1); +by (REPEAT (eresolve_tac [asm_rl, map_list_Collect, list_CollectD] 1)); +val term_induct_eqn = result(); + +(** Lemmas to justify using "term" in other recursive type definitions **) + +goalw Term.thy Term.defs "!!A B. A<=B ==> term(A) <= term(B)"; +by (rtac lfp_mono 1); +by (REPEAT (rtac Term.bnd_mono 1)); +by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); +val term_mono = result(); + +(*Easily provable by induction also*) +goalw Term.thy (Term.defs@Term.con_defs) "term(univ(A)) <= univ(A)"; +by (rtac lfp_lowerbound 1); +by (rtac (A_subset_univ RS univ_mono) 2); +by (safe_tac ZF_cs); +by (REPEAT (ares_tac [Pair_in_univ, list_univ RS subsetD] 1)); +val term_univ = result(); + +val term_subset_univ = standard + (term_mono RS (term_univ RSN (2,subset_trans))); + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/termfn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/termfn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,192 @@ +(* Title: ZF/ex/term + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Terms over a given alphabet -- function applications; illustrates list functor + (essentially the same type as in Trees & Forests) +*) + +writeln"File ZF/ex/term-fn."; + +open TermFn; + +(*** term_rec -- by Vset recursion ***) + +(*Lemma: map works correctly on the underlying list of terms*) +val [major,ordi] = goal ListFn.thy + "[| l: list(A); Ord(i) |] ==> \ +\ rank(l): i --> map(%z. (lam x:Vset(i).h(x)) ` z, l) = map(h,l)"; +by (rtac (major RS List.induct) 1); +by (SIMP_TAC list_ss 1); +by (rtac impI 1); +by (forward_tac [rank_Cons1 RS Ord_trans] 1); +by (dtac (rank_Cons2 RS Ord_trans) 2); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [ordi, VsetI]))); +val map_lemma = result(); + +(*Typing premise is necessary to invoke map_lemma*) +val [prem] = goal TermFn.thy + "ts: list(A) ==> \ +\ term_rec(Apply(a,ts), d) = d(a, ts, map (%z. term_rec(z,d), ts))"; +by (rtac (term_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac Term.con_defs); +val term_rec_ss = ZF_ss + addcongs (mk_typed_congs TermFn.thy [("d", "[i,i,i]=>i")]) + addrews [Ord_rank, rank_pair2, prem RS map_lemma]; +by (SIMP_TAC term_rec_ss 1); +val term_rec = result(); + +(*Slightly odd typing condition on r in the second premise!*) +val major::prems = goal TermFn.thy + "[| t: term(A); \ +\ !!x zs r. [| x: A; zs: list(term(A)); \ +\ r: list(UN t:term(A). C(t)) |] \ +\ ==> d(x, zs, r): C(Apply(x,zs)) \ +\ |] ==> term_rec(t,d) : C(t)"; +by (rtac (major RS Term.induct) 1); +by (forward_tac [list_CollectD] 1); +by (rtac (term_rec RS ssubst) 1); +by (REPEAT (ares_tac prems 1)); +by (etac List.induct 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [term_rec]))); +by (etac CollectE 1); +by (REPEAT (ares_tac [ConsI, UN_I] 1)); +val term_rec_type = result(); + +val [rew,tslist] = goal TermFn.thy + "[| !!t. j(t)==term_rec(t,d); ts: list(A) |] ==> \ +\ j(Apply(a,ts)) = d(a, ts, map(%Z.j(Z), ts))"; +by (rewtac rew); +by (rtac (tslist RS term_rec) 1); +val def_term_rec = result(); + +val prems = goal TermFn.thy + "[| t: term(A); \ +\ !!x zs r. [| x: A; zs: list(term(A)); r: list(C) |] \ +\ ==> d(x, zs, r): C \ +\ |] ==> term_rec(t,d) : C"; +by (REPEAT (ares_tac (term_rec_type::prems) 1)); +by (etac (subset_refl RS UN_least RS list_mono RS subsetD) 1); +val term_rec_simple_type = result(); + + +(** term_map **) + +val term_map = standard (term_map_def RS def_term_rec); + +val prems = goalw TermFn.thy [term_map_def] + "[| t: term(A); !!x. x: A ==> f(x): B |] ==> term_map(f,t) : term(B)"; +by (REPEAT (ares_tac ([term_rec_simple_type, ApplyI] @ prems) 1)); +val term_map_type = result(); + +val [major] = goal TermFn.thy + "t: term(A) ==> term_map(f,t) : term({f(u). u:A})"; +by (rtac (major RS term_map_type) 1); +by (etac RepFunI 1); +val term_map_type2 = result(); + + +(** term_size **) + +val term_size = standard (term_size_def RS def_term_rec); + +goalw TermFn.thy [term_size_def] "!!t A. t: term(A) ==> term_size(t) : nat"; +by (REPEAT (ares_tac [term_rec_simple_type, list_add_type, nat_succI] 1)); +val term_size_type = result(); + + +(** reflect **) + +val reflect = standard (reflect_def RS def_term_rec); + +goalw TermFn.thy [reflect_def] "!!t A. t: term(A) ==> reflect(t) : term(A)"; +by (REPEAT (ares_tac [term_rec_simple_type, rev_type, ApplyI] 1)); +val reflect_type = result(); + +(** preorder **) + +val preorder = standard (preorder_def RS def_term_rec); + +goalw TermFn.thy [preorder_def] + "!!t A. t: term(A) ==> preorder(t) : list(A)"; +by (REPEAT (ares_tac [term_rec_simple_type, ConsI, flat_type] 1)); +val preorder_type = result(); + + +(** Term simplification **) + +val term_typechecks = + [ApplyI, term_map_type, term_map_type2, term_size_type, reflect_type, + preorder_type]; + +(*map_type2 and term_map_type2 instantiate variables*) +val term_ss = list_ss + addcongs (Term.congs @ + mk_congs TermFn.thy ["term_rec","term_map","term_size", + "reflect","preorder"]) + addrews [term_rec, term_map, term_size, reflect, + preorder] + setauto type_auto_tac (list_typechecks@term_typechecks); + + +(** theorems about term_map **) + +goal TermFn.thy "!!t A. t: term(A) ==> term_map(%u.u, t) = t"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [map_ident]) 1); +val term_map_ident = result(); + +goal TermFn.thy + "!!t A. t: term(A) ==> term_map(f, term_map(g,t)) = term_map(%u.f(g(u)), t)"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [map_compose]) 1); +val term_map_compose = result(); + +goal TermFn.thy + "!!t A. t: term(A) ==> term_map(f, reflect(t)) = reflect(term_map(f,t))"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [rev_map_distrib RS sym, map_compose]) 1); +val term_map_reflect = result(); + + +(** theorems about term_size **) + +goal TermFn.thy + "!!t A. t: term(A) ==> term_size(term_map(f,t)) = term_size(t)"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [map_compose]) 1); +val term_size_term_map = result(); + +goal TermFn.thy "!!t A. t: term(A) ==> term_size(reflect(t)) = term_size(t)"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [rev_map_distrib RS sym, map_compose, + list_add_rev]) 1); +val term_size_reflect = result(); + +goal TermFn.thy "!!t A. t: term(A) ==> term_size(t) = length(preorder(t))"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [length_flat, map_compose]) 1); +val term_size_length = result(); + + +(** theorems about reflect **) + +goal TermFn.thy "!!t A. t: term(A) ==> reflect(reflect(t)) = t"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [rev_map_distrib, map_compose, + map_ident, rev_rev_ident]) 1); +val reflect_reflect_ident = result(); + + +(** theorems about preorder **) + +goal TermFn.thy + "!!t A. t: term(A) ==> preorder(term_map(f,t)) = map(f, preorder(t))"; +by (etac term_induct_eqn 1); +by (ASM_SIMP_TAC (term_ss addrews [map_compose, map_flat]) 1); +val preorder_term_map = result(); + +(** preorder(reflect(t)) = rev(postorder(t)) **) + +writeln"Reached end of file."; diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/termfn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/termfn.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,31 @@ +(* Title: ZF/ex/term-fn.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Terms over an alphabet. +Illustrates the list functor (essentially the same type as in Trees & Forests) +*) + +TermFn = Term + ListFn + +consts + term_rec :: "[i, [i,i,i]=>i] => i" + term_map :: "[i=>i, i] => i" + term_size :: "i=>i" + reflect :: "i=>i" + preorder :: "i=>i" + +rules + term_rec_def + "term_rec(t,d) == \ +\ Vrec(t, %t g. term_case(%x zs. d(x, zs, map(%z.g`z, zs)), t))" + + term_map_def "term_map(f,t) == term_rec(t, %x zs rs. Apply(f(x), rs))" + + term_size_def "term_size(t) == term_rec(t, %x zs rs. succ(list_add(rs)))" + + reflect_def "reflect(t) == term_rec(t, %x zs rs. Apply(x, rev(rs)))" + + preorder_def "preorder(t) == term_rec(t, %x zs rs. Cons(x, flat(rs)))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/tf.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/tf.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,65 @@ +(* Title: ZF/ex/tf.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Trees & forests, a mutually recursive type definition. +*) + +structure TF = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("tree", "univ(A)", + [(["Tcons"], "[i,i]=>i")]), + ("forest", "univ(A)", + [(["Fnil"], "i"), + (["Fcons"], "[i,i]=>i")])]; + val rec_styp = "i=>i"; + val ext = None + val sintrs = + ["[| a:A; tf: forest(A) |] ==> Tcons(a,tf) : tree(A)", + "Fnil : forest(A)", + "[| t: tree(A); tf: forest(A) |] ==> Fcons(t,tf) : forest(A)"]; + val monos = []; + val type_intrs = data_typechecks + val type_elims = []); + +val [TconsI, FnilI, FconsI] = TF.intrs; + +(** tree_forest(A) as the union of tree(A) and forest(A) **) + +goalw TF.thy (tl TF.defs) "tree(A) <= tree_forest(A)"; +by (rtac Part_subset 1); +val tree_subset_TF = result(); + +goalw TF.thy (tl TF.defs) "forest(A) <= tree_forest(A)"; +by (rtac Part_subset 1); +val forest_subset_TF = result(); + +goalw TF.thy (tl TF.defs) "tree(A) Un forest(A) = tree_forest(A)"; +by (rtac (TF.dom_subset RS Part_sum_equality) 1); +val TF_equals_Un = result(); + +(** NOT useful, but interesting... **) + +goalw TF.thy (tl TF.defs) "tree(A) = {Inl(x). x: A*forest(A)}"; +by (res_inst_tac [("P", "%x.?t(x) = ?u::i")] (TF.unfold RS ssubst) 1); +by (rewrite_goals_tac TF.con_defs); +by (rtac equalityI 1); +by (fast_tac (ZF_cs addIs [PartI] addSEs (PartE::TF.free_SEs)) 1); +by (fast_tac (ZF_cs addIs ([PartI,InlI,InrI] @ [TF.dom_subset RS subsetD] + @ data_typechecks) + addSEs (PartE::TF.free_SEs)) 1); +val tree_unfold = result(); + +goalw TF.thy (tl TF.defs) "forest(A) = {Inr(x). x: {0} + tree(A)*forest(A)}"; +by (res_inst_tac [("P", "%x.?t(x) = ?u::i")] (TF.unfold RS ssubst) 1); +by (rewrite_goals_tac TF.con_defs); +by (rtac equalityI 1); +by (fast_tac (ZF_cs addIs [PartI,InlI,InrI] + addSEs (PartE::TF.free_SEs)) 1); +by (fast_tac (ZF_cs addIs ([PartI,InlI,InrI] @ [TF.dom_subset RS subsetD] + @ data_typechecks) + addSEs ([PartE, sumE] @ TF.free_SEs)) 1); +val forest_unfold = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/tf_fn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/tf_fn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,233 @@ +(* Title: ZF/ex/tf.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For tf.thy. Trees & forests, a mutually recursive type definition. + +Still needs + +"TF_reflect == (%z. TF_rec(z, %x ts r. Tcons(x,r), 0, + %t ts r1 r2. TF_of_list(list_of_TF(r2) @ )))" +*) + +open TF_Fn; + + +(*** TF_rec -- by Vset recursion ***) + +(*Used only to verify TF_rec*) +val TF_congs = mk_typed_congs TF.thy + [("b", "[i,i,i]=>i"), ("d", "[i,i,i,i]=>i")]; + +(** conversion rules **) + +goal TF_Fn.thy "TF_rec(Tcons(a,tf), b, c, d) = b(a, tf, TF_rec(tf,b,c,d))"; +by (rtac (TF_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac TF.con_defs); +by (SIMP_TAC (rank_ss addcongs TF_congs) 1); +val TF_rec_Tcons = result(); + +goal TF_Fn.thy "TF_rec(Fnil, b, c, d) = c"; +by (rtac (TF_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac TF.con_defs); +by (SIMP_TAC rank_ss 1); +val TF_rec_Fnil = result(); + +goal TF_Fn.thy "TF_rec(Fcons(t,tf), b, c, d) = \ +\ d(t, tf, TF_rec(t, b, c, d), TF_rec(tf, b, c, d))"; +by (rtac (TF_rec_def RS def_Vrec RS trans) 1); +by (rewrite_goals_tac TF.con_defs); +by (SIMP_TAC (rank_ss addcongs TF_congs) 1); +val TF_rec_Fcons = result(); + +(*list_ss includes list operations as well as arith_ss*) +val TF_rec_ss = list_ss addrews + [TF_rec_Tcons, TF_rec_Fnil, TF_rec_Fcons, TconsI, FnilI, FconsI]; + +(** Type checking **) + +val major::prems = goal TF_Fn.thy + "[| z: tree_forest(A); \ +\ !!x tf r. [| x: A; tf: forest(A); r: C(tf) \ +\ |] ==> b(x,tf,r): C(Tcons(x,tf)); \ +\ c : C(Fnil); \ +\ !!t tf r1 r2. [| t: tree(A); tf: forest(A); r1: C(t); r2: C(tf) \ +\ |] ==> d(t,tf,r1,r2): C(Fcons(t,tf)) \ +\ |] ==> TF_rec(z,b,c,d) : C(z)"; +by (rtac (major RS TF.induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (TF_rec_ss addrews prems))); +val TF_rec_type = result(); + +(*Mutually recursive version*) +val prems = goal TF_Fn.thy + "[| !!x tf r. [| x: A; tf: forest(A); r: D(tf) \ +\ |] ==> b(x,tf,r): C(Tcons(x,tf)); \ +\ c : D(Fnil); \ +\ !!t tf r1 r2. [| t: tree(A); tf: forest(A); r1: C(t); r2: D(tf) \ +\ |] ==> d(t,tf,r1,r2): D(Fcons(t,tf)) \ +\ |] ==> (ALL t:tree(A). TF_rec(t,b,c,d) : C(t)) & \ +\ (ALL tf: forest(A). TF_rec(tf,b,c,d) : D(tf))"; +by (rewtac Ball_def); +by (rtac TF.mutual_induct 1); +by (ALLGOALS (ASM_SIMP_TAC (TF_rec_ss addrews prems))); +val tree_forest_rec_type = result(); + + +(** Versions for use with definitions **) + +val [rew] = goal TF_Fn.thy + "[| !!z. j(z)==TF_rec(z,b,c,d) |] ==> j(Tcons(a,tf)) = b(a,tf,j(tf))"; +by (rewtac rew); +by (rtac TF_rec_Tcons 1); +val def_TF_rec_Tcons = result(); + +val [rew] = goal TF_Fn.thy + "[| !!z. j(z)==TF_rec(z,b,c,d) |] ==> j(Fnil) = c"; +by (rewtac rew); +by (rtac TF_rec_Fnil 1); +val def_TF_rec_Fnil = result(); + +val [rew] = goal TF_Fn.thy + "[| !!z. j(z)==TF_rec(z,b,c,d) |] ==> j(Fcons(t,tf)) = d(t,tf,j(t),j(tf))"; +by (rewtac rew); +by (rtac TF_rec_Fcons 1); +val def_TF_rec_Fcons = result(); + +fun TF_recs def = map standard + ([def] RL [def_TF_rec_Tcons, def_TF_rec_Fnil, def_TF_rec_Fcons]); + + +(** list_of_TF and TF_of_list **) + +val [list_of_TF_Tcons, list_of_TF_Fnil, list_of_TF_Fcons] = + TF_recs list_of_TF_def; + +goalw TF_Fn.thy [list_of_TF_def] + "!!z A. z: tree_forest(A) ==> list_of_TF(z) : list(tree(A))"; +by (REPEAT (ares_tac ([TF_rec_type,TconsI] @ list_typechecks) 1)); +val list_of_TF_type = result(); + +val [TF_of_list_Nil,TF_of_list_Cons] = list_recs TF_of_list_def; + +goalw TF_Fn.thy [TF_of_list_def] + "!!l A. l: list(tree(A)) ==> TF_of_list(l) : forest(A)"; +by (REPEAT (ares_tac [list_rec_type, FnilI, FconsI] 1)); +val TF_of_list_type = result(); + + +(** TF_map **) + +val [TF_map_Tcons, TF_map_Fnil, TF_map_Fcons] = TF_recs TF_map_def; + +val prems = goalw TF_Fn.thy [TF_map_def] + "[| !!x. x: A ==> h(x): B |] ==> \ +\ (ALL t:tree(A). TF_map(h,t) : tree(B)) & \ +\ (ALL tf: forest(A). TF_map(h,tf) : forest(B))"; +by (REPEAT + (ares_tac ([tree_forest_rec_type, TconsI, FnilI, FconsI] @ prems) 1)); +val TF_map_type = result(); + + +(** TF_size **) + +val [TF_size_Tcons, TF_size_Fnil, TF_size_Fcons] = TF_recs TF_size_def; + +goalw TF_Fn.thy [TF_size_def] + "!!z A. z: tree_forest(A) ==> TF_size(z) : nat"; +by (REPEAT (ares_tac [TF_rec_type, add_type, nat_0I, nat_succI] 1)); +val TF_size_type = result(); + + +(** TF_preorder **) + +val [TF_preorder_Tcons, TF_preorder_Fnil, TF_preorder_Fcons] = + TF_recs TF_preorder_def; + +goalw TF_Fn.thy [TF_preorder_def] + "!!z A. z: tree_forest(A) ==> TF_preorder(z) : list(A)"; +by (REPEAT (ares_tac [TF_rec_type, app_type,NilI, ConsI] 1)); +val TF_preorder_type = result(); + + +(** Term simplification **) + +val treeI = tree_subset_TF RS subsetD +and forestI = forest_subset_TF RS subsetD; + +val TF_typechecks = + [TconsI, FnilI, FconsI, treeI, forestI, + list_of_TF_type, TF_map_type, TF_size_type, TF_preorder_type]; + +val TF_congs = TF.congs @ + mk_congs TF_Fn.thy + ["TF_rec", "list_of_TF", "TF_of_list", "TF_map", "TF_size", "TF_preorder"]; + +val TF_rewrites = + [TF_rec_Tcons, TF_rec_Fnil, TF_rec_Fcons, + list_of_TF_Tcons, list_of_TF_Fnil, list_of_TF_Fcons, + TF_of_list_Nil,TF_of_list_Cons, + TF_map_Tcons, TF_map_Fnil, TF_map_Fcons, + TF_size_Tcons, TF_size_Fnil, TF_size_Fcons, + TF_preorder_Tcons, TF_preorder_Fnil, TF_preorder_Fcons]; + +val TF_ss = list_ss addcongs TF_congs + addrews (TF_rewrites@TF_typechecks); + +(** theorems about list_of_TF and TF_of_list **) + +(*essentially the same as list induction*) +val major::prems = goal TF_Fn.thy + "[| tf: forest(A); \ +\ R(Fnil); \ +\ !!t tf. [| t: tree(A); tf: forest(A); R(tf) |] ==> R(Fcons(t,tf)) \ +\ |] ==> R(tf)"; +by (rtac (major RS (TF.mutual_induct RS conjunct2 RS spec RSN (2,rev_mp))) 1); +by (REPEAT (ares_tac (TrueI::prems) 1)); +val forest_induct = result(); + +goal TF_Fn.thy "!!tf A. tf: forest(A) ==> TF_of_list(list_of_TF(tf)) = tf"; +by (etac forest_induct 1); +by (ALLGOALS (ASM_SIMP_TAC TF_ss)); +val forest_iso = result(); + +goal TF_Fn.thy + "!!ts. ts: list(tree(A)) ==> list_of_TF(TF_of_list(ts)) = ts"; +by (etac List.induct 1); +by (ALLGOALS (ASM_SIMP_TAC TF_ss)); +val tree_list_iso = result(); + +(** theorems about TF_map **) + +goal TF_Fn.thy "!!z A. z: tree_forest(A) ==> TF_map(%u.u, z) = z"; +by (etac TF.induct 1); +by (ALLGOALS (ASM_SIMP_TAC TF_ss)); +val TF_map_ident = result(); + +goal TF_Fn.thy + "!!z A. z: tree_forest(A) ==> TF_map(h, TF_map(j,z)) = TF_map(%u.h(j(u)), z)"; +by (etac TF.induct 1); +by (ALLGOALS (ASM_SIMP_TAC TF_ss)); +val TF_map_compose = result(); + +(** theorems about TF_size **) + +goal TF_Fn.thy + "!!z A. z: tree_forest(A) ==> TF_size(TF_map(h,z)) = TF_size(z)"; +by (etac TF.induct 1); +by (ALLGOALS (ASM_SIMP_TAC TF_ss)); +val TF_size_TF_map = result(); + +goal TF_Fn.thy + "!!z A. z: tree_forest(A) ==> TF_size(z) = length(TF_preorder(z))"; +by (etac TF.induct 1); +by (ALLGOALS (ASM_SIMP_TAC (TF_ss addrews [length_app]))); +val TF_size_length = result(); + +(** theorems about TF_preorder **) + +goal TF_Fn.thy "!!z A. z: tree_forest(A) ==> \ +\ TF_preorder(TF_map(h,z)) = map(h, TF_preorder(z))"; +by (etac TF.induct 1); +by (ALLGOALS (ASM_SIMP_TAC (TF_ss addrews [map_app_distrib]))); +val TF_preorder_TF_map = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/tf_fn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/tf_fn.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,42 @@ +(* Title: ZF/ex/TF.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Trees & forests, a mutually recursive type definition. +*) + +TF_Fn = TF + ListFn + +consts + TF_rec :: "[i, [i,i,i]=>i, i, [i,i,i,i]=>i] => i" + TF_map :: "[i=>i, i] => i" + TF_size :: "i=>i" + TF_preorder :: "i=>i" + list_of_TF :: "i=>i" + TF_of_list :: "i=>i" + +rules + TF_rec_def + "TF_rec(z,b,c,d) == Vrec(z, \ +\ %z r. tree_forest_case(%x tf. b(x, tf, r`tf), \ +\ c, \ +\ %t tf. d(t, tf, r`t, r`tf), z))" + + list_of_TF_def + "list_of_TF(z) == TF_rec(z, %x tf r. [Tcons(x,tf)], [], \ +\ %t tf r1 r2. Cons(t, r2))" + + TF_of_list_def + "TF_of_list(tf) == list_rec(tf, Fnil, %t tf r. Fcons(t,r))" + + TF_map_def + "TF_map(h,z) == TF_rec(z, %x tf r.Tcons(h(x),r), Fnil, \ +\ %t tf r1 r2. Fcons(r1,r2))" + + TF_size_def + "TF_size(z) == TF_rec(z, %x tf r.succ(r), 0, %t tf r1 r2. r1#+r2)" + + TF_preorder_def + "TF_preorder(z) == TF_rec(z, %x tf r.Cons(x,r), Nil, %t tf r1 r2. r1@r2)" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/twos-compl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/twos-compl.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,103 @@ +(* Title: ZF/ex/twos-compl.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +ML code for Arithmetic on binary integers; the model for theory BinFn + + The sign Plus stands for an infinite string of leading 0's. + The sign Minus stands for an infinite string of leading 1's. + +A number can have multiple representations, namely leading 0's with sign +Plus and leading 1's with sign Minus. See int_of_binary for the numerical +interpretation. + +The representation expects that (m mod 2) is 0 or 1, even if m is negative; +For instance, ~5 div 2 = ~3 and ~5 mod 2 = 1; thus ~5 = (~3)*2 + 1 + +Still needs division! + +print_depth 40; +System.Control.Print.printDepth := 350; +*) + +infix 5 $$ + +(*Recursive datatype of binary integers*) +datatype bin = Plus | Minus | op $$ of bin * int; + +(** Conversions between bin and int **) + +fun int_of_binary Plus = 0 + | int_of_binary Minus = ~1 + | int_of_binary (w$$b) = 2 * int_of_binary w + b; + +fun binary_of_int 0 = Plus + | binary_of_int ~1 = Minus + | binary_of_int n = binary_of_int (n div 2) $$ (n mod 2); + +(*** Addition ***) + +(*Adding one to a number*) +fun bin_succ Plus = Plus$$1 + | bin_succ Minus = Plus + | bin_succ (w$$1) = bin_succ(w) $$ 0 + | bin_succ (w$$0) = w$$1; + +(*Subtracing one from a number*) +fun bin_pred Plus = Minus + | bin_pred Minus = Minus$$0 + | bin_pred (w$$1) = w$$0 + | bin_pred (w$$0) = bin_pred(w) $$ 1; + +(*sum of two binary integers*) +fun bin_add (Plus, w) = w + | bin_add (Minus, w) = bin_pred w + | bin_add (v$$x, Plus) = v$$x + | bin_add (v$$x, Minus) = bin_pred (v$$x) + | bin_add (v$$x, w$$y) = bin_add(v, if x+y=2 then bin_succ w else w) $$ + ((x+y) mod 2); + +(*** Subtraction ***) + +(*Unary minus*) +fun bin_minus Plus = Plus + | bin_minus Minus = Plus$$1 + | bin_minus (w$$1) = bin_pred (bin_minus(w) $$ 0) + | bin_minus (w$$0) = bin_minus(w) $$ 0; + +(*** Multiplication ***) + +(*product of two bins*) +fun bin_mult (Plus, _) = Plus + | bin_mult (Minus, v) = bin_minus v + | bin_mult (w$$1, v) = bin_add(bin_mult(w,v) $$ 0, v) + | bin_mult (w$$0, v) = bin_mult(w,v) $$ 0; + +(*** Testing ***) + +(*tests addition*) +fun checksum m n = + let val wm = binary_of_int m + and wn = binary_of_int n + val wsum = bin_add(wm,wn) + in if m+n = int_of_binary wsum then (wm, wn, wsum, m+n) + else raise Match + end; + +fun bfact n = if n=0 then Plus$$1 + else bin_mult(binary_of_int n, bfact(n-1)); + +(*Examples... +bfact 5; +int_of_binary it; +bfact 69; +int_of_binary it; + +(*leading zeros!*) +bin_add(binary_of_int 1234, binary_of_int ~1234); +bin_mult(binary_of_int 1234, Plus); + +(*leading ones!*) +bin_add(binary_of_int 1234, binary_of_int ~1235); +*) diff -r 000000000000 -r a5a9c433f639 src/ZF/ex/twos_compl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ex/twos_compl.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,103 @@ +(* Title: ZF/ex/twos-compl.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +ML code for Arithmetic on binary integers; the model for theory BinFn + + The sign Plus stands for an infinite string of leading 0's. + The sign Minus stands for an infinite string of leading 1's. + +A number can have multiple representations, namely leading 0's with sign +Plus and leading 1's with sign Minus. See int_of_binary for the numerical +interpretation. + +The representation expects that (m mod 2) is 0 or 1, even if m is negative; +For instance, ~5 div 2 = ~3 and ~5 mod 2 = 1; thus ~5 = (~3)*2 + 1 + +Still needs division! + +print_depth 40; +System.Control.Print.printDepth := 350; +*) + +infix 5 $$ + +(*Recursive datatype of binary integers*) +datatype bin = Plus | Minus | op $$ of bin * int; + +(** Conversions between bin and int **) + +fun int_of_binary Plus = 0 + | int_of_binary Minus = ~1 + | int_of_binary (w$$b) = 2 * int_of_binary w + b; + +fun binary_of_int 0 = Plus + | binary_of_int ~1 = Minus + | binary_of_int n = binary_of_int (n div 2) $$ (n mod 2); + +(*** Addition ***) + +(*Adding one to a number*) +fun bin_succ Plus = Plus$$1 + | bin_succ Minus = Plus + | bin_succ (w$$1) = bin_succ(w) $$ 0 + | bin_succ (w$$0) = w$$1; + +(*Subtracing one from a number*) +fun bin_pred Plus = Minus + | bin_pred Minus = Minus$$0 + | bin_pred (w$$1) = w$$0 + | bin_pred (w$$0) = bin_pred(w) $$ 1; + +(*sum of two binary integers*) +fun bin_add (Plus, w) = w + | bin_add (Minus, w) = bin_pred w + | bin_add (v$$x, Plus) = v$$x + | bin_add (v$$x, Minus) = bin_pred (v$$x) + | bin_add (v$$x, w$$y) = bin_add(v, if x+y=2 then bin_succ w else w) $$ + ((x+y) mod 2); + +(*** Subtraction ***) + +(*Unary minus*) +fun bin_minus Plus = Plus + | bin_minus Minus = Plus$$1 + | bin_minus (w$$1) = bin_pred (bin_minus(w) $$ 0) + | bin_minus (w$$0) = bin_minus(w) $$ 0; + +(*** Multiplication ***) + +(*product of two bins*) +fun bin_mult (Plus, _) = Plus + | bin_mult (Minus, v) = bin_minus v + | bin_mult (w$$1, v) = bin_add(bin_mult(w,v) $$ 0, v) + | bin_mult (w$$0, v) = bin_mult(w,v) $$ 0; + +(*** Testing ***) + +(*tests addition*) +fun checksum m n = + let val wm = binary_of_int m + and wn = binary_of_int n + val wsum = bin_add(wm,wn) + in if m+n = int_of_binary wsum then (wm, wn, wsum, m+n) + else raise Match + end; + +fun bfact n = if n=0 then Plus$$1 + else bin_mult(binary_of_int n, bfact(n-1)); + +(*Examples... +bfact 5; +int_of_binary it; +bfact 69; +int_of_binary it; + +(*leading zeros!*) +bin_add(binary_of_int 1234, binary_of_int ~1234); +bin_mult(binary_of_int 1234, Plus); + +(*leading ones!*) +bin_add(binary_of_int 1234, binary_of_int ~1235); +*) diff -r 000000000000 -r a5a9c433f639 src/ZF/fin.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/fin.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,104 @@ +(* Title: ZF/ex/fin.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Finite powerset operator + +could define cardinality? + +prove X:Fin(A) ==> EX n:nat. EX f. f:bij(X,n) + card(0)=0 + [| ~ a:b; b: Fin(A) |] ==> card(cons(a,b)) = succ(card(b)) + +b: Fin(A) ==> inj(b,b)<=surj(b,b) + +Limit(i) ==> Fin(Vfrom(A,i)) <= Un j:i. Fin(Vfrom(A,j)) +Fin(univ(A)) <= univ(A) +*) + +structure Fin = Inductive_Fun + (val thy = Arith.thy addconsts [(["Fin"],"i=>i")]; + val rec_doms = [("Fin","Pow(A)")]; + val sintrs = + ["0 : Fin(A)", + "[| a: A; b: Fin(A) |] ==> cons(a,b) : Fin(A)"]; + val monos = []; + val con_defs = []; + val type_intrs = [Pow_bottom, cons_subsetI, PowI] + val type_elims = [make_elim PowD]); + +val [Fin_0I, Fin_consI] = Fin.intrs; + + +goalw Fin.thy Fin.defs "!!A B. A<=B ==> Fin(A) <= Fin(B)"; +by (rtac lfp_mono 1); +by (REPEAT (rtac Fin.bnd_mono 1)); +by (REPEAT (ares_tac (Pow_mono::basic_monos) 1)); +val Fin_mono = result(); + +(* A : Fin(B) ==> A <= B *) +val FinD = Fin.dom_subset RS subsetD RS PowD; + +(** Induction on finite sets **) + +(*Discharging ~ x:y entails extra work*) +val major::prems = goal Fin.thy + "[| b: Fin(A); \ +\ P(0); \ +\ !!x y. [| x: A; y: Fin(A); ~ x:y; P(y) |] ==> P(cons(x,y)) \ +\ |] ==> P(b)"; +by (rtac (major RS Fin.induct) 1); +by (res_inst_tac [("Q","a:b")] (excluded_middle RS disjE) 2); +by (etac (cons_absorb RS ssubst) 3 THEN assume_tac 3); (*backtracking!*) +by (REPEAT (ares_tac prems 1)); +val Fin_induct = result(); + +(** Simplification for Fin **) +val Fin_ss = arith_ss addcongs mk_congs Fin.thy ["Fin"] addrews Fin.intrs; + +(*The union of two finite sets is fin*) +val major::prems = goal Fin.thy + "[| b: Fin(A); c: Fin(A) |] ==> b Un c : Fin(A)"; +by (rtac (major RS Fin_induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (Fin_ss addrews (prems@[Un_0, Un_cons])))); +val Fin_UnI = result(); + +(*The union of a set of finite sets is fin*) +val [major] = goal Fin.thy "C : Fin(Fin(A)) ==> Union(C) : Fin(A)"; +by (rtac (major RS Fin_induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (Fin_ss addrews [Union_0, Union_cons, Fin_UnI]))); +val Fin_UnionI = result(); + +(*Every subset of a finite set is fin*) +val [subs,fin] = goal Fin.thy "[| c<=b; b: Fin(A) |] ==> c: Fin(A)"; +by (EVERY1 [subgoal_tac "(ALL z. z<=b --> z: Fin(A))", + etac (spec RS mp), + rtac subs]); +by (rtac (fin RS Fin_induct) 1); +by (SIMP_TAC (Fin_ss addrews [subset_empty_iff]) 1); +by (safe_tac (ZF_cs addSDs [subset_cons_iff RS iffD1])); +by (eres_inst_tac [("b","z")] (cons_Diff RS subst) 2); +by (ALLGOALS (ASM_SIMP_TAC Fin_ss)); +val Fin_subset = result(); + +val major::prems = goal Fin.thy + "[| c: Fin(A); b: Fin(A); \ +\ P(b); \ +\ !!x y. [| x: A; y: Fin(A); x:y; P(y) |] ==> P(y-{x}) \ +\ |] ==> c<=b --> P(b-c)"; +by (rtac (major RS Fin_induct) 1); +by (rtac (Diff_cons RS ssubst) 2); +by (ALLGOALS (ASM_SIMP_TAC (Fin_ss addrews (prems@[Diff_0, cons_subset_iff, + Diff_subset RS Fin_subset])))); +val Fin_0_induct_lemma = result(); + +val prems = goal Fin.thy + "[| b: Fin(A); \ +\ P(b); \ +\ !!x y. [| x: A; y: Fin(A); x:y; P(y) |] ==> P(y-{x}) \ +\ |] ==> P(0)"; +by (rtac (Diff_cancel RS subst) 1); +by (rtac (Fin_0_induct_lemma RS mp) 1); +by (REPEAT (ares_tac (subset_refl::prems) 1)); +val Fin_0_induct = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/fixedpt.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/fixedpt.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,317 @@ +(* Title: ZF/fixedpt.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For fixedpt.thy. Least and greatest fixed points; the Knaster-Tarski Theorem + +Proved in the lattice of subsets of D, namely Pow(D), with Inter as glb +*) + +open Fixedpt; + +(*** Monotone operators ***) + +val prems = goalw Fixedpt.thy [bnd_mono_def] + "[| h(D)<=D; \ +\ !!W X. [| W<=D; X<=D; W<=X |] ==> h(W) <= h(X) \ +\ |] ==> bnd_mono(D,h)"; +by (REPEAT (ares_tac (prems@[conjI,allI,impI]) 1 + ORELSE etac subset_trans 1)); +val bnd_monoI = result(); + +val [major] = goalw Fixedpt.thy [bnd_mono_def] "bnd_mono(D,h) ==> h(D) <= D"; +by (rtac (major RS conjunct1) 1); +val bnd_monoD1 = result(); + +val major::prems = goalw Fixedpt.thy [bnd_mono_def] + "[| bnd_mono(D,h); W<=X; X<=D |] ==> h(W) <= h(X)"; +by (rtac (major RS conjunct2 RS spec RS spec RS mp RS mp) 1); +by (REPEAT (resolve_tac prems 1)); +val bnd_monoD2 = result(); + +val [major,minor] = goal Fixedpt.thy + "[| bnd_mono(D,h); X<=D |] ==> h(X) <= D"; +by (rtac (major RS bnd_monoD2 RS subset_trans) 1); +by (rtac (major RS bnd_monoD1) 3); +by (rtac minor 1); +by (rtac subset_refl 1); +val bnd_mono_subset = result(); + +goal Fixedpt.thy "!!A B. [| bnd_mono(D,h); A <= D; B <= D |] ==> \ +\ h(A) Un h(B) <= h(A Un B)"; +by (REPEAT (ares_tac [Un_upper1, Un_upper2, Un_least] 1 + ORELSE etac bnd_monoD2 1)); +val bnd_mono_Un = result(); + +(*Useful??*) +goal Fixedpt.thy "!!A B. [| bnd_mono(D,h); A <= D; B <= D |] ==> \ +\ h(A Int B) <= h(A) Int h(B)"; +by (REPEAT (ares_tac [Int_lower1, Int_lower2, Int_greatest] 1 + ORELSE etac bnd_monoD2 1)); +val bnd_mono_Int = result(); + +(**** Proof of Knaster-Tarski Theorem for the lfp ****) + +(*lfp is contained in each pre-fixedpoint*) +val prems = goalw Fixedpt.thy [lfp_def] + "[| h(A) <= A; A<=D |] ==> lfp(D,h) <= A"; +by (rtac (PowI RS CollectI RS Inter_lower) 1); +by (REPEAT (resolve_tac prems 1)); +val lfp_lowerbound = result(); + +(*Unfolding the defn of Inter dispenses with the premise bnd_mono(D,h)!*) +goalw Fixedpt.thy [lfp_def,Inter_def] "lfp(D,h) <= D"; +by (fast_tac ZF_cs 1); +val lfp_subset = result(); + +(*Used in datatype package*) +val [rew] = goal Fixedpt.thy "A==lfp(D,h) ==> A <= D"; +by (rewtac rew); +by (rtac lfp_subset 1); +val def_lfp_subset = result(); + +val subset0_cs = FOL_cs + addSIs [ballI, InterI, CollectI, PowI, empty_subsetI] + addIs [bexI, UnionI, ReplaceI, RepFunI] + addSEs [bexE, make_elim PowD, UnionE, ReplaceE, RepFunE, + CollectE, emptyE] + addEs [rev_ballE, InterD, make_elim InterD, subsetD]; + +val subset_cs = subset0_cs + addSIs [subset_refl,cons_subsetI,subset_consI,Union_least,UN_least,Un_least, + Inter_greatest,Int_greatest,RepFun_subset] + addSIs [Un_upper1,Un_upper2,Int_lower1,Int_lower2] + addIs [Union_upper,Inter_lower] + addSEs [cons_subsetE]; + +val prems = goalw Fixedpt.thy [lfp_def] + "[| h(D) <= D; !!X. [| h(X) <= X; X<=D |] ==> A<=X |] ==> \ +\ A <= lfp(D,h)"; +br (Pow_top RS CollectI RS Inter_greatest) 1; +by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [CollectE,PowD] 1)); +val lfp_greatest = result(); + +val hmono::prems = goal Fixedpt.thy + "[| bnd_mono(D,h); h(A)<=A; A<=D |] ==> h(lfp(D,h)) <= A"; +by (rtac (hmono RS bnd_monoD2 RS subset_trans) 1); +by (rtac lfp_lowerbound 1); +by (REPEAT (resolve_tac prems 1)); +val lfp_lemma1 = result(); + +val [hmono] = goal Fixedpt.thy + "bnd_mono(D,h) ==> h(lfp(D,h)) <= lfp(D,h)"; +by (rtac (bnd_monoD1 RS lfp_greatest) 1); +by (rtac lfp_lemma1 2); +by (REPEAT (ares_tac [hmono] 1)); +val lfp_lemma2 = result(); + +val [hmono] = goal Fixedpt.thy + "bnd_mono(D,h) ==> lfp(D,h) <= h(lfp(D,h))"; +by (rtac lfp_lowerbound 1); +by (rtac (hmono RS bnd_monoD2) 1); +by (rtac (hmono RS lfp_lemma2) 1); +by (rtac (hmono RS bnd_mono_subset) 2); +by (REPEAT (rtac lfp_subset 1)); +val lfp_lemma3 = result(); + +val prems = goal Fixedpt.thy + "bnd_mono(D,h) ==> lfp(D,h) = h(lfp(D,h))"; +by (REPEAT (resolve_tac (prems@[equalityI,lfp_lemma2,lfp_lemma3]) 1)); +val lfp_Tarski = result(); + +(*Definition form, to control unfolding*) +val [rew,mono] = goal Fixedpt.thy + "[| A==lfp(D,h); bnd_mono(D,h) |] ==> A = h(A)"; +by (rewtac rew); +by (rtac (mono RS lfp_Tarski) 1); +val def_lfp_Tarski = result(); + +(*** General induction rule for least fixedpoints ***) + +val [hmono,indstep] = goal Fixedpt.thy + "[| bnd_mono(D,h); !!x. x : h(Collect(lfp(D,h),P)) ==> P(x) \ +\ |] ==> h(Collect(lfp(D,h),P)) <= Collect(lfp(D,h),P)"; +by (rtac subsetI 1); +by (rtac CollectI 1); +by (etac indstep 2); +by (rtac (hmono RS lfp_lemma2 RS subsetD) 1); +by (rtac (hmono RS bnd_monoD2 RS subsetD) 1); +by (REPEAT (ares_tac [Collect_subset, lfp_subset] 1)); +val Collect_is_pre_fixedpt = result(); + +(*This rule yields an induction hypothesis in which the components of a + data structure may be assumed to be elements of lfp(D,h)*) +val prems = goal Fixedpt.thy + "[| bnd_mono(D,h); a : lfp(D,h); \ +\ !!x. x : h(Collect(lfp(D,h),P)) ==> P(x) \ +\ |] ==> P(a)"; +by (rtac (Collect_is_pre_fixedpt RS lfp_lowerbound RS subsetD RS CollectD2) 1); +by (rtac (lfp_subset RS (Collect_subset RS subset_trans)) 3); +by (REPEAT (ares_tac prems 1)); +val induct = result(); + +(*Definition form, to control unfolding*) +val rew::prems = goal Fixedpt.thy + "[| A == lfp(D,h); bnd_mono(D,h); a:A; \ +\ !!x. x : h(Collect(A,P)) ==> P(x) \ +\ |] ==> P(a)"; +by (rtac induct 1); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems) 1)); +val def_induct = result(); + +(*This version is useful when "A" is not a subset of D; + second premise could simply be h(D Int A) <= D or !!X. X<=D ==> h(X)<=D *) +val [hsub,hmono] = goal Fixedpt.thy + "[| h(D Int A) <= A; bnd_mono(D,h) |] ==> lfp(D,h) <= A"; +by (rtac (lfp_lowerbound RS subset_trans) 1); +by (rtac (hmono RS bnd_mono_subset RS Int_greatest) 1); +by (REPEAT (resolve_tac [hsub,Int_lower1,Int_lower2] 1)); +val lfp_Int_lowerbound = result(); + +(*Monotonicity of lfp, where h precedes i under a domain-like partial order + monotonicity of h is not strictly necessary; h must be bounded by D*) +val [hmono,imono,subhi] = goal Fixedpt.thy + "[| bnd_mono(D,h); bnd_mono(E,i); \ +\ !!X. X<=D ==> h(X) <= i(X) |] ==> lfp(D,h) <= lfp(E,i)"; +br (bnd_monoD1 RS lfp_greatest) 1; +br imono 1; +by (rtac (hmono RSN (2, lfp_Int_lowerbound)) 1); +by (rtac (Int_lower1 RS subhi RS subset_trans) 1); +by (rtac (imono RS bnd_monoD2 RS subset_trans) 1); +by (REPEAT (ares_tac [Int_lower2] 1)); +val lfp_mono = result(); + +(*This (unused) version illustrates that monotonicity is not really needed, + but both lfp's must be over the SAME set D; Inter is anti-monotonic!*) +val [isubD,subhi] = goal Fixedpt.thy + "[| i(D) <= D; !!X. X<=D ==> h(X) <= i(X) |] ==> lfp(D,h) <= lfp(D,i)"; +br lfp_greatest 1; +br isubD 1; +by (rtac lfp_lowerbound 1); +be (subhi RS subset_trans) 1; +by (REPEAT (assume_tac 1)); +val lfp_mono2 = result(); + + +(**** Proof of Knaster-Tarski Theorem for the gfp ****) + +(*gfp contains each post-fixedpoint that is contained in D*) +val prems = goalw Fixedpt.thy [gfp_def] + "[| A <= h(A); A<=D |] ==> A <= gfp(D,h)"; +by (rtac (PowI RS CollectI RS Union_upper) 1); +by (REPEAT (resolve_tac prems 1)); +val gfp_upperbound = result(); + +goalw Fixedpt.thy [gfp_def] "gfp(D,h) <= D"; +by (fast_tac ZF_cs 1); +val gfp_subset = result(); + +(*Used in datatype package*) +val [rew] = goal Fixedpt.thy "A==gfp(D,h) ==> A <= D"; +by (rewtac rew); +by (rtac gfp_subset 1); +val def_gfp_subset = result(); + +val hmono::prems = goalw Fixedpt.thy [gfp_def] + "[| bnd_mono(D,h); !!X. [| X <= h(X); X<=D |] ==> X<=A |] ==> \ +\ gfp(D,h) <= A"; +by (fast_tac (subset_cs addIs ((hmono RS bnd_monoD1)::prems)) 1); +val gfp_least = result(); + +val hmono::prems = goal Fixedpt.thy + "[| bnd_mono(D,h); A<=h(A); A<=D |] ==> A <= h(gfp(D,h))"; +by (rtac (hmono RS bnd_monoD2 RSN (2,subset_trans)) 1); +by (rtac gfp_subset 3); +by (rtac gfp_upperbound 2); +by (REPEAT (resolve_tac prems 1)); +val gfp_lemma1 = result(); + +val [hmono] = goal Fixedpt.thy + "bnd_mono(D,h) ==> gfp(D,h) <= h(gfp(D,h))"; +by (rtac gfp_least 1); +by (rtac gfp_lemma1 2); +by (REPEAT (ares_tac [hmono] 1)); +val gfp_lemma2 = result(); + +val [hmono] = goal Fixedpt.thy + "bnd_mono(D,h) ==> h(gfp(D,h)) <= gfp(D,h)"; +by (rtac gfp_upperbound 1); +by (rtac (hmono RS bnd_monoD2) 1); +by (rtac (hmono RS gfp_lemma2) 1); +by (REPEAT (rtac ([hmono, gfp_subset] MRS bnd_mono_subset) 1)); +val gfp_lemma3 = result(); + +val prems = goal Fixedpt.thy + "bnd_mono(D,h) ==> gfp(D,h) = h(gfp(D,h))"; +by (REPEAT (resolve_tac (prems@[equalityI,gfp_lemma2,gfp_lemma3]) 1)); +val gfp_Tarski = result(); + +(*Definition form, to control unfolding*) +val [rew,mono] = goal Fixedpt.thy + "[| A==gfp(D,h); bnd_mono(D,h) |] ==> A = h(A)"; +by (rewtac rew); +by (rtac (mono RS gfp_Tarski) 1); +val def_gfp_Tarski = result(); + + +(*** Coinduction rules for greatest fixed points ***) + +(*weak version*) +goal Fixedpt.thy "!!X h. [| a: X; X <= h(X); X <= D |] ==> a : gfp(D,h)"; +by (REPEAT (ares_tac [gfp_upperbound RS subsetD] 1)); +val weak_coinduct = result(); + +val [subs_h,subs_D,mono] = goal Fixedpt.thy + "[| X <= h(X Un gfp(D,h)); X <= D; bnd_mono(D,h) |] ==> \ +\ X Un gfp(D,h) <= h(X Un gfp(D,h))"; +by (rtac (subs_h RS Un_least) 1); +by (rtac (mono RS gfp_lemma2 RS subset_trans) 1); +by (rtac (Un_upper2 RS subset_trans) 1); +by (rtac ([mono, subs_D, gfp_subset] MRS bnd_mono_Un) 1); +val coinduct_lemma = result(); + +(*strong version*) +goal Fixedpt.thy + "!!X D. [| bnd_mono(D,h); a: X; X <= h(X Un gfp(D,h)); X <= D |] ==> \ +\ a : gfp(D,h)"; +by (rtac (coinduct_lemma RSN (2, weak_coinduct)) 1); +by (REPEAT (ares_tac [gfp_subset, UnI1, Un_least] 1)); +val coinduct = result(); + +(*Definition form, to control unfolding*) +val rew::prems = goal Fixedpt.thy + "[| A == gfp(D,h); bnd_mono(D,h); a: X; X <= h(X Un A); X <= D |] ==> \ +\ a : A"; +by (rewtac rew); +by (rtac coinduct 1); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems) 1)); +val def_coinduct = result(); + +(*Lemma used immediately below!*) +val [subsA,XimpP] = goal ZF.thy + "[| X <= A; !!z. z:X ==> P(z) |] ==> X <= Collect(A,P)"; +by (rtac (subsA RS subsetD RS CollectI RS subsetI) 1); +by (assume_tac 1); +by (etac XimpP 1); +val subset_Collect = result(); + +(*The version used in the induction/coinduction package*) +val prems = goal Fixedpt.thy + "[| A == gfp(D, %w. Collect(D,P(w))); bnd_mono(D, %w. Collect(D,P(w))); \ +\ a: X; X <= D; !!z. z: X ==> P(X Un A, z) |] ==> \ +\ a : A"; +by (rtac def_coinduct 1); +by (REPEAT (ares_tac (subset_Collect::prems) 1)); +val def_Collect_coinduct = result(); + +(*Monotonicity of gfp!*) +val [hmono,subde,subhi] = goal Fixedpt.thy + "[| bnd_mono(D,h); D <= E; \ +\ !!X. X<=D ==> h(X) <= i(X) |] ==> gfp(D,h) <= gfp(E,i)"; +by (rtac gfp_upperbound 1); +by (rtac (hmono RS gfp_lemma2 RS subset_trans) 1); +by (rtac (gfp_subset RS subhi) 1); +by (rtac ([gfp_subset, subde] MRS subset_trans) 1); +val gfp_mono = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/fixedpt.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/fixedpt.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,23 @@ +(* Title: ZF/fixedpt.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Least and greatest fixed points +*) + +Fixedpt = ZF + +consts + bnd_mono :: "[i,i=>i]=>o" + lfp, gfp :: "[i,i=>i]=>i" + +rules + (*monotone operator from Pow(D) to itself*) + bnd_mono_def + "bnd_mono(D,h) == h(D)<=D & (ALL W X. W<=X --> X<=D --> h(W) <= h(X))" + + lfp_def "lfp(D,h) == Inter({X: Pow(D). h(X) <= X})" + + gfp_def "gfp(D,h) == Union({X: Pow(D). X <= h(X)})" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/func.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/func.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,367 @@ +(* Title: ZF/func + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Functions in Zermelo-Fraenkel Set Theory +*) + +(*** The Pi operator -- dependent function space ***) + +val prems = goalw ZF.thy [Pi_def] + "[| f <= Sigma(A,B); !!x. x:A ==> EX! y. : f |] ==> \ +\ f: Pi(A,B)"; +by (REPEAT (ares_tac (prems @ [CollectI,PowI,ballI,impI]) 1)); +val PiI = result(); + +(**Two "destruct" rules for Pi **) + +val [major] = goalw ZF.thy [Pi_def] "f: Pi(A,B) ==> f <= Sigma(A,B)"; +by (rtac (major RS CollectE) 1); +by (etac PowD 1); +val fun_is_rel = result(); + +val major::prems = goalw ZF.thy [Pi_def] + "[| f: Pi(A,B); a:A |] ==> EX! y. : f"; +by (rtac (major RS CollectE) 1); +by (etac bspec 1); +by (resolve_tac prems 1); +val fun_unique_Pair = result(); + +val prems = goal ZF.thy + "[| f: Pi(A,B); \ +\ [| f <= Sigma(A,B); ALL x:A. EX! y. : f |] ==> P \ +\ |] ==> P"; +by (REPEAT (ares_tac (prems@[ballI,fun_is_rel,fun_unique_Pair]) 1)); +val PiE = result(); + +val prems = goalw ZF.thy [Pi_def] + "[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> Pi(A,B) = Pi(A',B')"; +by (prove_cong_tac (prems@[Collect_cong,Sigma_cong,ball_cong]) 1); +val Pi_cong = result(); + +(*Weaking one function type to another*) +goalw ZF.thy [Pi_def] "!!f. [| f: A->B; B<=D |] ==> f: A->D"; +by (safe_tac ZF_cs); +by (set_mp_tac 1); +by (fast_tac ZF_cs 1); +val fun_weaken_type = result(); + +(*Empty function spaces*) +goalw ZF.thy [Pi_def] "Pi(0,A) = {0}"; +by (fast_tac (ZF_cs addIs [equalityI]) 1); +val Pi_empty1 = result(); + +goalw ZF.thy [Pi_def] "!!A a. a:A ==> A->0 = 0"; +by (fast_tac (ZF_cs addIs [equalityI]) 1); +val Pi_empty2 = result(); + + +(*** Function Application ***) + +goal ZF.thy "!!a b f. [| : f; : f; f: Pi(A,B) |] ==> b=c"; +by (etac PiE 1); +by (etac (bspec RS ex1_equalsE) 1); +by (etac (subsetD RS SigmaD1) 1); +by (REPEAT (assume_tac 1)); +val apply_equality2 = result(); + +goalw ZF.thy [apply_def] "!!a b f. [| : f; f: Pi(A,B) |] ==> f`a = b"; +by (rtac the_equality 1); +by (rtac apply_equality2 2); +by (REPEAT (assume_tac 1)); +val apply_equality = result(); + +val prems = goal ZF.thy + "[| f: Pi(A,B); c: f; !!x. [| x:A; c = |] ==> P \ +\ |] ==> P"; +by (cut_facts_tac prems 1); +by (etac (fun_is_rel RS subsetD RS SigmaE) 1); +by (REPEAT (ares_tac prems 1)); +by (hyp_subst_tac 1); +by (etac (apply_equality RS ssubst) 1); +by (resolve_tac prems 1); +by (rtac refl 1); +val memberPiE = result(); + +(*Conclusion is flexible -- use res_inst_tac or else RS with a theorem + of the form f:A->B *) +goal ZF.thy "!!f. [| f: Pi(A,B); a:A |] ==> f`a : B(a)"; +by (rtac (fun_unique_Pair RS ex1E) 1); +by (REPEAT (assume_tac 1)); +by (rtac (fun_is_rel RS subsetD RS SigmaE2) 1); +by (etac (apply_equality RS ssubst) 3); +by (REPEAT (assume_tac 1)); +val apply_type = result(); + +goal ZF.thy "!!f. [| f: Pi(A,B); a:A |] ==> : f"; +by (rtac (fun_unique_Pair RS ex1E) 1); +by (resolve_tac [apply_equality RS ssubst] 3); +by (REPEAT (assume_tac 1)); +val apply_Pair = result(); + +val [major] = goal ZF.thy + "f: Pi(A,B) ==> : f <-> a:A & f`a = b"; +by (rtac (major RS PiE) 1); +by (fast_tac (ZF_cs addSIs [major RS apply_Pair, + major RSN (2,apply_equality)]) 1); +val apply_iff = result(); + +(*Refining one Pi type to another*) +val prems = goal ZF.thy + "[| f: Pi(A,C); !!x. x:A ==> f`x : B(x) |] ==> f : Pi(A,B)"; +by (rtac (subsetI RS PiI) 1); +by (eresolve_tac (prems RL [memberPiE]) 1); +by (etac ssubst 1); +by (REPEAT (ares_tac (prems@[SigmaI,fun_unique_Pair]) 1)); +val Pi_type = result(); + + +(** Elimination of membership in a function **) + +goal ZF.thy "!!a A. [| : f; f: Pi(A,B) |] ==> a : A"; +by (REPEAT (ares_tac [fun_is_rel RS subsetD RS SigmaD1] 1)); +val domain_type = result(); + +goal ZF.thy "!!b B a. [| : f; f: Pi(A,B) |] ==> b : B(a)"; +by (etac (fun_is_rel RS subsetD RS SigmaD2) 1); +by (assume_tac 1); +val range_type = result(); + +val prems = goal ZF.thy + "[| : f; f: Pi(A,B); \ +\ [| a:A; b:B(a); f`a = b |] ==> P \ +\ |] ==> P"; +by (cut_facts_tac prems 1); +by (resolve_tac prems 1); +by (REPEAT (eresolve_tac [asm_rl,domain_type,range_type,apply_equality] 1)); +val Pair_mem_PiE = result(); + +(*** Lambda Abstraction ***) + +goalw ZF.thy [lam_def] "!!A b. a:A ==> : (lam x:A. b(x))"; +by (etac RepFunI 1); +val lamI = result(); + +val major::prems = goalw ZF.thy [lam_def] + "[| p: (lam x:A. b(x)); !!x.[| x:A; p= |] ==> P \ +\ |] ==> P"; +by (rtac (major RS RepFunE) 1); +by (REPEAT (ares_tac prems 1)); +val lamE = result(); + +goal ZF.thy "!!a b c. [| : (lam x:A. b(x)) |] ==> c = b(a)"; +by (REPEAT (eresolve_tac [asm_rl,lamE,Pair_inject,ssubst] 1)); +val lamD = result(); + +val prems = goalw ZF.thy [lam_def] + "[| !!x. x:A ==> b(x): B(x) |] ==> (lam x:A.b(x)) : Pi(A,B)"; +by (fast_tac (ZF_cs addIs (PiI::prems)) 1); +val lam_type = result(); + +goal ZF.thy "(lam x:A.b(x)) : A -> {b(x). x:A}"; +by (REPEAT (ares_tac [refl,lam_type,RepFunI] 1)); +val lam_funtype = result(); + +goal ZF.thy "!!a A. a : A ==> (lam x:A.b(x)) ` a = b(a)"; +by (REPEAT (ares_tac [apply_equality,lam_funtype,lamI] 1)); +val beta = result(); + +(*congruence rule for lambda abstraction*) +val prems = goalw ZF.thy [lam_def] + "[| A=A'; !!x. x:A' ==> b(x)=b'(x) |] ==> \ +\ (lam x:A.b(x)) = (lam x:A'.b'(x))"; +by (rtac RepFun_cong 1); +by (res_inst_tac [("t","Pair")] subst_context2 2); +by (REPEAT (ares_tac (prems@[refl]) 1)); +val lam_cong = result(); + +val [major] = goal ZF.thy + "(!!x. x:A ==> EX! y. Q(x,y)) ==> EX f. ALL x:A. Q(x, f`x)"; +by (res_inst_tac [("x", "lam x: A. THE y. Q(x,y)")] exI 1); +by (rtac ballI 1); +by (rtac (beta RS ssubst) 1); +by (assume_tac 1); +by (etac (major RS theI) 1); +val lam_theI = result(); + + +(** Extensionality **) + +(*Semi-extensionality!*) +val prems = goal ZF.thy + "[| f : Pi(A,B); g: Pi(C,D); A<=C; \ +\ !!x. x:A ==> f`x = g`x |] ==> f<=g"; +by (rtac subsetI 1); +by (eresolve_tac (prems RL [memberPiE]) 1); +by (etac ssubst 1); +by (resolve_tac (prems RL [ssubst]) 1); +by (REPEAT (ares_tac (prems@[apply_Pair,subsetD]) 1)); +val fun_subset = result(); + +val prems = goal ZF.thy + "[| f : Pi(A,B); g: Pi(A,D); \ +\ !!x. x:A ==> f`x = g`x |] ==> f=g"; +by (REPEAT (ares_tac (prems @ (prems RL [sym]) @ + [subset_refl,equalityI,fun_subset]) 1)); +val fun_extension = result(); + +goal ZF.thy "!!f A B. f : Pi(A,B) ==> (lam x:A. f`x) = f"; +by (rtac fun_extension 1); +by (REPEAT (ares_tac [lam_type,apply_type,beta] 1)); +val eta = result(); + +(*Every element of Pi(A,B) may be expressed as a lambda abstraction!*) +val prems = goal ZF.thy + "[| f: Pi(A,B); \ +\ !!b. [| ALL x:A. b(x):B(x); f = (lam x:A.b(x)) |] ==> P \ +\ |] ==> P"; +by (resolve_tac prems 1); +by (rtac (eta RS sym) 2); +by (REPEAT (ares_tac (prems@[ballI,apply_type]) 1)); +val Pi_lamE = result(); + + +(*** properties of "restrict" ***) + +goalw ZF.thy [restrict_def,lam_def] + "!!f A. [| f: Pi(C,B); A<=C |] ==> restrict(f,A) <= f"; +by (fast_tac (ZF_cs addIs [apply_Pair]) 1); +val restrict_subset = result(); + +val prems = goalw ZF.thy [restrict_def] + "[| !!x. x:A ==> f`x: B(x) |] ==> restrict(f,A) : Pi(A,B)"; +by (rtac lam_type 1); +by (eresolve_tac prems 1); +val restrict_type = result(); + +val [pi,subs] = goal ZF.thy + "[| f: Pi(C,B); A<=C |] ==> restrict(f,A) : Pi(A,B)"; +by (rtac (pi RS apply_type RS restrict_type) 1); +by (etac (subs RS subsetD) 1); +val restrict_type2 = result(); + +goalw ZF.thy [restrict_def] "!!a A. a : A ==> restrict(f,A) ` a = f`a"; +by (etac beta 1); +val restrict = result(); + +(*NOT SAFE as a congruence rule for the simplifier! Can cause it to fail!*) +val prems = goalw ZF.thy [restrict_def] + "[| A=B; !!x. x:B ==> f`x=g`x |] ==> restrict(f,A) = restrict(g,B)"; +by (REPEAT (ares_tac (prems@[lam_cong]) 1)); +val restrict_eqI = result(); + +goalw ZF.thy [restrict_def] "domain(restrict(f,C)) = C"; +by (REPEAT (ares_tac [equalityI,subsetI,domainI,lamI] 1 + ORELSE eresolve_tac [domainE,lamE,Pair_inject,ssubst] 1)); +val domain_restrict = result(); + +val [prem] = goalw ZF.thy [restrict_def] + "A<=C ==> restrict(lam x:C. b(x), A) = (lam x:A. b(x))"; +by (rtac (refl RS lam_cong) 1); +be (prem RS subsetD RS beta) 1; (*easier than calling SIMP_TAC*) +val restrict_lam_eq = result(); + + + +(*** Unions of functions ***) + +(** The Union of a set of COMPATIBLE functions is a function **) +val [ex_prem,disj_prem] = goal ZF.thy + "[| ALL x:S. EX C D. x:C->D; \ +\ !!x y. [| x:S; y:S |] ==> x<=y | y<=x |] ==> \ +\ Union(S) : domain(Union(S)) -> range(Union(S))"; +val premE = ex_prem RS bspec RS exE; +by (REPEAT (eresolve_tac [exE,PiE,premE] 1 + ORELSE ares_tac [PiI, ballI RS rel_Union, exI] 1)); +by (REPEAT (eresolve_tac [asm_rl,domainE,UnionE,exE] 1 + ORELSE ares_tac [allI,impI,ex1I,UnionI] 1)); +by (res_inst_tac [ ("x1","B") ] premE 1); +by (res_inst_tac [ ("x1","Ba") ] premE 2); +by (REPEAT (eresolve_tac [asm_rl,exE] 1)); +by (eresolve_tac [disj_prem RS disjE] 1); +by (DEPTH_SOLVE (set_mp_tac 1 + ORELSE eresolve_tac [asm_rl, apply_equality2] 1)); +val fun_Union = result(); + + +(** The Union of 2 disjoint functions is a function **) + +val prems = goal ZF.thy + "[| f: A->B; g: C->D; A Int C = 0 |] ==> \ +\ (f Un g) : (A Un C) -> (B Un D)"; + (*Contradiction if A Int C = 0, a:A, a:B*) +val [disjoint] = prems RL ([IntI] RLN (2, [equals0D])); +by (cut_facts_tac prems 1); +by (rtac PiI 1); +(*solve subgoal 2 first!!*) +by (DEPTH_SOLVE_1 (eresolve_tac [UnE, Pair_mem_PiE, sym, disjoint] 2 + INTLEAVE ares_tac [ex1I, apply_Pair RS UnI1, apply_Pair RS UnI2] 2)); +by (REPEAT (eresolve_tac [asm_rl,UnE,rel_Un,PiE] 1)); +val fun_disjoint_Un = result(); + +goal ZF.thy + "!!f g a. [| a:A; f: A->B; g: C->D; A Int C = 0 |] ==> \ +\ (f Un g)`a = f`a"; +by (REPEAT (ares_tac [apply_equality,UnI1,apply_Pair, + fun_disjoint_Un] 1)); +val fun_disjoint_apply1 = result(); + +goal ZF.thy + "!!f g c. [| c:C; f: A->B; g: C->D; A Int C = 0 |] ==> \ +\ (f Un g)`c = g`c"; +by (REPEAT (ares_tac [apply_equality,UnI2,apply_Pair, + fun_disjoint_Un] 1)); +val fun_disjoint_apply2 = result(); + +(** Domain and range of a function/relation **) + +val [major] = goal ZF.thy "f : Pi(A,B) ==> domain(f)=A"; +by (rtac equalityI 1); +by (fast_tac (ZF_cs addIs [major RS apply_Pair]) 2); +by (rtac (major RS PiE) 1); +by (fast_tac ZF_cs 1); +val domain_of_fun = result(); + +val [major] = goal ZF.thy "f : Pi(A,B) ==> f : A->range(f)"; +by (rtac (major RS Pi_type) 1); +by (etac (major RS apply_Pair RS rangeI) 1); +val range_of_fun = result(); + +(*** Extensions of functions ***) + +(*Singleton function -- in the underlying form of singletons*) +goal ZF.thy "Upair(,) : Upair(a,a) -> Upair(b,b)"; +by (fast_tac (ZF_cs addIs [PiI]) 1); +val fun_single_lemma = result(); + +goalw ZF.thy [cons_def] + "!!f A B. [| f: A->B; ~c:A |] ==> cons(,f) : cons(c,A) -> cons(b,B)"; +by (rtac (fun_single_lemma RS fun_disjoint_Un) 1); +by (assume_tac 1); +by (rtac equals0I 1); +by (fast_tac ZF_cs 1); +val fun_extend = result(); + +goal ZF.thy "!!f A B. [| f: A->B; a:A; ~ c:A |] ==> cons(,f)`a = f`a"; +by (rtac (apply_Pair RS consI2 RS apply_equality) 1); +by (rtac fun_extend 3); +by (REPEAT (assume_tac 1)); +val fun_extend_apply1 = result(); + +goal ZF.thy "!!f A B. [| f: A->B; ~ c:A |] ==> cons(,f)`c = b"; +by (rtac (consI1 RS apply_equality) 1); +by (rtac fun_extend 1); +by (REPEAT (assume_tac 1)); +val fun_extend_apply2 = result(); + +(*The empty function*) +goal ZF.thy "0: 0->A"; +by (fast_tac (ZF_cs addIs [PiI]) 1); +val fun_empty = result(); + +(*The singleton function*) +goal ZF.thy "{} : {a} -> cons(b,C)"; +by (REPEAT (ares_tac [fun_extend,fun_empty,notI] 1 ORELSE etac emptyE 1)); +val fun_single = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/ind-syntax.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ind-syntax.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,175 @@ +(* Title: ZF/ind-syntax.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Abstract Syntax functions for Inductive Definitions +*) + + +(*SHOULD BE ABLE TO DELETE THESE!*) +fun flatten_typ sign T = + let val {syn,...} = Sign.rep_sg sign + in Pretty.str_of (Syntax.pretty_typ syn T) + end; +fun flatten_term sign t = Pretty.str_of (Sign.pretty_term sign t); + +(*Add constants to a theory*) +infix addconsts; +fun thy addconsts const_decs = + extend_theory thy (space_implode "_" (flat (map #1 const_decs)) + ^ "_Theory") + ([], [], [], [], const_decs, None) []; + + +(*Make a definition, lhs==rhs, checking that vars on lhs contain *) +fun mk_defpair sign (lhs,rhs) = + let val Const(name,_) = head_of lhs + val dummy = assert (term_vars rhs subset term_vars lhs + andalso + term_frees rhs subset term_frees lhs + andalso + term_tvars rhs subset term_tvars lhs + andalso + term_tfrees rhs subset term_tfrees lhs) + ("Extra variables on RHS in definition of " ^ name) + in (name ^ "_def", + flatten_term sign (Logic.mk_equals (lhs,rhs))) + end; + +(*export to Pure/sign? Used in Provers/simp.ML...*) +fun lookup_const sign a = Symtab.lookup(#const_tab (Sign.rep_sg sign), a); + +(*export to Pure/library? *) +fun assert_all pred l msg_fn = + let fun asl [] = () + | asl (x::xs) = if pred x then asl xs + else error (msg_fn x) + in asl l end; + + +(** Abstract syntax definitions for FOL and ZF **) + +val iT = Type("i",[]) +and oT = Type("o",[]); + +fun ap t u = t$u; +fun app t (u1,u2) = t $ u1 $ u2; + +(*Given u expecting arguments of types [T1,...,Tn], create term of + type T1*...*Tn => i using split*) +fun ap_split split u [ ] = Abs("null", iT, u) + | ap_split split u [_] = u + | ap_split split u [_,_] = split $ u + | ap_split split u (T::Ts) = + split $ (Abs("v", T, ap_split split (u $ Bound(length Ts - 2)) Ts)); + +val conj = Const("op &", [oT,oT]--->oT) +and disj = Const("op |", [oT,oT]--->oT) +and imp = Const("op -->", [oT,oT]--->oT); + +val eq_const = Const("op =", [iT,iT]--->oT); + +val mem_const = Const("op :", [iT,iT]--->oT); + +val exists_const = Const("Ex", [iT-->oT]--->oT); +fun mk_exists (Free(x,T),P) = exists_const $ (absfree (x,T,P)); + +val all_const = Const("All", [iT-->oT]--->oT); +fun mk_all (Free(x,T),P) = all_const $ (absfree (x,T,P)); + +(*Creates All(%v.v:A --> P(v)) rather than Ball(A,P) *) +fun mk_all_imp (A,P) = + all_const $ Abs("v", iT, imp $ (mem_const $ Bound 0 $ A) $ (P $ Bound 0)); + + +val Part_const = Const("Part", [iT,iT-->iT]--->iT); + +val Collect_const = Const("Collect", [iT,iT-->oT]--->iT); +fun mk_Collect (a,D,t) = Collect_const $ D $ absfree(a, iT, t); + +val Trueprop = Const("Trueprop",oT-->propT); +fun mk_tprop P = Trueprop $ P; +fun dest_tprop (Const("Trueprop",_) $ P) = P; + +(*** Tactic for folding constructor definitions ***) + +(*The depth of injections in a constructor function*) +fun inject_depth (Const _ $ t) = 1 + inject_depth t + | inject_depth t = 0; + +val rhs_of_thm = #2 o Logic.dest_equals o #prop o rep_thm; + +(*There are critical pairs! E.g. K == Inl(0), S == Inr(Inl(0)) + Folds longest definitions first to avoid folding subexpressions of an rhs.*) +fun fold_con_tac defs = + let val keylist = make_keylist (inject_depth o rhs_of_thm) defs; + val keys = distinct (sort op> (map #2 keylist)); + val deflists = map (keyfilter keylist) keys + in EVERY (map fold_tac deflists) end; + +(*Prove a goal stated as a term, with exception handling*) +fun prove_term sign defs (P,tacsf) = + let val ct = Sign.cterm_of sign P + in prove_goalw_cterm defs ct tacsf + handle e => (writeln ("Exception in proof of\n" ^ + Sign.string_of_cterm ct); + raise e) + end; + +(*Read an assumption in the given theory*) +fun assume_read thy a = assume (Sign.read_cterm (sign_of thy) (a,propT)); + +(*Make distinct individual variables a1, a2, a3, ..., an. *) +fun mk_frees a [] = [] + | mk_frees a (T::Ts) = Free(a,T) :: mk_frees (bump_string a) Ts; + +(*Used by intr-elim.ML and in individual datatype definitions*) +val basic_monos = [subset_refl, imp_refl, disj_mono, conj_mono, + ex_mono, Collect_mono, Part_mono, in_mono]; + +fun rule_concl rl = + let val Const("op :",_) $ t $ X = dest_tprop (Logic.strip_imp_concl rl) + in (t,X) end + handle _ => error "Conclusion of rule should be a set membership"; + +(*For deriving cases rules. CollectD2 discards the domain, which is redundant; + read_instantiate replaces a propositional variable by a formula variable*) +val equals_CollectD = + read_instantiate [("W","?Q")] + (make_elim (equalityD1 RS subsetD RS CollectD2)); + + +(*From HOL/ex/meson.ML: raises exception if no rules apply -- unlike RL*) +fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls)) + | tryres (th, []) = raise THM("tryres", 0, [th]); + +fun gen_make_elim elim_rls rl = + standard (tryres (rl, elim_rls @ [revcut_rl])); + +(** For constructor.ML **) + +(*Avoids duplicate definitions by removing constants already declared mixfix*) +fun remove_mixfixes None decs = decs + | remove_mixfixes (Some sext) decs = + let val mixtab = Symtab.st_of_declist(Syntax.constants sext, Symtab.null) + fun is_mix c = case Symtab.lookup(mixtab,c) of + None=>false | Some _ => true + in map (fn (cs,styp)=> (filter_out is_mix cs, styp)) decs + end; + +fun ext_constants None = [] + | ext_constants (Some sext) = Syntax.constants sext; + + +(*Could go to FOL, but it's hardly general*) +val [def] = goal IFOL.thy "a==b ==> a=c <-> c=b"; +by (rewtac def); +by (rtac iffI 1); +by (REPEAT (etac sym 1)); +val def_swap_iff = result(); + +val def_trans = prove_goal IFOL.thy "[| f==g; g(a)=b |] ==> f(a)=b" + (fn [rew,prem] => [ rewtac rew, rtac prem 1 ]); + + diff -r 000000000000 -r a5a9c433f639 src/ZF/ind_syntax.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ind_syntax.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,175 @@ +(* Title: ZF/ind-syntax.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Abstract Syntax functions for Inductive Definitions +*) + + +(*SHOULD BE ABLE TO DELETE THESE!*) +fun flatten_typ sign T = + let val {syn,...} = Sign.rep_sg sign + in Pretty.str_of (Syntax.pretty_typ syn T) + end; +fun flatten_term sign t = Pretty.str_of (Sign.pretty_term sign t); + +(*Add constants to a theory*) +infix addconsts; +fun thy addconsts const_decs = + extend_theory thy (space_implode "_" (flat (map #1 const_decs)) + ^ "_Theory") + ([], [], [], [], const_decs, None) []; + + +(*Make a definition, lhs==rhs, checking that vars on lhs contain *) +fun mk_defpair sign (lhs,rhs) = + let val Const(name,_) = head_of lhs + val dummy = assert (term_vars rhs subset term_vars lhs + andalso + term_frees rhs subset term_frees lhs + andalso + term_tvars rhs subset term_tvars lhs + andalso + term_tfrees rhs subset term_tfrees lhs) + ("Extra variables on RHS in definition of " ^ name) + in (name ^ "_def", + flatten_term sign (Logic.mk_equals (lhs,rhs))) + end; + +(*export to Pure/sign? Used in Provers/simp.ML...*) +fun lookup_const sign a = Symtab.lookup(#const_tab (Sign.rep_sg sign), a); + +(*export to Pure/library? *) +fun assert_all pred l msg_fn = + let fun asl [] = () + | asl (x::xs) = if pred x then asl xs + else error (msg_fn x) + in asl l end; + + +(** Abstract syntax definitions for FOL and ZF **) + +val iT = Type("i",[]) +and oT = Type("o",[]); + +fun ap t u = t$u; +fun app t (u1,u2) = t $ u1 $ u2; + +(*Given u expecting arguments of types [T1,...,Tn], create term of + type T1*...*Tn => i using split*) +fun ap_split split u [ ] = Abs("null", iT, u) + | ap_split split u [_] = u + | ap_split split u [_,_] = split $ u + | ap_split split u (T::Ts) = + split $ (Abs("v", T, ap_split split (u $ Bound(length Ts - 2)) Ts)); + +val conj = Const("op &", [oT,oT]--->oT) +and disj = Const("op |", [oT,oT]--->oT) +and imp = Const("op -->", [oT,oT]--->oT); + +val eq_const = Const("op =", [iT,iT]--->oT); + +val mem_const = Const("op :", [iT,iT]--->oT); + +val exists_const = Const("Ex", [iT-->oT]--->oT); +fun mk_exists (Free(x,T),P) = exists_const $ (absfree (x,T,P)); + +val all_const = Const("All", [iT-->oT]--->oT); +fun mk_all (Free(x,T),P) = all_const $ (absfree (x,T,P)); + +(*Creates All(%v.v:A --> P(v)) rather than Ball(A,P) *) +fun mk_all_imp (A,P) = + all_const $ Abs("v", iT, imp $ (mem_const $ Bound 0 $ A) $ (P $ Bound 0)); + + +val Part_const = Const("Part", [iT,iT-->iT]--->iT); + +val Collect_const = Const("Collect", [iT,iT-->oT]--->iT); +fun mk_Collect (a,D,t) = Collect_const $ D $ absfree(a, iT, t); + +val Trueprop = Const("Trueprop",oT-->propT); +fun mk_tprop P = Trueprop $ P; +fun dest_tprop (Const("Trueprop",_) $ P) = P; + +(*** Tactic for folding constructor definitions ***) + +(*The depth of injections in a constructor function*) +fun inject_depth (Const _ $ t) = 1 + inject_depth t + | inject_depth t = 0; + +val rhs_of_thm = #2 o Logic.dest_equals o #prop o rep_thm; + +(*There are critical pairs! E.g. K == Inl(0), S == Inr(Inl(0)) + Folds longest definitions first to avoid folding subexpressions of an rhs.*) +fun fold_con_tac defs = + let val keylist = make_keylist (inject_depth o rhs_of_thm) defs; + val keys = distinct (sort op> (map #2 keylist)); + val deflists = map (keyfilter keylist) keys + in EVERY (map fold_tac deflists) end; + +(*Prove a goal stated as a term, with exception handling*) +fun prove_term sign defs (P,tacsf) = + let val ct = Sign.cterm_of sign P + in prove_goalw_cterm defs ct tacsf + handle e => (writeln ("Exception in proof of\n" ^ + Sign.string_of_cterm ct); + raise e) + end; + +(*Read an assumption in the given theory*) +fun assume_read thy a = assume (Sign.read_cterm (sign_of thy) (a,propT)); + +(*Make distinct individual variables a1, a2, a3, ..., an. *) +fun mk_frees a [] = [] + | mk_frees a (T::Ts) = Free(a,T) :: mk_frees (bump_string a) Ts; + +(*Used by intr-elim.ML and in individual datatype definitions*) +val basic_monos = [subset_refl, imp_refl, disj_mono, conj_mono, + ex_mono, Collect_mono, Part_mono, in_mono]; + +fun rule_concl rl = + let val Const("op :",_) $ t $ X = dest_tprop (Logic.strip_imp_concl rl) + in (t,X) end + handle _ => error "Conclusion of rule should be a set membership"; + +(*For deriving cases rules. CollectD2 discards the domain, which is redundant; + read_instantiate replaces a propositional variable by a formula variable*) +val equals_CollectD = + read_instantiate [("W","?Q")] + (make_elim (equalityD1 RS subsetD RS CollectD2)); + + +(*From HOL/ex/meson.ML: raises exception if no rules apply -- unlike RL*) +fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls)) + | tryres (th, []) = raise THM("tryres", 0, [th]); + +fun gen_make_elim elim_rls rl = + standard (tryres (rl, elim_rls @ [revcut_rl])); + +(** For constructor.ML **) + +(*Avoids duplicate definitions by removing constants already declared mixfix*) +fun remove_mixfixes None decs = decs + | remove_mixfixes (Some sext) decs = + let val mixtab = Symtab.st_of_declist(Syntax.constants sext, Symtab.null) + fun is_mix c = case Symtab.lookup(mixtab,c) of + None=>false | Some _ => true + in map (fn (cs,styp)=> (filter_out is_mix cs, styp)) decs + end; + +fun ext_constants None = [] + | ext_constants (Some sext) = Syntax.constants sext; + + +(*Could go to FOL, but it's hardly general*) +val [def] = goal IFOL.thy "a==b ==> a=c <-> c=b"; +by (rewtac def); +by (rtac iffI 1); +by (REPEAT (etac sym 1)); +val def_swap_iff = result(); + +val def_trans = prove_goal IFOL.thy "[| f==g; g(a)=b |] ==> f(a)=b" + (fn [rew,prem] => [ rewtac rew, rtac prem 1 ]); + + diff -r 000000000000 -r a5a9c433f639 src/ZF/indrule.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/indrule.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,162 @@ +(* Title: ZF/indrule.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Induction rule module -- for Inductive/Coinductive Definitions + +Proves a strong induction rule and a mutual induction rule +*) + +signature INDRULE = + sig + val induct : thm (*main induction rule*) + val mutual_induct : thm (*mutual induction rule*) + end; + + +functor Indrule_Fun (structure Ind: INDUCTIVE and + Pr: PR and Intr_elim: INTR_ELIM) : INDRULE = +struct +open Logic Ind Intr_elim; + +val dummy = writeln "Proving the induction rules..."; + +(*** Prove the main induction rule ***) + +val pred_name = "P"; (*name for predicate variables*) + +val prove = prove_term (sign_of Intr_elim.thy); + +val big_rec_def::part_rec_defs = Intr_elim.defs; + +(*Used to make induction rules; + ind_alist = [(rec_tm1,pred1),...] -- associates predicates with rec ops + prem is a premise of an intr rule*) +fun add_induct_prem ind_alist (prem as Const("Trueprop",_) $ + (Const("op :",_)$t$X), iprems) = + (case gen_assoc (op aconv) (ind_alist, X) of + Some pred => prem :: mk_tprop (pred $ t) :: iprems + | None => (*possibly membership in M(rec_tm), for M monotone*) + let fun mk_sb (rec_tm,pred) = (rec_tm, Collect_const$rec_tm$pred) + in subst_free (map mk_sb ind_alist) prem :: iprems end) + | add_induct_prem ind_alist (prem,iprems) = prem :: iprems; + +(*Make a premise of the induction rule.*) +fun induct_prem ind_alist intr = + let val quantfrees = map dest_Free (term_frees intr \\ rec_params) + val iprems = foldr (add_induct_prem ind_alist) + (strip_imp_prems intr,[]) + val (t,X) = rule_concl intr + val (Some pred) = gen_assoc (op aconv) (ind_alist, X) + val concl = mk_tprop (pred $ t) + in list_all_free (quantfrees, list_implies (iprems,concl)) end + handle Bind => error"Recursion term not found in conclusion"; + +(*Avoids backtracking by delivering the correct premise to each goal*) +fun ind_tac [] 0 = all_tac + | ind_tac(prem::prems) i = REPEAT (ares_tac [Part_eqI,prem] i) THEN + ind_tac prems (i-1); + +val pred = Free(pred_name, iT-->oT); + +val ind_prems = map (induct_prem (map (rpair pred) rec_tms)) intr_tms; + +val quant_induct = + prove part_rec_defs + (list_implies (ind_prems, mk_tprop (mk_all_imp(big_rec_tm,pred))), + fn prems => + [rtac (impI RS allI) 1, + etac raw_induct 1, + REPEAT (FIRSTGOAL (eresolve_tac [CollectE,exE,conjE,disjE,ssubst])), + REPEAT (FIRSTGOAL (eresolve_tac [PartE,CollectE])), + ind_tac (rev prems) (length prems) ]); + +(*** Prove the simultaneous induction rule ***) + +(*Make distinct predicates for each inductive set*) + +(*Sigmas and Cartesian products may nest ONLY to the right!*) +fun mk_pred_typ (t $ A $ B) = + if t = Pr.sigma then iT --> mk_pred_typ B + else iT --> oT + | mk_pred_typ _ = iT --> oT + +(*Given a recursive set and its domain, return the "fsplit" predicate + and a conclusion for the simultaneous induction rule*) +fun mk_predpair (rec_tm,domt) = + let val rec_name = (#1 o dest_Const o head_of) rec_tm + val T = mk_pred_typ domt + val pfree = Free(pred_name ^ "_" ^ rec_name, T) + val frees = mk_frees "za" (binder_types T) + val qconcl = + foldr mk_all (frees, + imp $ (mem_const $ foldr1 (app Pr.pair) frees $ rec_tm) + $ (list_comb (pfree,frees))) + in (ap_split Pr.fsplit_const pfree (binder_types T), + qconcl) + end; + +val (preds,qconcls) = split_list (map mk_predpair (rec_tms~~domts)); + +(*Used to form simultaneous induction lemma*) +fun mk_rec_imp (rec_tm,pred) = + imp $ (mem_const $ Bound 0 $ rec_tm) $ (pred $ Bound 0); + +(*To instantiate the main induction rule*) +val induct_concl = + mk_tprop(mk_all_imp(big_rec_tm, + Abs("z", iT, + fold_bal (app conj) + (map mk_rec_imp (rec_tms~~preds))))) +and mutual_induct_concl = mk_tprop(fold_bal (app conj) qconcls); + +val lemma = (*makes the link between the two induction rules*) + prove part_rec_defs + (mk_implies (induct_concl,mutual_induct_concl), + fn prems => + [cut_facts_tac prems 1, + REPEAT (eresolve_tac [asm_rl,conjE,PartE,mp] 1 + ORELSE resolve_tac [allI,impI,conjI,Part_eqI] 1 + ORELSE dresolve_tac [spec, mp, Pr.fsplitD] 1)]); + +(*Mutual induction follows by freeness of Inl/Inr.*) + +(*Removes Collects caused by M-operators in the intro rules*) +val cmonos = [subset_refl RS Collect_mono] RL monos RLN (2,[rev_subsetD]); + +(*Avoids backtracking by delivering the correct premise to each goal*) +fun mutual_ind_tac [] 0 = all_tac + | mutual_ind_tac(prem::prems) i = + SELECT_GOAL + ((*unpackage and use "prem" in the corresponding place*) + REPEAT (FIRSTGOAL + (eresolve_tac ([conjE,mp]@cmonos) ORELSE' + ares_tac [prem,impI,conjI])) + (*prove remaining goals by contradiction*) + THEN rewrite_goals_tac (con_defs@part_rec_defs) + THEN REPEAT (eresolve_tac (PartE :: sumprod_free_SEs) 1)) + i THEN mutual_ind_tac prems (i-1); + +val mutual_induct_fsplit = + prove [] + (list_implies (map (induct_prem (rec_tms~~preds)) intr_tms, + mutual_induct_concl), + fn prems => + [rtac (quant_induct RS lemma) 1, + mutual_ind_tac (rev prems) (length prems)]); + +(*Attempts to remove all occurrences of fsplit*) +val fsplit_tac = + REPEAT (SOMEGOAL (FIRST' [rtac Pr.fsplitI, + dtac Pr.fsplitD, + etac Pr.fsplitE, + bound_hyp_subst_tac])) + THEN prune_params_tac; + +(*strip quantifier*) +val induct = standard (quant_induct RS spec RSN (2,rev_mp)); + +val mutual_induct = rule_by_tactic fsplit_tac mutual_induct_fsplit; + +end; diff -r 000000000000 -r a5a9c433f639 src/ZF/inductive.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/inductive.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,63 @@ +(* Title: ZF/inductive.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Inductive Definitions for Zermelo-Fraenkel Set Theory + +Uses least fixedpoints with standard products and sums + +Sums are used only for mutual recursion; +Products are used only to derive "streamlined" induction rules for relations +*) + + +structure Lfp = + struct + val oper = Const("lfp", [iT,iT-->iT]--->iT) + val bnd_mono = Const("bnd_mono", [iT,iT-->iT]--->oT) + val bnd_monoI = bnd_monoI + val subs = def_lfp_subset + val Tarski = def_lfp_Tarski + val induct = def_induct + end; + +structure Standard_Prod = + struct + val sigma = Const("Sigma", [iT, iT-->iT]--->iT) + val pair = Const("Pair", [iT,iT]--->iT) + val split_const = Const("split", [[iT,iT]--->iT, iT]--->iT) + val fsplit_const = Const("fsplit", [[iT,iT]--->oT, iT]--->oT) + val pair_iff = Pair_iff + val split_eq = split + val fsplitI = fsplitI + val fsplitD = fsplitD + val fsplitE = fsplitE + end; + +structure Standard_Sum = + struct + val sum = Const("op +", [iT,iT]--->iT) + val inl = Const("Inl", iT-->iT) + val inr = Const("Inr", iT-->iT) + val elim = Const("case", [iT-->iT, iT-->iT, iT]--->iT) + val case_inl = case_Inl + val case_inr = case_Inr + val inl_iff = Inl_iff + val inr_iff = Inr_iff + val distinct = Inl_Inr_iff + val distinct' = Inr_Inl_iff + end; + +functor Inductive_Fun (Ind: INDUCTIVE) : sig include INTR_ELIM INDRULE end = +struct +structure Intr_elim = + Intr_elim_Fun(structure Ind=Ind and Fp=Lfp and + Pr=Standard_Prod and Su=Standard_Sum); + +structure Indrule = Indrule_Fun (structure Ind=Ind and + Pr=Standard_Prod and Intr_elim=Intr_elim); + +open Intr_elim Indrule +end; + diff -r 000000000000 -r a5a9c433f639 src/ZF/intr-elim.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/intr-elim.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,279 @@ +(* Title: ZF/intr-elim.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Introduction/elimination rule module -- for Inductive/Coinductive Definitions + +Features: +* least or greatest fixedpoints +* user-specified product and sum constructions +* mutually recursive definitions +* definitions involving arbitrary monotone operators +* automatically proves introduction and elimination rules + +The recursive sets must *already* be declared as constants in parent theory! + + Introduction rules have the form + [| ti:M(Sj), ..., P(x), ... |] ==> t: Sk |] + where M is some monotone operator (usually the identity) + P(x) is any (non-conjunctive) side condition on the free variables + ti, t are any terms + Sj, Sk are two of the sets being defiend in mutual recursion + +Sums are used only for mutual recursion; +Products are used only to derive "streamlined" induction rules for relations +*) + +signature FP = (** Description of a fixed point operator **) + sig + val oper : term (*fixed point operator*) + val bnd_mono : term (*monotonicity predicate*) + val bnd_monoI : thm (*intro rule for bnd_mono*) + val subs : thm (*subset theorem for fp*) + val Tarski : thm (*Tarski's fixed point theorem*) + val induct : thm (*induction/coinduction rule*) + end; + +signature PR = (** Description of a Cartesian product **) + sig + val sigma : term (*Cartesian product operator*) + val pair : term (*pairing operator*) + val split_const : term (*splitting operator*) + val fsplit_const : term (*splitting operator for formulae*) + val pair_iff : thm (*injectivity of pairing, using <->*) + val split_eq : thm (*equality rule for split*) + val fsplitI : thm (*intro rule for fsplit*) + val fsplitD : thm (*destruct rule for fsplit*) + val fsplitE : thm (*elim rule for fsplit*) + end; + +signature SU = (** Description of a disjoint sum **) + sig + val sum : term (*disjoint sum operator*) + val inl : term (*left injection*) + val inr : term (*right injection*) + val elim : term (*case operator*) + val case_inl : thm (*inl equality rule for case*) + val case_inr : thm (*inr equality rule for case*) + val inl_iff : thm (*injectivity of inl, using <->*) + val inr_iff : thm (*injectivity of inr, using <->*) + val distinct : thm (*distinctness of inl, inr using <->*) + val distinct' : thm (*distinctness of inr, inl using <->*) + end; + +signature INDUCTIVE = (** Description of a (co)inductive defn **) + sig + val thy : theory (*parent theory*) + val rec_doms : (string*string) list (*recursion ops and their domains*) + val sintrs : string list (*desired introduction rules*) + val monos : thm list (*monotonicity of each M operator*) + val con_defs : thm list (*definitions of the constructors*) + val type_intrs : thm list (*type-checking intro rules*) + val type_elims : thm list (*type-checking elim rules*) + end; + +signature INTR_ELIM = + sig + val thy : theory (*new theory*) + val defs : thm list (*definitions made in thy*) + val bnd_mono : thm (*monotonicity for the lfp definition*) + val unfold : thm (*fixed-point equation*) + val dom_subset : thm (*inclusion of recursive set in dom*) + val intrs : thm list (*introduction rules*) + val elim : thm (*case analysis theorem*) + val raw_induct : thm (*raw induction rule from Fp.induct*) + val mk_cases : thm list -> string -> thm (*generates case theorems*) + (*internal items...*) + val big_rec_tm : term (*the lhs of the fixedpoint defn*) + val rec_tms : term list (*the mutually recursive sets*) + val domts : term list (*domains of the recursive sets*) + val intr_tms : term list (*terms for the introduction rules*) + val rec_params : term list (*parameters of the recursion*) + val sumprod_free_SEs : thm list (*destruct rules for Su and Pr*) + end; + + +functor Intr_elim_Fun (structure Ind: INDUCTIVE and + Fp: FP and Pr : PR and Su : SU) : INTR_ELIM = +struct +open Logic Ind; + +(*** Extract basic information from arguments ***) + +val sign = sign_of Ind.thy; + +fun rd T a = + Sign.read_cterm sign (a,T) + handle ERROR => error ("The error above occurred for " ^ a); + +val rec_names = map #1 rec_doms +and domts = map (Sign.term_of o rd iT o #2) rec_doms; + +val dummy = assert_all Syntax.is_identifier rec_names + (fn a => "Name of recursive set not an identifier: " ^ a); + +val dummy = assert_all (is_some o lookup_const sign) rec_names + (fn a => "Name of recursive set not declared as constant: " ^ a); + +val intr_tms = map (Sign.term_of o rd propT) sintrs; +val (Const(_,recT),rec_params) = strip_comb (#2 (rule_concl(hd intr_tms))) +val rec_hds = map (fn a=> Const(a,recT)) rec_names; +val rec_tms = map (fn rec_hd=> list_comb(rec_hd,rec_params)) rec_hds; + +val dummy = assert_all is_Free rec_params + (fn t => "Param in recursion term not a free variable: " ^ + Sign.string_of_term sign t); + +(*** Construct the lfp definition ***) + +val mk_variant = variant (foldr add_term_names (intr_tms,[])); + +val z' = mk_variant"z" +and X' = mk_variant"X" +and w' = mk_variant"w"; + +(*simple error-checking in the premises*) +fun chk_prem rec_hd (Const("op &",_) $ _ $ _) = + error"Premises may not be conjuctive" + | chk_prem rec_hd (Const("op :",_) $ t $ X) = + deny (rec_hd occs t) "Recursion term on left of member symbol" + | chk_prem rec_hd t = + deny (rec_hd occs t) "Recursion term in side formula"; + +(*Makes a disjunct from an introduction rule*) +fun lfp_part intr = (*quantify over rule's free vars except parameters*) + let val prems = map dest_tprop (strip_imp_prems intr) + val dummy = seq (fn rec_hd => seq (chk_prem rec_hd) prems) rec_hds + val exfrees = term_frees intr \\ rec_params + val zeq = eq_const $ (Free(z',iT)) $ (#1 (rule_concl intr)) + in foldr mk_exists (exfrees, fold_bal (app conj) (zeq::prems)) end; + +val dom_sum = fold_bal (app Su.sum) domts; + +(*The Part(A,h) terms -- compose injections to make h*) +fun mk_Part (Bound 0) = Free(X',iT) (*no mutual rec, no Part needed*) + | mk_Part h = Part_const $ Free(X',iT) $ Abs(w',iT,h); + +(*Access to balanced disjoint sums via injections*) +val parts = + map mk_Part (accesses_bal (ap Su.inl, ap Su.inr, Bound 0) + (length rec_doms)); + +(*replace each set by the corresponding Part(A,h)*) +val part_intrs = map (subst_free (rec_tms ~~ parts) o lfp_part) intr_tms; + +val lfp_abs = absfree(X', iT, + mk_Collect(z', dom_sum, fold_bal (app disj) part_intrs)); + +val lfp_rhs = Fp.oper $ dom_sum $ lfp_abs + +val dummy = seq (fn rec_hd => deny (rec_hd occs lfp_rhs) + "Illegal occurrence of recursion operator") + rec_hds; + +(*** Make the new theory ***) + +(*A key definition: + If no mutual recursion then it equals the one recursive set. + If mutual recursion then it differs from all the recursive sets. *) +val big_rec_name = space_implode "_" rec_names; + +(*Big_rec... is the union of the mutually recursive sets*) +val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); + +(*The individual sets must already be declared*) +val axpairs = map (mk_defpair sign) + ((big_rec_tm, lfp_rhs) :: + (case parts of + [_] => [] (*no mutual recursion*) + | _ => rec_tms ~~ (*define the sets as Parts*) + map (subst_atomic [(Free(X',iT),big_rec_tm)]) parts)); + +val thy = extend_theory Ind.thy (big_rec_name ^ "_Inductive") + ([], [], [], [], [], None) axpairs; + +val defs = map (get_axiom thy o #1) axpairs; + +val big_rec_def::part_rec_defs = defs; + +val prove = prove_term (sign_of thy); + +(********) +val dummy = writeln "Proving monotonocity..."; + +val bnd_mono = + prove [] (mk_tprop (Fp.bnd_mono $ dom_sum $ lfp_abs), + fn _ => + [rtac (Collect_subset RS bnd_monoI) 1, + REPEAT (ares_tac (basic_monos @ monos) 1)]); + +val dom_subset = standard (big_rec_def RS Fp.subs); + +val unfold = standard (bnd_mono RS (big_rec_def RS Fp.Tarski)); + +(********) +val dummy = writeln "Proving the introduction rules..."; + +(*Mutual recursion: Needs subset rules for the individual sets???*) +val rec_typechecks = [dom_subset] RL (asm_rl::monos) RL [subsetD]; + +(*Type-checking is hardest aspect of proof; + disjIn selects the correct disjunct after unfolding*) +fun intro_tacsf disjIn prems = + [(*insert prems and underlying sets*) + cut_facts_tac (prems @ (prems RL [PartD1])) 1, + rtac (unfold RS ssubst) 1, + REPEAT (resolve_tac [Part_eqI,CollectI] 1), + (*Now 2-3 subgoals: typechecking, the disjunction, perhaps equality.*) + rtac disjIn 2, + REPEAT (ares_tac [refl,exI,conjI] 2), + rewrite_goals_tac con_defs, + (*Now can solve the trivial equation*) + REPEAT (rtac refl 2), + REPEAT (FIRSTGOAL (eresolve_tac (asm_rl::type_elims) + ORELSE' dresolve_tac rec_typechecks)), + DEPTH_SOLVE (ares_tac type_intrs 1)]; + +(*combines disjI1 and disjI2 to access the corresponding nested disjunct...*) +val mk_disj_rls = + let fun f rl = rl RS disjI1 + and g rl = rl RS disjI2 + in accesses_bal(f, g, asm_rl) end; + +val intrs = map (prove part_rec_defs) + (intr_tms ~~ map intro_tacsf (mk_disj_rls(length intr_tms))); + +(********) +val dummy = writeln "Proving the elimination rule..."; + +val elim_rls = [asm_rl, FalseE, conjE, exE, disjE]; + +val sumprod_free_SEs = + map (gen_make_elim [conjE,FalseE]) + ([Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff, Pr.pair_iff] + RL [iffD1]); + +(*Breaks down logical connectives in the monotonic function*) +val basic_elim_tac = + REPEAT (SOMEGOAL (eresolve_tac (elim_rls@sumprod_free_SEs) + ORELSE' bound_hyp_subst_tac)) + THEN prune_params_tac; + +val elim = rule_by_tactic basic_elim_tac (unfold RS equals_CollectD); + +(*Applies freeness of the given constructors. + NB for datatypes, defs=con_defs; for inference systems, con_defs=[]! *) +fun con_elim_tac defs = + rewrite_goals_tac defs THEN basic_elim_tac THEN fold_con_tac defs; + +(*String s should have the form t:Si where Si is an inductive set*) +fun mk_cases defs s = + rule_by_tactic (con_elim_tac defs) + (assume_read thy s RS elim); + +val defs = big_rec_def::part_rec_defs; + +val raw_induct = standard ([big_rec_def, bnd_mono] MRS Fp.induct); + +end; diff -r 000000000000 -r a5a9c433f639 src/ZF/intr_elim.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/intr_elim.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,279 @@ +(* Title: ZF/intr-elim.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Introduction/elimination rule module -- for Inductive/Coinductive Definitions + +Features: +* least or greatest fixedpoints +* user-specified product and sum constructions +* mutually recursive definitions +* definitions involving arbitrary monotone operators +* automatically proves introduction and elimination rules + +The recursive sets must *already* be declared as constants in parent theory! + + Introduction rules have the form + [| ti:M(Sj), ..., P(x), ... |] ==> t: Sk |] + where M is some monotone operator (usually the identity) + P(x) is any (non-conjunctive) side condition on the free variables + ti, t are any terms + Sj, Sk are two of the sets being defiend in mutual recursion + +Sums are used only for mutual recursion; +Products are used only to derive "streamlined" induction rules for relations +*) + +signature FP = (** Description of a fixed point operator **) + sig + val oper : term (*fixed point operator*) + val bnd_mono : term (*monotonicity predicate*) + val bnd_monoI : thm (*intro rule for bnd_mono*) + val subs : thm (*subset theorem for fp*) + val Tarski : thm (*Tarski's fixed point theorem*) + val induct : thm (*induction/coinduction rule*) + end; + +signature PR = (** Description of a Cartesian product **) + sig + val sigma : term (*Cartesian product operator*) + val pair : term (*pairing operator*) + val split_const : term (*splitting operator*) + val fsplit_const : term (*splitting operator for formulae*) + val pair_iff : thm (*injectivity of pairing, using <->*) + val split_eq : thm (*equality rule for split*) + val fsplitI : thm (*intro rule for fsplit*) + val fsplitD : thm (*destruct rule for fsplit*) + val fsplitE : thm (*elim rule for fsplit*) + end; + +signature SU = (** Description of a disjoint sum **) + sig + val sum : term (*disjoint sum operator*) + val inl : term (*left injection*) + val inr : term (*right injection*) + val elim : term (*case operator*) + val case_inl : thm (*inl equality rule for case*) + val case_inr : thm (*inr equality rule for case*) + val inl_iff : thm (*injectivity of inl, using <->*) + val inr_iff : thm (*injectivity of inr, using <->*) + val distinct : thm (*distinctness of inl, inr using <->*) + val distinct' : thm (*distinctness of inr, inl using <->*) + end; + +signature INDUCTIVE = (** Description of a (co)inductive defn **) + sig + val thy : theory (*parent theory*) + val rec_doms : (string*string) list (*recursion ops and their domains*) + val sintrs : string list (*desired introduction rules*) + val monos : thm list (*monotonicity of each M operator*) + val con_defs : thm list (*definitions of the constructors*) + val type_intrs : thm list (*type-checking intro rules*) + val type_elims : thm list (*type-checking elim rules*) + end; + +signature INTR_ELIM = + sig + val thy : theory (*new theory*) + val defs : thm list (*definitions made in thy*) + val bnd_mono : thm (*monotonicity for the lfp definition*) + val unfold : thm (*fixed-point equation*) + val dom_subset : thm (*inclusion of recursive set in dom*) + val intrs : thm list (*introduction rules*) + val elim : thm (*case analysis theorem*) + val raw_induct : thm (*raw induction rule from Fp.induct*) + val mk_cases : thm list -> string -> thm (*generates case theorems*) + (*internal items...*) + val big_rec_tm : term (*the lhs of the fixedpoint defn*) + val rec_tms : term list (*the mutually recursive sets*) + val domts : term list (*domains of the recursive sets*) + val intr_tms : term list (*terms for the introduction rules*) + val rec_params : term list (*parameters of the recursion*) + val sumprod_free_SEs : thm list (*destruct rules for Su and Pr*) + end; + + +functor Intr_elim_Fun (structure Ind: INDUCTIVE and + Fp: FP and Pr : PR and Su : SU) : INTR_ELIM = +struct +open Logic Ind; + +(*** Extract basic information from arguments ***) + +val sign = sign_of Ind.thy; + +fun rd T a = + Sign.read_cterm sign (a,T) + handle ERROR => error ("The error above occurred for " ^ a); + +val rec_names = map #1 rec_doms +and domts = map (Sign.term_of o rd iT o #2) rec_doms; + +val dummy = assert_all Syntax.is_identifier rec_names + (fn a => "Name of recursive set not an identifier: " ^ a); + +val dummy = assert_all (is_some o lookup_const sign) rec_names + (fn a => "Name of recursive set not declared as constant: " ^ a); + +val intr_tms = map (Sign.term_of o rd propT) sintrs; +val (Const(_,recT),rec_params) = strip_comb (#2 (rule_concl(hd intr_tms))) +val rec_hds = map (fn a=> Const(a,recT)) rec_names; +val rec_tms = map (fn rec_hd=> list_comb(rec_hd,rec_params)) rec_hds; + +val dummy = assert_all is_Free rec_params + (fn t => "Param in recursion term not a free variable: " ^ + Sign.string_of_term sign t); + +(*** Construct the lfp definition ***) + +val mk_variant = variant (foldr add_term_names (intr_tms,[])); + +val z' = mk_variant"z" +and X' = mk_variant"X" +and w' = mk_variant"w"; + +(*simple error-checking in the premises*) +fun chk_prem rec_hd (Const("op &",_) $ _ $ _) = + error"Premises may not be conjuctive" + | chk_prem rec_hd (Const("op :",_) $ t $ X) = + deny (rec_hd occs t) "Recursion term on left of member symbol" + | chk_prem rec_hd t = + deny (rec_hd occs t) "Recursion term in side formula"; + +(*Makes a disjunct from an introduction rule*) +fun lfp_part intr = (*quantify over rule's free vars except parameters*) + let val prems = map dest_tprop (strip_imp_prems intr) + val dummy = seq (fn rec_hd => seq (chk_prem rec_hd) prems) rec_hds + val exfrees = term_frees intr \\ rec_params + val zeq = eq_const $ (Free(z',iT)) $ (#1 (rule_concl intr)) + in foldr mk_exists (exfrees, fold_bal (app conj) (zeq::prems)) end; + +val dom_sum = fold_bal (app Su.sum) domts; + +(*The Part(A,h) terms -- compose injections to make h*) +fun mk_Part (Bound 0) = Free(X',iT) (*no mutual rec, no Part needed*) + | mk_Part h = Part_const $ Free(X',iT) $ Abs(w',iT,h); + +(*Access to balanced disjoint sums via injections*) +val parts = + map mk_Part (accesses_bal (ap Su.inl, ap Su.inr, Bound 0) + (length rec_doms)); + +(*replace each set by the corresponding Part(A,h)*) +val part_intrs = map (subst_free (rec_tms ~~ parts) o lfp_part) intr_tms; + +val lfp_abs = absfree(X', iT, + mk_Collect(z', dom_sum, fold_bal (app disj) part_intrs)); + +val lfp_rhs = Fp.oper $ dom_sum $ lfp_abs + +val dummy = seq (fn rec_hd => deny (rec_hd occs lfp_rhs) + "Illegal occurrence of recursion operator") + rec_hds; + +(*** Make the new theory ***) + +(*A key definition: + If no mutual recursion then it equals the one recursive set. + If mutual recursion then it differs from all the recursive sets. *) +val big_rec_name = space_implode "_" rec_names; + +(*Big_rec... is the union of the mutually recursive sets*) +val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); + +(*The individual sets must already be declared*) +val axpairs = map (mk_defpair sign) + ((big_rec_tm, lfp_rhs) :: + (case parts of + [_] => [] (*no mutual recursion*) + | _ => rec_tms ~~ (*define the sets as Parts*) + map (subst_atomic [(Free(X',iT),big_rec_tm)]) parts)); + +val thy = extend_theory Ind.thy (big_rec_name ^ "_Inductive") + ([], [], [], [], [], None) axpairs; + +val defs = map (get_axiom thy o #1) axpairs; + +val big_rec_def::part_rec_defs = defs; + +val prove = prove_term (sign_of thy); + +(********) +val dummy = writeln "Proving monotonocity..."; + +val bnd_mono = + prove [] (mk_tprop (Fp.bnd_mono $ dom_sum $ lfp_abs), + fn _ => + [rtac (Collect_subset RS bnd_monoI) 1, + REPEAT (ares_tac (basic_monos @ monos) 1)]); + +val dom_subset = standard (big_rec_def RS Fp.subs); + +val unfold = standard (bnd_mono RS (big_rec_def RS Fp.Tarski)); + +(********) +val dummy = writeln "Proving the introduction rules..."; + +(*Mutual recursion: Needs subset rules for the individual sets???*) +val rec_typechecks = [dom_subset] RL (asm_rl::monos) RL [subsetD]; + +(*Type-checking is hardest aspect of proof; + disjIn selects the correct disjunct after unfolding*) +fun intro_tacsf disjIn prems = + [(*insert prems and underlying sets*) + cut_facts_tac (prems @ (prems RL [PartD1])) 1, + rtac (unfold RS ssubst) 1, + REPEAT (resolve_tac [Part_eqI,CollectI] 1), + (*Now 2-3 subgoals: typechecking, the disjunction, perhaps equality.*) + rtac disjIn 2, + REPEAT (ares_tac [refl,exI,conjI] 2), + rewrite_goals_tac con_defs, + (*Now can solve the trivial equation*) + REPEAT (rtac refl 2), + REPEAT (FIRSTGOAL (eresolve_tac (asm_rl::type_elims) + ORELSE' dresolve_tac rec_typechecks)), + DEPTH_SOLVE (ares_tac type_intrs 1)]; + +(*combines disjI1 and disjI2 to access the corresponding nested disjunct...*) +val mk_disj_rls = + let fun f rl = rl RS disjI1 + and g rl = rl RS disjI2 + in accesses_bal(f, g, asm_rl) end; + +val intrs = map (prove part_rec_defs) + (intr_tms ~~ map intro_tacsf (mk_disj_rls(length intr_tms))); + +(********) +val dummy = writeln "Proving the elimination rule..."; + +val elim_rls = [asm_rl, FalseE, conjE, exE, disjE]; + +val sumprod_free_SEs = + map (gen_make_elim [conjE,FalseE]) + ([Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff, Pr.pair_iff] + RL [iffD1]); + +(*Breaks down logical connectives in the monotonic function*) +val basic_elim_tac = + REPEAT (SOMEGOAL (eresolve_tac (elim_rls@sumprod_free_SEs) + ORELSE' bound_hyp_subst_tac)) + THEN prune_params_tac; + +val elim = rule_by_tactic basic_elim_tac (unfold RS equals_CollectD); + +(*Applies freeness of the given constructors. + NB for datatypes, defs=con_defs; for inference systems, con_defs=[]! *) +fun con_elim_tac defs = + rewrite_goals_tac defs THEN basic_elim_tac THEN fold_con_tac defs; + +(*String s should have the form t:Si where Si is an inductive set*) +fun mk_cases defs s = + rule_by_tactic (con_elim_tac defs) + (assume_read thy s RS elim); + +val defs = big_rec_def::part_rec_defs; + +val raw_induct = standard ([big_rec_def, bnd_mono] MRS Fp.induct); + +end; diff -r 000000000000 -r a5a9c433f639 src/ZF/list.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/list.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,80 @@ +(* Title: ZF/list.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Datatype definition of Lists +*) + +structure List = Datatype_Fun + (val thy = Univ.thy; + val rec_specs = + [("list", "univ(A)", + [(["Nil"], "i"), + (["Cons"], "[i,i]=>i")])]; + val rec_styp = "i=>i"; + val ext = None + val sintrs = + ["Nil : list(A)", + "[| a: A; l: list(A) |] ==> Cons(a,l) : list(A)"]; + val monos = []; + val type_intrs = data_typechecks + val type_elims = []); + +val [NilI, ConsI] = List.intrs; + +(*An elimination rule, for type-checking*) +val ConsE = List.mk_cases List.con_defs "Cons(a,l) : list(A)"; + +(*Proving freeness results*) +val Cons_iff = List.mk_free "Cons(a,l)=Cons(a',l') <-> a=a' & l=l'"; +val Nil_Cons_iff = List.mk_free "~ Nil=Cons(a,l)"; + +(*Perform induction on l, then prove the major premise using prems. *) +fun list_ind_tac a prems i = + EVERY [res_inst_tac [("x",a)] List.induct i, + rename_last_tac a ["1"] (i+2), + ares_tac prems i]; + +(** Lemmas to justify using "list" in other recursive type definitions **) + +goalw List.thy List.defs "!!A B. A<=B ==> list(A) <= list(B)"; +by (rtac lfp_mono 1); +by (REPEAT (rtac List.bnd_mono 1)); +by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); +val list_mono = result(); + +(*There is a similar proof by list induction.*) +goalw List.thy (List.defs@List.con_defs) "list(univ(A)) <= univ(A)"; +by (rtac lfp_lowerbound 1); +by (rtac (A_subset_univ RS univ_mono) 2); +by (fast_tac (ZF_cs addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ, + Pair_in_univ]) 1); +val list_univ = result(); + +val list_subset_univ = standard + (list_mono RS (list_univ RSN (2,subset_trans))); + +(***** +val major::prems = goal List.thy + "[| l: list(A); \ +\ c: C(0); \ +\ !!x y. [| x: A; y: list(A) |] ==> h(x,y): C() \ +\ |] ==> list_case(l,c,h) : C(l)"; +by (rtac (major RS list_induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (ZF_ss addrews + ([list_case_0,list_case_Pair]@prems)))); +val list_case_type = result(); +****) + + +(** For recursion **) + +goalw List.thy List.con_defs "rank(a) : rank(Cons(a,l))"; +by (SIMP_TAC rank_ss 1); +val rank_Cons1 = result(); + +goalw List.thy List.con_defs "rank(l) : rank(Cons(a,l))"; +by (SIMP_TAC rank_ss 1); +val rank_Cons2 = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/listfn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/listfn.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,306 @@ +(* Title: ZF/list-fn.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For list-fn.thy. Lists in Zermelo-Fraenkel Set Theory +*) + +open ListFn; + + +(** list_rec -- by Vset recursion **) + +(*Used to verify list_rec*) +val list_rec_ss = ZF_ss + addcongs (mk_typed_congs ListFn.thy [("h", "[i,i,i]=>i")]) + addrews List.case_eqns; + +goal ListFn.thy "list_rec(Nil,c,h) = c"; +by (rtac (list_rec_def RS def_Vrec RS trans) 1); +by (SIMP_TAC list_rec_ss 1); +val list_rec_Nil = result(); + +goal ListFn.thy "list_rec(Cons(a,l), c, h) = h(a, l, list_rec(l,c,h))"; +by (rtac (list_rec_def RS def_Vrec RS trans) 1); +by (SIMP_TAC (list_rec_ss addrews [Vset_rankI, rank_Cons2]) 1); +val list_rec_Cons = result(); + +(*Type checking -- proved by induction, as usual*) +val prems = goal ListFn.thy + "[| l: list(A); \ +\ c: C(Nil); \ +\ !!x y r. [| x:A; y: list(A); r: C(y) |] ==> h(x,y,r): C(Cons(x,y)) \ +\ |] ==> list_rec(l,c,h) : C(l)"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS (ASM_SIMP_TAC + (ZF_ss addrews (prems@[list_rec_Nil,list_rec_Cons])))); +val list_rec_type = result(); + +(** Versions for use with definitions **) + +val [rew] = goal ListFn.thy + "[| !!l. j(l)==list_rec(l, c, h) |] ==> j(Nil) = c"; +by (rewtac rew); +by (rtac list_rec_Nil 1); +val def_list_rec_Nil = result(); + +val [rew] = goal ListFn.thy + "[| !!l. j(l)==list_rec(l, c, h) |] ==> j(Cons(a,l)) = h(a,l,j(l))"; +by (rewtac rew); +by (rtac list_rec_Cons 1); +val def_list_rec_Cons = result(); + +fun list_recs def = map standard + ([def] RL [def_list_rec_Nil, def_list_rec_Cons]); + +(** map **) + +val [map_Nil,map_Cons] = list_recs map_def; + +val prems = goalw ListFn.thy [map_def] + "[| l: list(A); !!x. x: A ==> h(x): B |] ==> map(h,l) : list(B)"; +by (REPEAT (ares_tac (prems@[list_rec_type, NilI, ConsI]) 1)); +val map_type = result(); + +val [major] = goal ListFn.thy "l: list(A) ==> map(h,l) : list({h(u). u:A})"; +by (rtac (major RS map_type) 1); +by (etac RepFunI 1); +val map_type2 = result(); + +(** length **) + +val [length_Nil,length_Cons] = list_recs length_def; + +val prems = goalw ListFn.thy [length_def] + "l: list(A) ==> length(l) : nat"; +by (REPEAT (ares_tac (prems @ [list_rec_type, nat_0I, nat_succI]) 1)); +val length_type = result(); + +(** app **) + +val [app_Nil,app_Cons] = list_recs app_def; + +val prems = goalw ListFn.thy [app_def] + "[| xs: list(A); ys: list(A) |] ==> xs@ys : list(A)"; +by (REPEAT (ares_tac (prems @ [list_rec_type, ConsI]) 1)); +val app_type = result(); + +(** rev **) + +val [rev_Nil,rev_Cons] = list_recs rev_def; + +val prems = goalw ListFn.thy [rev_def] + "xs: list(A) ==> rev(xs) : list(A)"; +by (REPEAT (ares_tac (prems @ [list_rec_type, NilI, ConsI, app_type]) 1)); +val rev_type = result(); + + +(** flat **) + +val [flat_Nil,flat_Cons] = list_recs flat_def; + +val prems = goalw ListFn.thy [flat_def] + "ls: list(list(A)) ==> flat(ls) : list(A)"; +by (REPEAT (ares_tac (prems @ [list_rec_type, NilI, ConsI, app_type]) 1)); +val flat_type = result(); + + +(** list_add **) + +val [list_add_Nil,list_add_Cons] = list_recs list_add_def; + +val prems = goalw ListFn.thy [list_add_def] + "xs: list(nat) ==> list_add(xs) : nat"; +by (REPEAT (ares_tac (prems @ [list_rec_type, nat_0I, add_type]) 1)); +val list_add_type = result(); + +(** ListFn simplification **) + +val list_typechecks = + [NilI, ConsI, list_rec_type, + map_type, map_type2, app_type, length_type, rev_type, flat_type, + list_add_type]; + +val list_congs = + List.congs @ + mk_congs ListFn.thy + ["list_rec","map","op @","length","rev","flat","list_add"]; + +val list_ss = arith_ss + addcongs list_congs + addrews List.case_eqns + addrews list_typechecks + addrews [list_rec_Nil, list_rec_Cons, + map_Nil, map_Cons, + app_Nil, app_Cons, + length_Nil, length_Cons, + rev_Nil, rev_Cons, + flat_Nil, flat_Cons, + list_add_Nil, list_add_Cons]; + +(*** theorems about map ***) + +val prems = goal ListFn.thy + "l: list(A) ==> map(%u.u, l) = l"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val map_ident = result(); + +val prems = goal ListFn.thy + "l: list(A) ==> map(h, map(j,l)) = map(%u.h(j(u)), l)"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val map_compose = result(); + +val prems = goal ListFn.thy + "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)"; +by (list_ind_tac "xs" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val map_app_distrib = result(); + +val prems = goal ListFn.thy + "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))"; +by (list_ind_tac "ls" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [map_app_distrib]))); +val map_flat = result(); + +val prems = goal ListFn.thy + "l: list(A) ==> \ +\ list_rec(map(h,l), c, d) = \ +\ list_rec(l, c, %x xs r. d(h(x), map(h,xs), r))"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS + (ASM_SIMP_TAC + (list_ss addcongs (mk_typed_congs ListFn.thy [("d", "[i,i,i]=>i")])))); +val list_rec_map = result(); + +(** theorems about list(Collect(A,P)) -- used in ex/term.ML **) + +(* c : list(Collect(B,P)) ==> c : list(B) *) +val list_CollectD = standard (Collect_subset RS list_mono RS subsetD); + +val prems = goal ListFn.thy + "l: list({x:A. h(x)=j(x)}) ==> map(h,l) = map(j,l)"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val map_list_Collect = result(); + +(*** theorems about length ***) + +val prems = goal ListFn.thy + "xs: list(A) ==> length(map(h,xs)) = length(xs)"; +by (list_ind_tac "xs" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val length_map = result(); + +val prems = goal ListFn.thy + "xs: list(A) ==> length(xs@ys) = length(xs) #+ length(ys)"; +by (list_ind_tac "xs" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val length_app = result(); + +(* [| m: nat; n: nat |] ==> m #+ succ(n) = succ(n) #+ m + Used for rewriting below*) +val add_commute_succ = nat_succI RSN (2,add_commute); + +val prems = goal ListFn.thy + "xs: list(A) ==> length(rev(xs)) = length(xs)"; +by (list_ind_tac "xs" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [length_app, add_commute_succ]))); +val length_rev = result(); + +val prems = goal ListFn.thy + "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))"; +by (list_ind_tac "ls" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [length_app]))); +val length_flat = result(); + +(*** theorems about app ***) + +val [major] = goal ListFn.thy "xs: list(A) ==> xs@Nil=xs"; +by (rtac (major RS List.induct) 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val app_right_Nil = result(); + +val prems = goal ListFn.thy "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)"; +by (list_ind_tac "xs" prems 1); +by (ALLGOALS (ASM_SIMP_TAC list_ss)); +val app_assoc = result(); + +val prems = goal ListFn.thy + "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)"; +by (list_ind_tac "ls" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [app_assoc]))); +val flat_app_distrib = result(); + +(*** theorems about rev ***) + +val prems = goal ListFn.thy "l: list(A) ==> rev(map(h,l)) = map(h,rev(l))"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [map_app_distrib]))); +val rev_map_distrib = result(); + +(*Simplifier needs the premises as assumptions because rewriting will not + instantiate the variable ?A in the rules' typing conditions; note that + rev_type does not instantiate ?A. Only the premises do. +*) +val prems = goal ListFn.thy + "[| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)"; +by (cut_facts_tac prems 1); +by (etac List.induct 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [app_right_Nil,app_assoc]))); +val rev_app_distrib = result(); + +val prems = goal ListFn.thy "l: list(A) ==> rev(rev(l))=l"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [rev_app_distrib]))); +val rev_rev_ident = result(); + +val prems = goal ListFn.thy + "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))"; +by (list_ind_tac "ls" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews + [map_app_distrib, flat_app_distrib, rev_app_distrib, app_right_Nil]))); +val rev_flat = result(); + + +(*** theorems about list_add ***) + +val prems = goal ListFn.thy + "[| xs: list(nat); ys: list(nat) |] ==> \ +\ list_add(xs@ys) = list_add(ys) #+ list_add(xs)"; +by (cut_facts_tac prems 1); +by (list_ind_tac "xs" prems 1); +by (ALLGOALS + (ASM_SIMP_TAC (list_ss addrews [add_0_right, add_assoc RS sym]))); +by (resolve_tac arith_congs 1); +by (REPEAT (ares_tac [refl, list_add_type, add_commute] 1)); +val list_add_app = result(); + +val prems = goal ListFn.thy + "l: list(nat) ==> list_add(rev(l)) = list_add(l)"; +by (list_ind_tac "l" prems 1); +by (ALLGOALS + (ASM_SIMP_TAC (list_ss addrews [list_add_app, add_0_right]))); +val list_add_rev = result(); + +val prems = goal ListFn.thy + "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))"; +by (list_ind_tac "ls" prems 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [list_add_app]))); +by (REPEAT (ares_tac [refl, list_add_type, map_type, add_commute] 1)); +val list_add_flat = result(); + +(** New induction rule **) + +val major::prems = goal ListFn.thy + "[| l: list(A); \ +\ P(Nil); \ +\ !!x y. [| x: A; y: list(A); P(y) |] ==> P(y @ [x]) \ +\ |] ==> P(l)"; +by (rtac (major RS rev_rev_ident RS subst) 1); +by (rtac (major RS rev_type RS List.induct) 1); +by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews prems))); +val list_append_induct = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/listfn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/listfn.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,44 @@ +(* Title: ZF/list-fn + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Functions for Lists in Zermelo-Fraenkel Set Theory + +map is a binding operator -- it applies to meta-level functions, not +object-level functions. This simplifies the final form of term_rec_conv, +although complicating its derivation. +*) + +ListFn = List + +consts + "@" :: "[i,i]=>i" (infixr 60) + list_rec :: "[i, i, [i,i,i]=>i] => i" + map :: "[i=>i, i] => i" + length,rev :: "i=>i" + flat :: "i=>i" + list_add :: "i=>i" + + (* List Enumeration *) + "[]" :: "i" ("[]") + "@List" :: "args => i" ("[(_)]") + + +translations + "[x, xs]" == "Cons(x, [xs])" + "[x]" == "Cons(x, [])" + "[]" == "Nil" + + +rules + list_rec_def + "list_rec(l,c,h) == Vrec(l, %l g.list_case(c, %x xs. h(x, xs, g`xs), l))" + + map_def "map(f,l) == list_rec(l, Nil, %x xs r. Cons(f(x), r))" + length_def "length(l) == list_rec(l, 0, %x xs r. succ(r))" + app_def "xs@ys == list_rec(xs, ys, %x xs r. Cons(x,r))" + rev_def "rev(l) == list_rec(l, Nil, %x xs r. r @ [x])" + flat_def "flat(ls) == list_rec(ls, Nil, %l ls r. l @ r)" + list_add_def "list_add(l) == list_rec(l, 0, %x xs r. x#+r)" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/mono.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/mono.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,192 @@ +(* Title: ZF/mono + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Monotonicity of various operations (for lattice properties see subset.ML) +*) + +(** Replacement, in its various formulations **) + +(*Not easy to express monotonicity in P, since any "bigger" predicate + would have to be single-valued*) +goal ZF.thy "!!A B. A<=B ==> Replace(A,P) <= Replace(B,P)"; +by (fast_tac ZF_cs 1); +val Replace_mono = result(); + +goal ZF.thy "!!A B. A<=B ==> {f(x). x:A} <= {f(x). x:B}"; +by (fast_tac ZF_cs 1); +val RepFun_mono = result(); + +goal ZF.thy "!!A B. A<=B ==> Pow(A) <= Pow(B)"; +by (fast_tac ZF_cs 1); +val Pow_mono = result(); + +goal ZF.thy "!!A B. A<=B ==> Union(A) <= Union(B)"; +by (fast_tac ZF_cs 1); +val Union_mono = result(); + +val prems = goal ZF.thy + "[| A<=C; !!x. x:A ==> B(x)<=D(x) \ +\ |] ==> (UN x:A. B(x)) <= (UN x:C. D(x))"; +by (fast_tac (ZF_cs addIs (prems RL [subsetD])) 1); +val UN_mono = result(); + +(*Intersection is ANTI-monotonic. There are TWO premises! *) +goal ZF.thy "!!A B. [| A<=B; a:A |] ==> Inter(B) <= Inter(A)"; +by (fast_tac ZF_cs 1); +val Inter_anti_mono = result(); + +goal ZF.thy "!!C D. C<=D ==> cons(a,C) <= cons(a,D)"; +by (fast_tac ZF_cs 1); +val cons_mono = result(); + +goal ZF.thy "!!A B C D. [| A<=C; B<=D |] ==> A Un B <= C Un D"; +by (fast_tac ZF_cs 1); +val Un_mono = result(); + +goal ZF.thy "!!A B C D. [| A<=C; B<=D |] ==> A Int B <= C Int D"; +by (fast_tac ZF_cs 1); +val Int_mono = result(); + +goal ZF.thy "!!A B C D. [| A<=C; D<=B |] ==> A-B <= C-D"; +by (fast_tac ZF_cs 1); +val Diff_mono = result(); + +(** Standard products, sums and function spaces **) + +goal ZF.thy "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \ +\ Sigma(A,B) <= Sigma(C,D)"; +by (fast_tac ZF_cs 1); +val Sigma_mono_lemma = result(); +val Sigma_mono = ballI RSN (2,Sigma_mono_lemma); + +goalw Sum.thy sum_defs "!!A B C D. [| A<=C; B<=D |] ==> A+B <= C+D"; +by (REPEAT (ares_tac [subset_refl,Un_mono,Sigma_mono] 1)); +val sum_mono = result(); + +(*Note that B->A and C->A are typically disjoint!*) +goal ZF.thy "!!A B C. B<=C ==> A->B <= A->C"; +by (fast_tac (ZF_cs addIs [lam_type] addEs [Pi_lamE]) 1); +val Pi_mono = result(); + +goalw ZF.thy [lam_def] "!!A B. A<=B ==> Lambda(A,c) <= Lambda(B,c)"; +by (etac RepFun_mono 1); +val lam_mono = result(); + +(** Quine-inspired ordered pairs, products, injections and sums **) + +goalw QPair.thy [QPair_def] "!!a b c d. [| a<=c; b<=d |] ==> <= "; +by (REPEAT (ares_tac [sum_mono] 1)); +val QPair_mono = result(); + +goal QPair.thy "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \ +\ QSigma(A,B) <= QSigma(C,D)"; +by (fast_tac (ZF_cs addIs [QSigmaI] addSEs [QSigmaE]) 1); +val QSigma_mono_lemma = result(); +val QSigma_mono = ballI RSN (2,QSigma_mono_lemma); + +goalw QPair.thy [QInl_def] "!!a b. a<=b ==> QInl(a) <= QInl(b)"; +by (REPEAT (ares_tac [subset_refl RS QPair_mono] 1)); +val QInl_mono = result(); + +goalw QPair.thy [QInr_def] "!!a b. a<=b ==> QInr(a) <= QInr(b)"; +by (REPEAT (ares_tac [subset_refl RS QPair_mono] 1)); +val QInr_mono = result(); + +goal QPair.thy "!!A B C D. [| A<=C; B<=D |] ==> A <+> B <= C <+> D"; +by (fast_tac qsum_cs 1); +val qsum_mono = result(); + + +(** Converse, domain, range, field **) + +goal ZF.thy "!!r s. r<=s ==> converse(r) <= converse(s)"; +by (fast_tac ZF_cs 1); +val converse_mono = result(); + +goal ZF.thy "!!r s. r<=s ==> domain(r)<=domain(s)"; +by (fast_tac ZF_cs 1); +val domain_mono = result(); + +val [prem] = goal ZF.thy "r <= Sigma(A,B) ==> domain(r) <= A"; +by (rtac (domain_subset RS (prem RS domain_mono RS subset_trans)) 1); +val domain_rel_subset = result(); + +goal ZF.thy "!!r s. r<=s ==> range(r)<=range(s)"; +by (fast_tac ZF_cs 1); +val range_mono = result(); + +val [prem] = goal ZF.thy "r <= A*B ==> range(r) <= B"; +by (rtac (range_subset RS (prem RS range_mono RS subset_trans)) 1); +val range_rel_subset = result(); + +goal ZF.thy "!!r s. r<=s ==> field(r)<=field(s)"; +by (fast_tac ZF_cs 1); +val field_mono = result(); + +goal ZF.thy "!!r A. r <= A*A ==> field(r) <= A"; +by (etac (field_mono RS subset_trans) 1); +by (fast_tac ZF_cs 1); +val field_rel_subset = result(); + + +(** Images **) + +val [prem1,prem2] = goal ZF.thy + "[| !! x y. :r ==> :s; A<=B |] ==> r``A <= s``B"; +by (fast_tac (ZF_cs addIs [prem1, prem2 RS subsetD]) 1); +val image_pair_mono = result(); + +val [prem1,prem2] = goal ZF.thy + "[| !! x y. :r ==> :s; A<=B |] ==> r-``A <= s-``B"; +by (fast_tac (ZF_cs addIs [prem1, prem2 RS subsetD]) 1); +val vimage_pair_mono = result(); + +goal ZF.thy "!!r s. [| r<=s; A<=B |] ==> r``A <= s``B"; +by (fast_tac ZF_cs 1); +val image_mono = result(); + +goal ZF.thy "!!r s. [| r<=s; A<=B |] ==> r-``A <= s-``B"; +by (fast_tac ZF_cs 1); +val vimage_mono = result(); + +val [sub,PQimp] = goal ZF.thy + "[| A<=B; !!x. x:A ==> P(x) --> Q(x) |] ==> Collect(A,P) <= Collect(B,Q)"; +by (fast_tac (ZF_cs addIs [sub RS subsetD, PQimp RS mp]) 1); +val Collect_mono = result(); + +(** Monotonicity of implications -- some could go to FOL **) + +goal ZF.thy "!!A B x. A<=B ==> x:A --> x:B"; +by (rtac impI 1); +by (etac subsetD 1); +by (assume_tac 1); +val in_mono = result(); + +goal IFOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1&P2) --> (Q1&Q2)"; +by (Int.fast_tac 1); +val conj_mono = result(); + +goal IFOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1|P2) --> (Q1|Q2)"; +by (Int.fast_tac 1); +val disj_mono = result(); + +goal IFOL.thy "!!P1 P2 Q1 Q2.[| Q1-->P1; P2-->Q2 |] ==> (P1-->P2)-->(Q1-->Q2)"; +by (Int.fast_tac 1); +val imp_mono = result(); + +goal IFOL.thy "P-->P"; +by (rtac impI 1); +by (assume_tac 1); +val imp_refl = result(); + +val [PQimp] = goal IFOL.thy + "[| !!x. P(x) --> Q(x) |] ==> (EX x.P(x)) --> (EX x.Q(x))"; +by (fast_tac (FOL_cs addIs [PQimp RS mp]) 1); +val ex_mono = result(); + +val [PQimp] = goal IFOL.thy + "[| !!x. P(x) --> Q(x) |] ==> (ALL x.P(x)) --> (ALL x.Q(x))"; +by (fast_tac (FOL_cs addIs [PQimp RS mp]) 1); +val all_mono = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/nat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/nat.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,182 @@ +(* Title: ZF/nat.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For nat.thy. Natural numbers in Zermelo-Fraenkel Set Theory +*) + +open Nat; + +goal Nat.thy "bnd_mono(Inf, %X. {0} Un {succ(i). i:X})"; +by (rtac bnd_monoI 1); +by (REPEAT (ares_tac [subset_refl, RepFun_mono, Un_mono] 2)); +by (cut_facts_tac [infinity] 1); +by (fast_tac ZF_cs 1); +val nat_bnd_mono = result(); + +(* nat = {0} Un {succ(x). x:nat} *) +val nat_unfold = nat_bnd_mono RS (nat_def RS def_lfp_Tarski); + +(** Type checking of 0 and successor **) + +goal Nat.thy "0 : nat"; +by (rtac (nat_unfold RS ssubst) 1); +by (rtac (singletonI RS UnI1) 1); +val nat_0I = result(); + +val prems = goal Nat.thy "n : nat ==> succ(n) : nat"; +by (rtac (nat_unfold RS ssubst) 1); +by (rtac (RepFunI RS UnI2) 1); +by (resolve_tac prems 1); +val nat_succI = result(); + +goalw Nat.thy [one_def] "1 : nat"; +by (rtac (nat_0I RS nat_succI) 1); +val nat_1I = result(); + +goal Nat.thy "bool <= nat"; +by (REPEAT (ares_tac [subsetI,nat_0I,nat_1I] 1 ORELSE etac boolE 1)); +val bool_subset_nat = result(); + +val bool_into_nat = bool_subset_nat RS subsetD; + + +(** Injectivity properties and induction **) + +(*Mathematical induction*) +val major::prems = goal Nat.thy + "[| n: nat; P(0); !!x. [| x: nat; P(x) |] ==> P(succ(x)) |] ==> P(n)"; +by (rtac ([nat_def, nat_bnd_mono, major] MRS def_induct) 1); +by (fast_tac (ZF_cs addIs prems) 1); +val nat_induct = result(); + +(*Perform induction on n, then prove the n:nat subgoal using prems. *) +fun nat_ind_tac a prems i = + EVERY [res_inst_tac [("n",a)] nat_induct i, + rename_last_tac a ["1"] (i+2), + ares_tac prems i]; + +val major::prems = goal Nat.thy + "[| n: nat; n=0 ==> P; !!x. [| x: nat; n=succ(x) |] ==> P |] ==> P"; +br (major RS (nat_unfold RS equalityD1 RS subsetD) RS UnE) 1; +by (DEPTH_SOLVE (eresolve_tac [singletonE,RepFunE] 1 + ORELSE ares_tac prems 1)); +val natE = result(); + +val prems = goal Nat.thy "n: nat ==> Ord(n)"; +by (nat_ind_tac "n" prems 1); +by (REPEAT (ares_tac [Ord_0, Ord_succ] 1)); +val naturals_are_ordinals = result(); + +goal Nat.thy "!!n. n: nat ==> n=0 | 0:n"; +by (etac nat_induct 1); +by (fast_tac ZF_cs 1); +by (fast_tac (ZF_cs addIs [naturals_are_ordinals RS Ord_0_mem_succ]) 1); +val natE0 = result(); + +goal Nat.thy "Ord(nat)"; +by (rtac OrdI 1); +by (etac (naturals_are_ordinals RS Ord_is_Transset) 2); +by (rewtac Transset_def); +by (rtac ballI 1); +by (etac nat_induct 1); +by (REPEAT (ares_tac [empty_subsetI,succ_subsetI] 1)); +val Ord_nat = result(); + +(** Variations on mathematical induction **) + +(*complete induction*) +val complete_induct = Ord_nat RSN (2, Ord_induct); + +val prems = goal Nat.thy + "[| m: nat; n: nat; \ +\ !!x. [| x: nat; m<=x; P(x) |] ==> P(succ(x)) \ +\ |] ==> m <= n --> P(m) --> P(n)"; +by (nat_ind_tac "n" prems 1); +by (ALLGOALS + (ASM_SIMP_TAC + (ZF_ss addrews (prems@distrib_rews@[subset_empty_iff, subset_succ_iff, + Ord_nat RS Ord_in_Ord])))); +val nat_induct_from_lemma = result(); + +(*Induction starting from m rather than 0*) +val prems = goal Nat.thy + "[| m <= n; m: nat; n: nat; \ +\ P(m); \ +\ !!x. [| x: nat; m<=x; P(x) |] ==> P(succ(x)) \ +\ |] ==> P(n)"; +by (rtac (nat_induct_from_lemma RS mp RS mp) 1); +by (REPEAT (ares_tac prems 1)); +val nat_induct_from = result(); + +(*Induction suitable for subtraction and less-than*) +val prems = goal Nat.thy + "[| m: nat; n: nat; \ +\ !!x. [| x: nat |] ==> P(x,0); \ +\ !!y. [| y: nat |] ==> P(0,succ(y)); \ +\ !!x y. [| x: nat; y: nat; P(x,y) |] ==> P(succ(x),succ(y)) \ +\ |] ==> P(m,n)"; +by (res_inst_tac [("x","m")] bspec 1); +by (resolve_tac prems 2); +by (nat_ind_tac "n" prems 1); +by (rtac ballI 2); +by (nat_ind_tac "x" [] 2); +by (REPEAT (ares_tac (prems@[ballI]) 1 ORELSE etac bspec 1)); +val diff_induct = result(); + +(** nat_case **) + +goalw Nat.thy [nat_case_def] "nat_case(0,a,b) = a"; +by (fast_tac (ZF_cs addIs [the_equality]) 1); +val nat_case_0 = result(); + +goalw Nat.thy [nat_case_def] "nat_case(succ(m),a,b) = b(m)"; +by (fast_tac (ZF_cs addIs [the_equality]) 1); +val nat_case_succ = result(); + +val major::prems = goal Nat.thy + "[| n: nat; a: C(0); !!m. m: nat ==> b(m): C(succ(m)) \ +\ |] ==> nat_case(n,a,b) : C(n)"; +by (rtac (major RS nat_induct) 1); +by (REPEAT (resolve_tac [nat_case_0 RS ssubst, + nat_case_succ RS ssubst] 1 + THEN resolve_tac prems 1)); +by (assume_tac 1); +val nat_case_type = result(); + +val prems = goalw Nat.thy [nat_case_def] + "[| n=n'; a=a'; !!m z. b(m)=b'(m) \ +\ |] ==> nat_case(n,a,b)=nat_case(n',a',b')"; +by (REPEAT (resolve_tac [the_cong,disj_cong,ex_cong] 1 + ORELSE EVERY1 (map rtac ((prems RL [ssubst]) @ [iff_refl])))); +val nat_case_cong = result(); + + +(** nat_rec -- used to define eclose and transrec, then obsolete **) + +val nat_rec_trans = wf_Memrel RS (nat_rec_def RS def_wfrec RS trans); + +goal Nat.thy "nat_rec(0,a,b) = a"; +by (rtac nat_rec_trans 1); +by (rtac nat_case_0 1); +val nat_rec_0 = result(); + +val [prem] = goal Nat.thy + "m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))"; +val nat_rec_ss = ZF_ss + addcongs (mk_typed_congs Nat.thy [("b", "[i,i]=>i")]) + addrews [prem, nat_case_succ, nat_succI, Memrel_iff, + vimage_singleton_iff]; +by (rtac nat_rec_trans 1); +by (SIMP_TAC nat_rec_ss 1); +val nat_rec_succ = result(); + +(** The union of two natural numbers is a natural number -- their maximum **) + +(* [| ?i : nat; ?j : nat |] ==> ?i Un ?j : nat *) +val Un_nat_type = standard (Ord_nat RSN (3,Ord_member_UnI)); + +(* [| ?i : nat; ?j : nat |] ==> ?i Int ?j : nat *) +val Int_nat_type = standard (Ord_nat RSN (3,Ord_member_IntI)); + diff -r 000000000000 -r a5a9c433f639 src/ZF/nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/nat.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,26 @@ +(* Title: ZF/nat.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Natural numbers in Zermelo-Fraenkel Set Theory +*) + +Nat = Ord + Bool + +consts + nat :: "i" + nat_case :: "[i, i, i=>i]=>i" + nat_rec :: "[i, i, [i,i]=>i]=>i" + +rules + + nat_def "nat == lfp(Inf, %X. {0} Un {succ(i). i:X})" + + nat_case_def + "nat_case(k,a,b) == THE y. k=0 & y=a | (EX x. k=succ(x) & y=b(x))" + + nat_rec_def + "nat_rec(k,a,b) == \ +\ wfrec(Memrel(nat), k, %n f. nat_case(n, a, %m. b(m, f`m)))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/ord.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ord.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,410 @@ +(* Title: ZF/ordinal.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For ordinal.thy. Ordinals in Zermelo-Fraenkel Set Theory +*) + +open Ord; + +(*** Rules for Transset ***) + +(** Two neat characterisations of Transset **) + +goalw Ord.thy [Transset_def] "Transset(A) <-> A<=Pow(A)"; +by (fast_tac ZF_cs 1); +val Transset_iff_Pow = result(); + +goalw Ord.thy [Transset_def] "Transset(A) <-> Union(succ(A)) = A"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val Transset_iff_Union_succ = result(); + +(** Consequences of downwards closure **) + +goalw Ord.thy [Transset_def] + "!!C a b. [| Transset(C); {a,b}: C |] ==> a:C & b: C"; +by (fast_tac ZF_cs 1); +val Transset_doubleton_D = result(); + +val [prem1,prem2] = goalw Ord.thy [Pair_def] + "[| Transset(C); : C |] ==> a:C & b: C"; +by (cut_facts_tac [prem2] 1); +by (fast_tac (ZF_cs addSDs [prem1 RS Transset_doubleton_D]) 1); +val Transset_Pair_D = result(); + +val prem1::prems = goal Ord.thy + "[| Transset(C); A*B <= C; b: B |] ==> A <= C"; +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addSDs [prem1 RS Transset_Pair_D]) 1); +val Transset_includes_domain = result(); + +val prem1::prems = goal Ord.thy + "[| Transset(C); A*B <= C; a: A |] ==> B <= C"; +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addSDs [prem1 RS Transset_Pair_D]) 1); +val Transset_includes_range = result(); + +val [prem1,prem2] = goalw (merge_theories(Ord.thy,Sum.thy)) [sum_def] + "[| Transset(C); A+B <= C |] ==> A <= C & B <= C"; +br (prem2 RS (Un_subset_iff RS iffD1) RS conjE) 1; +by (REPEAT (etac (prem1 RS Transset_includes_range) 1 + ORELSE resolve_tac [conjI, singletonI] 1)); +val Transset_includes_summands = result(); + +val [prem] = goalw (merge_theories(Ord.thy,Sum.thy)) [sum_def] + "Transset(C) ==> (A+B) Int C <= (A Int C) + (B Int C)"; +br (Int_Un_distrib RS ssubst) 1; +by (fast_tac (ZF_cs addSDs [prem RS Transset_Pair_D]) 1); +val Transset_sum_Int_subset = result(); + +(** Closure properties **) + +goalw Ord.thy [Transset_def] "Transset(0)"; +by (fast_tac ZF_cs 1); +val Transset_0 = result(); + +goalw Ord.thy [Transset_def] + "!!i j. [| Transset(i); Transset(j) |] ==> Transset(i Un j)"; +by (fast_tac ZF_cs 1); +val Transset_Un = result(); + +goalw Ord.thy [Transset_def] + "!!i j. [| Transset(i); Transset(j) |] ==> Transset(i Int j)"; +by (fast_tac ZF_cs 1); +val Transset_Int = result(); + +goalw Ord.thy [Transset_def] "!!i. Transset(i) ==> Transset(succ(i))"; +by (fast_tac ZF_cs 1); +val Transset_succ = result(); + +goalw Ord.thy [Transset_def] "!!i. Transset(i) ==> Transset(Pow(i))"; +by (fast_tac ZF_cs 1); +val Transset_Pow = result(); + +goalw Ord.thy [Transset_def] "!!A. Transset(A) ==> Transset(Union(A))"; +by (fast_tac ZF_cs 1); +val Transset_Union = result(); + +val [Transprem] = goalw Ord.thy [Transset_def] + "[| !!i. i:A ==> Transset(i) |] ==> Transset(Union(A))"; +by (fast_tac (ZF_cs addEs [Transprem RS bspec RS subsetD]) 1); +val Transset_Union_family = result(); + +val [prem,Transprem] = goalw Ord.thy [Transset_def] + "[| j:A; !!i. i:A ==> Transset(i) |] ==> Transset(Inter(A))"; +by (cut_facts_tac [prem] 1); +by (fast_tac (ZF_cs addEs [Transprem RS bspec RS subsetD]) 1); +val Transset_Inter_family = result(); + +(*** Natural Deduction rules for Ord ***) + +val prems = goalw Ord.thy [Ord_def] + "[| Transset(i); !!x. x:i ==> Transset(x) |] ==> Ord(i) "; +by (REPEAT (ares_tac (prems@[ballI,conjI]) 1)); +val OrdI = result(); + +val [major] = goalw Ord.thy [Ord_def] + "Ord(i) ==> Transset(i)"; +by (rtac (major RS conjunct1) 1); +val Ord_is_Transset = result(); + +val [major,minor] = goalw Ord.thy [Ord_def] + "[| Ord(i); j:i |] ==> Transset(j) "; +by (rtac (minor RS (major RS conjunct2 RS bspec)) 1); +val Ord_contains_Transset = result(); + +(*** Lemmas for ordinals ***) + +goalw Ord.thy [Ord_def,Transset_def] "!!i j. [| Ord(i); j:i |] ==> Ord(j) "; +by (fast_tac ZF_cs 1); +val Ord_in_Ord = result(); + +goal Ord.thy "!!i j. [| Ord(i); Transset(j); j<=i |] ==> Ord(j)"; +by (REPEAT (ares_tac [OrdI] 1 + ORELSE eresolve_tac [Ord_contains_Transset, subsetD] 1)); +val Ord_subset_Ord = result(); + +goalw Ord.thy [Ord_def,Transset_def] "!!i j. [| j:i; Ord(i) |] ==> j<=i"; +by (fast_tac ZF_cs 1); +val OrdmemD = result(); + +goal Ord.thy "!!i j k. [| i:j; j:k; Ord(k) |] ==> i:k"; +by (REPEAT (ares_tac [OrdmemD RS subsetD] 1)); +val Ord_trans = result(); + +goal Ord.thy "!!i j. [| i:j; Ord(j) |] ==> succ(i) <= j"; +by (REPEAT (ares_tac [OrdmemD RSN (2,succ_subsetI)] 1)); +val Ord_succ_subsetI = result(); + + +(*** The construction of ordinals: 0, succ, Union ***) + +goal Ord.thy "Ord(0)"; +by (REPEAT (ares_tac [OrdI,Transset_0] 1 ORELSE etac emptyE 1)); +val Ord_0 = result(); + +goal Ord.thy "!!i. Ord(i) ==> Ord(succ(i))"; +by (REPEAT (ares_tac [OrdI,Transset_succ] 1 + ORELSE eresolve_tac [succE,ssubst,Ord_is_Transset, + Ord_contains_Transset] 1)); +val Ord_succ = result(); + +val nonempty::prems = goal Ord.thy + "[| j:A; !!i. i:A ==> Ord(i) |] ==> Ord(Inter(A))"; +by (rtac (nonempty RS Transset_Inter_family RS OrdI) 1); +by (rtac Ord_is_Transset 1); +by (REPEAT (ares_tac ([Ord_contains_Transset,nonempty]@prems) 1 + ORELSE etac InterD 1)); +val Ord_Inter = result(); + +val jmemA::prems = goal Ord.thy + "[| j:A; !!x. x:A ==> Ord(B(x)) |] ==> Ord(INT x:A. B(x))"; +by (rtac (jmemA RS RepFunI RS Ord_Inter) 1); +by (etac RepFunE 1); +by (etac ssubst 1); +by (eresolve_tac prems 1); +val Ord_INT = result(); + + +(*** Natural Deduction rules for Memrel ***) + +goalw Ord.thy [Memrel_def] " : Memrel(A) <-> a:b & a:A & b:A"; +by (fast_tac ZF_cs 1); +val Memrel_iff = result(); + +val prems = goal Ord.thy "[| a: b; a: A; b: A |] ==> : Memrel(A)"; +by (REPEAT (resolve_tac (prems@[conjI, Memrel_iff RS iffD2]) 1)); +val MemrelI = result(); + +val [major,minor] = goal Ord.thy + "[| : Memrel(A); \ +\ [| a: A; b: A; a:b |] ==> P \ +\ |] ==> P"; +by (rtac (major RS (Memrel_iff RS iffD1) RS conjE) 1); +by (etac conjE 1); +by (rtac minor 1); +by (REPEAT (assume_tac 1)); +val MemrelE = result(); + +(*The membership relation (as a set) is well-founded. + Proof idea: show A<=B by applying the foundation axiom to A-B *) +goalw Ord.thy [wf_def] "wf(Memrel(A))"; +by (EVERY1 [rtac (foundation RS disjE RS allI), + etac disjI1, + etac bexE, + rtac (impI RS allI RS bexI RS disjI2), + etac MemrelE, + etac bspec, + REPEAT o assume_tac]); +val wf_Memrel = result(); + +(*** Transfinite induction ***) + +(*Epsilon induction over a transitive set*) +val major::prems = goalw Ord.thy [Transset_def] + "[| i: k; Transset(k); \ +\ !!x.[| x: k; ALL y:x. P(y) |] ==> P(x) \ +\ |] ==> P(i)"; +by (rtac (major RS (wf_Memrel RS wf_induct2)) 1); +by (fast_tac (ZF_cs addEs [MemrelE]) 1); +by (resolve_tac prems 1); +by (assume_tac 1); +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addIs [MemrelI]) 1); +val Transset_induct = result(); + +(*Induction over an ordinal*) +val Ord_induct = Ord_is_Transset RSN (2, Transset_induct); + +(*Induction over the class of ordinals -- a useful corollary of Ord_induct*) +val [major,indhyp] = goal Ord.thy + "[| Ord(i); \ +\ !!x.[| Ord(x); ALL y:x. P(y) |] ==> P(x) \ +\ |] ==> P(i)"; +by (rtac (major RS Ord_succ RS (succI1 RS Ord_induct)) 1); +by (rtac indhyp 1); +by (rtac (major RS Ord_succ RS Ord_in_Ord) 1); +by (REPEAT (assume_tac 1)); +val trans_induct = result(); + +(*Perform induction on i, then prove the Ord(i) subgoal using prems. *) +fun trans_ind_tac a prems i = + EVERY [res_inst_tac [("i",a)] trans_induct i, + rename_last_tac a ["1"] (i+1), + ares_tac prems i]; + + +(*** Fundamental properties of the epsilon ordering (< on ordinals) ***) + +(*Finds contradictions for the following proof*) +val Ord_trans_tac = EVERY' [etac notE, etac Ord_trans, REPEAT o atac]; + +(** Proving that "member" is a linear ordering on the ordinals **) + +val prems = goal Ord.thy + "Ord(i) ==> (ALL j. Ord(j) --> i:j | i=j | j:i)"; +by (trans_ind_tac "i" prems 1); +by (rtac (impI RS allI) 1); +by (trans_ind_tac "j" [] 1); +by (DEPTH_SOLVE (swap_res_tac [disjCI,equalityI,subsetI] 1 + ORELSE ball_tac 1 + ORELSE eresolve_tac [impE,disjE,allE] 1 + ORELSE hyp_subst_tac 1 + ORELSE Ord_trans_tac 1)); +val Ord_linear_lemma = result(); + +val ordi::ordj::prems = goal Ord.thy + "[| Ord(i); Ord(j); i:j ==> P; i=j ==> P; j:i ==> P |] \ +\ ==> P"; +by (rtac (ordi RS Ord_linear_lemma RS spec RS impE) 1); +by (rtac ordj 1); +by (REPEAT (eresolve_tac (prems@[asm_rl,disjE]) 1)); +val Ord_linear = result(); + +val prems = goal Ord.thy + "[| Ord(i); Ord(j); i<=j ==> P; j<=i ==> P |] \ +\ ==> P"; +by (res_inst_tac [("i","i"),("j","j")] Ord_linear 1); +by (DEPTH_SOLVE (ares_tac (prems@[subset_refl]) 1 + ORELSE eresolve_tac [asm_rl,OrdmemD,ssubst] 1)); +val Ord_subset = result(); + +goal Ord.thy "!!i j. [| j<=i; ~ i<=j; Ord(i); Ord(j) |] ==> j:i"; +by (etac Ord_linear 1); +by (REPEAT (ares_tac [subset_refl] 1 + ORELSE eresolve_tac [notE,OrdmemD,ssubst] 1)); +val Ord_member = result(); + +val [prem] = goal Ord.thy "Ord(i) ==> 0: succ(i)"; +by (rtac (empty_subsetI RS Ord_member) 1); +by (fast_tac ZF_cs 1); +by (rtac (prem RS Ord_succ) 1); +by (rtac Ord_0 1); +val Ord_0_mem_succ = result(); + +goal Ord.thy "!!i j. [| Ord(i); Ord(j) |] ==> j:i <-> j<=i & ~(i<=j)"; +by (rtac iffI 1); +by (rtac conjI 1); +by (etac OrdmemD 1); +by (rtac (mem_anti_refl RS notI) 2); +by (etac subsetD 2); +by (REPEAT (eresolve_tac [asm_rl, conjE, Ord_member] 1)); +val Ord_member_iff = result(); + +goal Ord.thy "!!i. Ord(i) ==> 0:i <-> ~ i=0"; +be (Ord_0 RSN (2,Ord_member_iff) RS iff_trans) 1; +by (fast_tac eq_cs 1); +val Ord_0_member_iff = result(); + +(** For ordinals, i: succ(j) means 'less-than or equals' **) + +goal Ord.thy "!!i j. [| j<=i; Ord(i); Ord(j) |] ==> j : succ(i)"; +by (rtac Ord_member 1); +by (REPEAT (ares_tac [Ord_succ] 3)); +by (rtac (mem_anti_refl RS notI) 2); +by (etac subsetD 2); +by (ALLGOALS (fast_tac ZF_cs)); +val member_succI = result(); + +goalw Ord.thy [Transset_def,Ord_def] + "!!i j. [| i : succ(j); Ord(j) |] ==> i<=j"; +by (fast_tac ZF_cs 1); +val member_succD = result(); + +goal Ord.thy "!!i j. [| Ord(i); Ord(j) |] ==> j:succ(i) <-> j<=i"; +by (fast_tac (subset_cs addSEs [member_succI, member_succD]) 1); +val member_succ_iff = result(); + +goal Ord.thy + "!!i j. [| Ord(i); Ord(j) |] ==> i<=succ(j) <-> i<=j | i=succ(j)"; +by (ASM_SIMP_TAC (ZF_ss addrews [member_succ_iff RS iff_sym, Ord_succ]) 1); +by (fast_tac ZF_cs 1); +val subset_succ_iff = result(); + +goal Ord.thy "!!i j. [| i:succ(j); j:k; Ord(k) |] ==> i:k"; +by (fast_tac (ZF_cs addEs [Ord_trans]) 1); +val Ord_trans1 = result(); + +goal Ord.thy "!!i j. [| i:j; j:succ(k); Ord(k) |] ==> i:k"; +by (fast_tac (ZF_cs addEs [Ord_trans]) 1); +val Ord_trans2 = result(); + +goal Ord.thy "!!i jk. [| i:j; j<=k; Ord(k) |] ==> i:k"; +by (fast_tac (ZF_cs addEs [Ord_trans]) 1); +val Ord_transsub2 = result(); + +goal Ord.thy "!!i j. [| i:j; Ord(j) |] ==> succ(i) : succ(j)"; +by (rtac member_succI 1); +by (REPEAT (ares_tac [subsetI,Ord_succ,Ord_in_Ord] 1 + ORELSE eresolve_tac [succE,Ord_trans,ssubst] 1)); +val succ_mem_succI = result(); + +goal Ord.thy "!!i j. [| succ(i) : succ(j); Ord(j) |] ==> i:j"; +by (REPEAT (eresolve_tac [asm_rl, make_elim member_succD, succ_subsetE] 1)); +val succ_mem_succE = result(); + +goal Ord.thy "!!i j. Ord(j) ==> succ(i) : succ(j) <-> i:j"; +by (REPEAT (ares_tac [iffI,succ_mem_succI,succ_mem_succE] 1)); +val succ_mem_succ_iff = result(); + +goal Ord.thy "!!i j. [| i<=j; Ord(i); Ord(j) |] ==> succ(i) <= succ(j)"; +by (rtac (member_succI RS succ_mem_succI RS member_succD) 1); +by (REPEAT (ares_tac [Ord_succ] 1)); +val Ord_succ_mono = result(); + +goal Ord.thy "!!i j k. [| i:k; j:k; Ord(k) |] ==> i Un j : k"; +by (res_inst_tac [("i","i"),("j","j")] Ord_subset 1); +by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 1)); +by (ASM_SIMP_TAC (ZF_ss addrews [subset_Un_iff RS iffD1]) 1); +by (rtac (Un_commute RS ssubst) 1); +by (ASM_SIMP_TAC (ZF_ss addrews [subset_Un_iff RS iffD1]) 1); +val Ord_member_UnI = result(); + +goal Ord.thy "!!i j k. [| i:k; j:k; Ord(k) |] ==> i Int j : k"; +by (res_inst_tac [("i","i"),("j","j")] Ord_subset 1); +by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 1)); +by (ASM_SIMP_TAC (ZF_ss addrews [subset_Int_iff RS iffD1]) 1); +by (rtac (Int_commute RS ssubst) 1); +by (ASM_SIMP_TAC (ZF_ss addrews [subset_Int_iff RS iffD1]) 1); +val Ord_member_IntI = result(); + + +(*** Results about limits ***) + +val prems = goal Ord.thy "[| !!i. i:A ==> Ord(i) |] ==> Ord(Union(A))"; +by (rtac (Ord_is_Transset RS Transset_Union_family RS OrdI) 1); +by (REPEAT (etac UnionE 1 ORELSE ares_tac ([Ord_contains_Transset]@prems) 1)); +val Ord_Union = result(); + +val prems = goal Ord.thy "[| !!x. x:A ==> Ord(B(x)) |] ==> Ord(UN x:A. B(x))"; +by (rtac Ord_Union 1); +by (etac RepFunE 1); +by (etac ssubst 1); +by (eresolve_tac prems 1); +val Ord_UN = result(); + +(*The upper bound must be a successor ordinal -- + consider that (UN i:nat.i)=nat although nat is an upper bound of each i*) +val [ordi,limit] = goal Ord.thy + "[| Ord(i); !!y. y:A ==> f(y): succ(i) |] ==> (UN y:A. f(y)) : succ(i)"; +by (rtac (member_succD RS UN_least RS member_succI) 1); +by (REPEAT (ares_tac [ordi, Ord_UN, ordi RS Ord_succ RS Ord_in_Ord, + limit] 1)); +val sup_least = result(); + +val [jmemi,ordi,limit] = goal Ord.thy + "[| j: i; Ord(i); !!y. y:A ==> f(y): j |] ==> (UN y:A. succ(f(y))) : i"; +by (cut_facts_tac [jmemi RS (ordi RS Ord_in_Ord)] 1); +by (rtac (sup_least RS Ord_trans2) 1); +by (REPEAT (ares_tac [jmemi, ordi, succ_mem_succI, limit] 1)); +val sup_least2 = result(); + +goal Ord.thy "!!i. Ord(i) ==> (UN y:i. succ(y)) = i"; +by (fast_tac (eq_cs addSEs [Ord_trans1]) 1); +val Ord_equality = result(); + +(*Holds for all transitive sets, not just ordinals*) +goal Ord.thy "!!i. Ord(i) ==> Union(i) <= i"; +by (fast_tac (ZF_cs addSEs [Ord_trans]) 1); +val Ord_Union_subset = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/ord.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/ord.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,19 @@ +(* Title: ZF/ordinal.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Ordinals in Zermelo-Fraenkel Set Theory +*) + +Ord = WF + +consts + Memrel :: "i=>i" + Transset,Ord :: "i=>o" + +rules + Memrel_def "Memrel(A) == {z: A*A . EX x y. z= & x:y }" + Transset_def "Transset(i) == ALL x:i. x<=i" + Ord_def "Ord(i) == Transset(i) & (ALL x:i. Transset(x))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/pair.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/pair.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,153 @@ +(* Title: ZF/pair + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Ordered pairs in Zermelo-Fraenkel Set Theory +*) + +(** Lemmas for showing that uniquely determines a and b **) + +val doubleton_iff = prove_goal ZF.thy + "{a,b} = {c,d} <-> (a=c & b=d) | (a=d & b=c)" + (fn _=> [ (resolve_tac [extension RS iff_trans] 1), + (fast_tac upair_cs 1) ]); + +val Pair_iff = prove_goalw ZF.thy [Pair_def] + " = <-> a=c & b=d" + (fn _=> [ (SIMP_TAC (FOL_ss addrews [doubleton_iff]) 1), + (fast_tac FOL_cs 1) ]); + +val Pair_inject = standard (Pair_iff RS iffD1 RS conjE); + +val Pair_inject1 = prove_goal ZF.thy " = ==> a=c" + (fn [major]=> + [ (rtac (major RS Pair_inject) 1), (assume_tac 1) ]); + +val Pair_inject2 = prove_goal ZF.thy " = ==> b=d" + (fn [major]=> + [ (rtac (major RS Pair_inject) 1), (assume_tac 1) ]); + +val Pair_neq_0 = prove_goalw ZF.thy [Pair_def] "=0 ==> P" + (fn [major]=> + [ (rtac (major RS equalityD1 RS subsetD RS emptyE) 1), + (rtac consI1 1) ]); + +val Pair_neq_fst = prove_goalw ZF.thy [Pair_def] "=a ==> P" + (fn [major]=> + [ (rtac (consI1 RS mem_anti_sym RS FalseE) 1), + (rtac (major RS subst) 1), + (rtac consI1 1) ]); + +val Pair_neq_snd = prove_goalw ZF.thy [Pair_def] "=b ==> P" + (fn [major]=> + [ (rtac (consI1 RS consI2 RS mem_anti_sym RS FalseE) 1), + (rtac (major RS subst) 1), + (rtac (consI1 RS consI2) 1) ]); + + +(*** Sigma: Disjoint union of a family of sets + Generalizes Cartesian product ***) + +val SigmaI = prove_goalw ZF.thy [Sigma_def] + "[| a:A; b:B(a) |] ==> : Sigma(A,B)" + (fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]); + +(*The general elimination rule*) +val SigmaE = prove_goalw ZF.thy [Sigma_def] + "[| c: Sigma(A,B); \ +\ !!x y.[| x:A; y:B(x); c= |] ==> P \ +\ |] ==> P" + (fn major::prems=> + [ (cut_facts_tac [major] 1), + (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]); + +(** Elimination of :A*B -- introduces no eigenvariables **) +val SigmaD1 = prove_goal ZF.thy " : Sigma(A,B) ==> a : A" + (fn [major]=> + [ (rtac (major RS SigmaE) 1), + (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); + +val SigmaD2 = prove_goal ZF.thy " : Sigma(A,B) ==> b : B(a)" + (fn [major]=> + [ (rtac (major RS SigmaE) 1), + (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); + +(*Also provable via + rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac) + THEN prune_params_tac) + (read_instantiate [("c","")] SigmaE); *) +val SigmaE2 = prove_goal ZF.thy + "[| : Sigma(A,B); \ +\ [| a:A; b:B(a) |] ==> P \ +\ |] ==> P" + (fn [major,minor]=> + [ (rtac minor 1), + (rtac (major RS SigmaD1) 1), + (rtac (major RS SigmaD2) 1) ]); + +val Sigma_cong = prove_goalw ZF.thy [Sigma_def] + "[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> \ +\ Sigma(A,B) = Sigma(A',B')" + (fn prems=> [ (prove_cong_tac (prems@[RepFun_cong]) 1) ]); + +val Sigma_empty1 = prove_goal ZF.thy "Sigma(0,B) = 0" + (fn _ => [ (fast_tac (lemmas_cs addIs [equalityI] addSEs [SigmaE]) 1) ]); + +val Sigma_empty2 = prove_goal ZF.thy "A*0 = 0" + (fn _ => [ (fast_tac (lemmas_cs addIs [equalityI] addSEs [SigmaE]) 1) ]); + + +(*** Eliminator - split ***) + +val split = prove_goalw ZF.thy [split_def] + "split(%x y.c(x,y), ) = c(a,b)" + (fn _ => + [ (fast_tac (upair_cs addIs [the_equality] addEs [Pair_inject]) 1) ]); + +val split_type = prove_goal ZF.thy + "[| p:Sigma(A,B); \ +\ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C() \ +\ |] ==> split(%x y.c(x,y), p) : C(p)" + (fn major::prems=> + [ (rtac (major RS SigmaE) 1), + (etac ssubst 1), + (REPEAT (ares_tac (prems @ [split RS ssubst]) 1)) ]); + +(*This congruence rule uses NO typing information...*) +val split_cong = prove_goalw ZF.thy [split_def] + "[| p=p'; !!x y.c(x,y) = c'(x,y) |] ==> \ +\ split(%x y.c(x,y), p) = split(%x y.c'(x,y), p')" + (fn prems=> [ (prove_cong_tac (prems@[the_cong]) 1) ]); + + +(*** conversions for fst and snd ***) + +val fst_conv = prove_goalw ZF.thy [fst_def] "fst() = a" + (fn _=> [ (rtac split 1) ]); + +val snd_conv = prove_goalw ZF.thy [snd_def] "snd() = b" + (fn _=> [ (rtac split 1) ]); + + +(*** split for predicates: result type o ***) + +goalw ZF.thy [fsplit_def] "!!R a b. R(a,b) ==> fsplit(R, )"; +by (REPEAT (ares_tac [refl,exI,conjI] 1)); +val fsplitI = result(); + +val major::prems = goalw ZF.thy [fsplit_def] + "[| fsplit(R,z); !!x y. [| z = ; R(x,y) |] ==> P |] ==> P"; +by (cut_facts_tac [major] 1); +by (REPEAT (eresolve_tac (prems@[asm_rl,exE,conjE]) 1)); +val fsplitE = result(); + +goal ZF.thy "!!R a b. fsplit(R,) ==> R(a,b)"; +by (REPEAT (eresolve_tac [asm_rl,fsplitE,Pair_inject,ssubst] 1)); +val fsplitD = result(); + +val pair_cs = upair_cs + addSIs [SigmaI] + addSEs [SigmaE2, SigmaE, Pair_inject, make_elim succ_inject, + Pair_neq_0, sym RS Pair_neq_0, succ_neq_0, sym RS succ_neq_0]; + diff -r 000000000000 -r a5a9c433f639 src/ZF/perm.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/perm.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,433 @@ +(* Title: ZF/perm.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For perm.thy. The theory underlying permutation groups + -- Composition of relations, the identity relation + -- Injections, surjections, bijections + -- Lemmas for the Schroeder-Bernstein Theorem +*) + +open Perm; + +(** Surjective function space **) + +goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> f: A->B"; +by (etac CollectD1 1); +val surj_is_fun = result(); + +goalw Perm.thy [surj_def] "!!f A B. f : Pi(A,B) ==> f: surj(A,range(f))"; +by (fast_tac (ZF_cs addIs [apply_equality] + addEs [range_of_fun,domain_type]) 1); +val fun_is_surj = result(); + +goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> range(f)=B"; +by (best_tac (ZF_cs addIs [equalityI,apply_Pair] addEs [range_type]) 1); +val surj_range = result(); + + +(** Injective function space **) + +goalw Perm.thy [inj_def] "!!f A B. f: inj(A,B) ==> f: A->B"; +by (etac CollectD1 1); +val inj_is_fun = result(); + +goalw Perm.thy [inj_def] + "!!f A B. [| :f; :f; f: inj(A,B) |] ==> a=c"; +by (REPEAT (eresolve_tac [asm_rl, Pair_mem_PiE, CollectE] 1)); +by (fast_tac ZF_cs 1); +val inj_equality = result(); + +(** Bijections -- simple lemmas but no intro/elim rules -- use unfolding **) + +goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> f: inj(A,B)"; +by (etac IntD1 1); +val bij_is_inj = result(); + +goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> f: surj(A,B)"; +by (etac IntD2 1); +val bij_is_surj = result(); + +(* f: bij(A,B) ==> f: A->B *) +val bij_is_fun = standard (bij_is_inj RS inj_is_fun); + +(** Identity function **) + +val [prem] = goalw Perm.thy [id_def] "a:A ==> : id(A)"; +by (rtac (prem RS lamI) 1); +val idI = result(); + +val major::prems = goalw Perm.thy [id_def] + "[| p: id(A); !!x.[| x:A; p= |] ==> P \ +\ |] ==> P"; +by (rtac (major RS lamE) 1); +by (REPEAT (ares_tac prems 1)); +val idE = result(); + +goalw Perm.thy [id_def] "id(A) : A->A"; +by (rtac lam_type 1); +by (assume_tac 1); +val id_type = result(); + +val [prem] = goalw Perm.thy [id_def] "A<=B ==> id(A) <= id(B)"; +by (rtac (prem RS lam_mono) 1); +val id_mono = result(); + +goalw Perm.thy [inj_def,id_def] "id(A): inj(A,A)"; +by (REPEAT (ares_tac [CollectI,lam_type] 1)); +by (SIMP_TAC ZF_ss 1); +val id_inj = result(); + +goalw Perm.thy [id_def,surj_def] "id(A): surj(A,A)"; +by (fast_tac (ZF_cs addIs [lam_type,beta]) 1); +val id_surj = result(); + +goalw Perm.thy [bij_def] "id(A): bij(A,A)"; +by (fast_tac (ZF_cs addIs [id_inj,id_surj]) 1); +val id_bij = result(); + + +(** Converse of a relation **) + +val [prem] = goal Perm.thy "f: inj(A,B) ==> converse(f) : range(f)->A"; +by (rtac (prem RS inj_is_fun RS PiE) 1); +by (rtac (converse_type RS PiI) 1); +by (fast_tac ZF_cs 1); +by (fast_tac (ZF_cs addEs [prem RSN (3,inj_equality)]) 1); +by flexflex_tac; +val inj_converse_fun = result(); + +val prems = goalw Perm.thy [surj_def] + "f: inj(A,B) ==> converse(f): surj(range(f), A)"; +by (fast_tac (ZF_cs addIs (prems@[inj_converse_fun,apply_Pair,apply_equality, + converseI,inj_is_fun])) 1); +val inj_converse_surj = result(); + +(*The premises are equivalent to saying that f is injective...*) +val prems = goal Perm.thy + "[| f: A->B; converse(f): C->A; a: A |] ==> converse(f)`(f`a) = a"; +by (fast_tac (ZF_cs addIs (prems@[apply_Pair,apply_equality,converseI])) 1); +val left_inverse_lemma = result(); + +val prems = goal Perm.thy + "[| f: inj(A,B); a: A |] ==> converse(f)`(f`a) = a"; +by (fast_tac (ZF_cs addIs (prems@ + [left_inverse_lemma,inj_converse_fun,inj_is_fun])) 1); +val left_inverse = result(); + +val prems = goal Perm.thy + "[| f: A->B; converse(f): C->A; b: C |] ==> f`(converse(f)`b) = b"; +by (rtac (apply_Pair RS (converseD RS apply_equality)) 1); +by (REPEAT (resolve_tac prems 1)); +val right_inverse_lemma = result(); + +val prems = goal Perm.thy + "[| f: inj(A,B); b: range(f) |] ==> f`(converse(f)`b) = b"; +by (rtac right_inverse_lemma 1); +by (REPEAT (resolve_tac (prems@ [inj_converse_fun,inj_is_fun]) 1)); +val right_inverse = result(); + +val prems = goal Perm.thy + "f: inj(A,B) ==> converse(f): inj(range(f), A)"; +bw inj_def; (*rewrite subgoal but not prems!!*) +by (cut_facts_tac prems 1); +by (safe_tac ZF_cs); +(*apply f to both sides and simplify using right_inverse + -- could also use etac[subst_context RS box_equals] in this proof *) +by (rtac simp_equals 2); +by (REPEAT (eresolve_tac [inj_converse_fun, right_inverse RS sym, ssubst] 1 + ORELSE ares_tac [refl,rangeI] 1)); +val inj_converse_inj = result(); + +goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> converse(f): bij(B,A)"; +by (etac IntE 1); +by (eresolve_tac [(surj_range RS subst)] 1); +by (rtac IntI 1); +by (etac inj_converse_inj 1); +by (etac inj_converse_surj 1); +val bij_converse_bij = result(); + + +(** Composition of two relations **) + +(*The inductive definition package could derive these theorems for R o S*) + +goalw Perm.thy [comp_def] "!!r s. [| :s; :r |] ==> : r O s"; +by (fast_tac ZF_cs 1); +val compI = result(); + +val prems = goalw Perm.thy [comp_def] + "[| xz : r O s; \ +\ !!x y z. [| xz=; :s; :r |] ==> P \ +\ |] ==> P"; +by (cut_facts_tac prems 1); +by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1)); +val compE = result(); + +val compEpair = + rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac) + THEN prune_params_tac) + (read_instantiate [("xz","")] compE); + +val comp_cs = ZF_cs addIs [compI,idI] addSEs [compE,idE]; + +(** Domain and Range -- see Suppes, section 3.1 **) + +(*Boyer et al., Set Theory in First-Order Logic, JAR 2 (1986), 287-327*) +goal Perm.thy "range(r O s) <= range(r)"; +by (fast_tac comp_cs 1); +val range_comp = result(); + +goal Perm.thy "!!r s. domain(r) <= range(s) ==> range(r O s) = range(r)"; +by (rtac (range_comp RS equalityI) 1); +by (fast_tac comp_cs 1); +val range_comp_eq = result(); + +goal Perm.thy "domain(r O s) <= domain(s)"; +by (fast_tac comp_cs 1); +val domain_comp = result(); + +goal Perm.thy "!!r s. range(s) <= domain(r) ==> domain(r O s) = domain(s)"; +by (rtac (domain_comp RS equalityI) 1); +by (fast_tac comp_cs 1); +val domain_comp_eq = result(); + +(** Other results **) + +goal Perm.thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; +by (fast_tac comp_cs 1); +val comp_mono = result(); + +(*composition preserves relations*) +goal Perm.thy "!!r s. [| s<=A*B; r<=B*C |] ==> (r O s) <= A*C"; +by (fast_tac comp_cs 1); +val comp_rel = result(); + +(*associative law for composition*) +goal Perm.thy "(r O s) O t = r O (s O t)"; +by (fast_tac (comp_cs addIs [equalityI]) 1); +val comp_assoc = result(); + +(*left identity of composition; provable inclusions are + id(A) O r <= r + and [| r<=A*B; B<=C |] ==> r <= id(C) O r *) +goal Perm.thy "!!r A B. r<=A*B ==> id(B) O r = r"; +by (fast_tac (comp_cs addIs [equalityI]) 1); +val left_comp_id = result(); + +(*right identity of composition; provable inclusions are + r O id(A) <= r + and [| r<=A*B; A<=C |] ==> r <= r O id(C) *) +goal Perm.thy "!!r A B. r<=A*B ==> r O id(A) = r"; +by (fast_tac (comp_cs addIs [equalityI]) 1); +val right_comp_id = result(); + + +(** Composition preserves functions, injections, and surjections **) + +goal Perm.thy "!!f g. [| g: A->B; f: B->C |] ==> (f O g) : A->C"; +by (REPEAT (ares_tac [PiI,comp_rel,ex1I,compI] 1 + ORELSE eresolve_tac [fun_is_rel,apply_Pair,apply_type] 1)); +by (fast_tac (comp_cs addDs [apply_equality]) 1); +val comp_func = result(); + +goal Perm.thy "!!f g. [| g: A->B; f: B->C; a:A |] ==> (f O g)`a = f`(g`a)"; +by (REPEAT (ares_tac [comp_func,apply_equality,compI, + apply_Pair,apply_type] 1)); +val comp_func_apply = result(); + +goalw Perm.thy [inj_def] + "!!f g. [| g: inj(A,B); f: inj(B,C) |] ==> (f O g) : inj(A,C)"; +by (REPEAT (eresolve_tac [bspec RS bspec RS mp, box_equals] 1 + ORELSE step_tac (ZF_cs addSIs [comp_func,apply_type,comp_func_apply]) 1)); +val comp_inj = result(); + +goalw Perm.thy [surj_def] + "!!f g. [| g: surj(A,B); f: surj(B,C) |] ==> (f O g) : surj(A,C)"; +by (best_tac (ZF_cs addSIs [comp_func,comp_func_apply]) 1); +val comp_surj = result(); + +goalw Perm.thy [bij_def] + "!!f g. [| g: bij(A,B); f: bij(B,C) |] ==> (f O g) : bij(A,C)"; +by (fast_tac (ZF_cs addIs [comp_inj,comp_surj]) 1); +val comp_bij = result(); + + +(** Dual properties of inj and surj -- useful for proofs from + D Pastre. Automatic theorem proving in set theory. + Artificial Intelligence, 10:1--27, 1978. **) + +goalw Perm.thy [inj_def] + "!!f g. [| (f O g): inj(A,C); g: A->B; f: B->C |] ==> g: inj(A,B)"; +by (safe_tac comp_cs); +by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1)); +by (ASM_SIMP_TAC (ZF_ss addrews [comp_func_apply]) 1); +val comp_mem_injD1 = result(); + +goalw Perm.thy [inj_def,surj_def] + "!!f g. [| (f O g): inj(A,C); g: surj(A,B); f: B->C |] ==> f: inj(B,C)"; +by (safe_tac comp_cs); +by (res_inst_tac [("x1", "x")] (bspec RS bexE) 1); +by (eres_inst_tac [("x1", "w")] (bspec RS bexE) 3); +by (REPEAT (assume_tac 1)); +by (safe_tac (comp_cs addSIs ZF_congs)); +by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1)); +by (ASM_SIMP_TAC (ZF_ss addrews [comp_func_apply]) 1); +val comp_mem_injD2 = result(); + +goalw Perm.thy [surj_def] + "!!f g. [| (f O g): surj(A,C); g: A->B; f: B->C |] ==> f: surj(B,C)"; +by (fast_tac (comp_cs addSIs [comp_func_apply RS sym, apply_type]) 1); +val comp_mem_surjD1 = result(); + +goal Perm.thy + "!!f g. [| (f O g)`a = c; g: A->B; f: B->C; a:A |] ==> f`(g`a) = c"; +by (REPEAT (ares_tac [comp_func_apply RS sym RS trans] 1)); +val comp_func_applyD = result(); + +goalw Perm.thy [inj_def,surj_def] + "!!f g. [| (f O g): surj(A,C); g: A->B; f: inj(B,C) |] ==> g: surj(A,B)"; +by (safe_tac comp_cs); +by (eres_inst_tac [("x1", "f`y")] (bspec RS bexE) 1); +by (REPEAT (ares_tac [apply_type] 1 ORELSE dtac comp_func_applyD 1)); +by (best_tac (comp_cs addSIs [apply_type]) 1); +val comp_mem_surjD2 = result(); + + +(** inverses of composition **) + +(*left inverse of composition; one inclusion is + f: A->B ==> id(A) <= converse(f) O f *) +val [prem] = goal Perm.thy + "f: inj(A,B) ==> converse(f) O f = id(A)"; +val injfD = prem RSN (3,inj_equality); +by (cut_facts_tac [prem RS inj_is_fun] 1); +by (fast_tac (comp_cs addIs [equalityI,apply_Pair] + addEs [domain_type, make_elim injfD]) 1); +val left_comp_inverse = result(); + +(*right inverse of composition; one inclusion is + f: A->B ==> f O converse(f) <= id(B) *) +val [prem] = goalw Perm.thy [surj_def] + "f: surj(A,B) ==> f O converse(f) = id(B)"; +val appfD = (prem RS CollectD1) RSN (3,apply_equality2); +by (cut_facts_tac [prem] 1); +by (rtac equalityI 1); +by (best_tac (comp_cs addEs [domain_type, range_type, make_elim appfD]) 1); +by (best_tac (comp_cs addIs [apply_Pair]) 1); +val right_comp_inverse = result(); + +(*Injective case applies converse(f) to both sides then simplifies + using left_inverse_lemma*) +goalw Perm.thy [bij_def,inj_def,surj_def] + "!!f A B. [| converse(f): B->A; f: A->B |] ==> f : bij(A,B)"; +val cf_cong = read_instantiate_sg (sign_of Perm.thy) + [("t","%x.?f`x")] subst_context; +by (fast_tac (ZF_cs addIs [left_inverse_lemma, right_inverse_lemma, apply_type] + addEs [cf_cong RS box_equals]) 1); +val invertible_imp_bijective = result(); + +(** Unions of functions -- cf similar theorems on func.ML **) + +goal Perm.thy "converse(r Un s) = converse(r) Un converse(s)"; +by (rtac equalityI 1); +by (DEPTH_SOLVE_1 + (resolve_tac [Un_least,converse_mono, Un_upper1,Un_upper2] 2)); +by (fast_tac ZF_cs 1); +val converse_of_Un = result(); + +goalw Perm.thy [surj_def] + "!!f g. [| f: surj(A,B); g: surj(C,D); A Int C = 0 |] ==> \ +\ (f Un g) : surj(A Un C, B Un D)"; +by (DEPTH_SOLVE_1 (eresolve_tac [fun_disjoint_apply1, fun_disjoint_apply2] 1 + ORELSE ball_tac 1 + ORELSE (rtac trans 1 THEN atac 2) + ORELSE step_tac (ZF_cs addIs [fun_disjoint_Un]) 1)); +val surj_disjoint_Un = result(); + +(*A simple, high-level proof; the version for injections follows from it, + using f:inj(A,B)<->f:bij(A,range(f)) *) +goal Perm.thy + "!!f g. [| f: bij(A,B); g: bij(C,D); A Int C = 0; B Int D = 0 |] ==> \ +\ (f Un g) : bij(A Un C, B Un D)"; +by (rtac invertible_imp_bijective 1); +by (rtac (converse_of_Un RS ssubst) 1); +by (REPEAT (ares_tac [fun_disjoint_Un, bij_is_fun, bij_converse_bij] 1)); +val bij_disjoint_Un = result(); + + +(** Restrictions as surjections and bijections *) + +val prems = goalw Perm.thy [surj_def] + "f: Pi(A,B) ==> f: surj(A, f``A)"; +val rls = apply_equality :: (prems RL [apply_Pair,Pi_type]); +by (fast_tac (ZF_cs addIs rls) 1); +val surj_image = result(); + +goal Perm.thy + "!!f. [| f: Pi(C,B); A<=C |] ==> restrict(f,A)``A = f``A"; +by (rtac equalityI 1); +by (SELECT_GOAL (rewtac restrict_def) 2); +by (REPEAT (eresolve_tac [imageE, apply_equality RS subst] 2 + ORELSE ares_tac [subsetI,lamI,imageI] 2)); +by (REPEAT (ares_tac [image_mono,restrict_subset,subset_refl] 1)); +val restrict_image = result(); + +goalw Perm.thy [inj_def] + "!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): inj(C,B)"; +by (safe_tac (ZF_cs addSEs [restrict_type2])); +by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp, subsetD, + box_equals, restrict] 1)); +val restrict_inj = result(); + +val prems = goal Perm.thy + "[| f: Pi(A,B); C<=A |] ==> restrict(f,C): surj(C, f``C)"; +by (rtac (restrict_image RS subst) 1); +by (rtac (restrict_type2 RS surj_image) 3); +by (REPEAT (resolve_tac prems 1)); +val restrict_surj = result(); + +goalw Perm.thy [inj_def,bij_def] + "!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)"; +by (safe_tac ZF_cs); +by (REPEAT (eresolve_tac [bspec RS bspec RS mp, subsetD, + box_equals, restrict] 1 + ORELSE ares_tac [surj_is_fun,restrict_surj] 1)); +val restrict_bij = result(); + + +(*** Lemmas for Ramsey's Theorem ***) + +goalw Perm.thy [inj_def] "!!f. [| f: inj(A,B); B<=D |] ==> f: inj(A,D)"; +by (fast_tac (ZF_cs addSEs [fun_weaken_type]) 1); +val inj_weaken_type = result(); + +val [major] = goal Perm.thy + "[| f: inj(succ(m), A) |] ==> restrict(f,m) : inj(m, A-{f`m})"; +by (rtac (major RS restrict_bij RS bij_is_inj RS inj_weaken_type) 1); +by (fast_tac ZF_cs 1); +by (cut_facts_tac [major] 1); +by (rewtac inj_def); +by (safe_tac ZF_cs); +by (etac range_type 1); +by (assume_tac 1); +by (dtac apply_equality 1); +by (assume_tac 1); +by (res_inst_tac [("a","m")] mem_anti_refl 1); +by (fast_tac ZF_cs 1); +val inj_succ_restrict = result(); + +goalw Perm.thy [inj_def] + "!!f. [| f: inj(A,B); ~ a:A; ~ b:B |] ==> \ +\ cons(,f) : inj(cons(a,A), cons(b,B))"; +(*cannot use safe_tac: must preserve the implication*) +by (etac CollectE 1); +by (rtac CollectI 1); +by (etac fun_extend 1); +by (REPEAT (ares_tac [ballI] 1)); +by (REPEAT_FIRST (eresolve_tac [consE,ssubst])); +(*Assumption ALL w:A. ALL x:A. f`w = f`x --> w=x makes ASM_SIMP_TAC loop!*) +by (ALLGOALS (SIMP_TAC (ZF_ss addrews [fun_extend_apply2,fun_extend_apply1]))); +by (ALLGOALS (fast_tac (ZF_cs addIs [apply_type]))); +val inj_extend = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/perm.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/perm.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,36 @@ +(* Title: ZF/perm + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +The theory underlying permutation groups + -- Composition of relations, the identity relation + -- Injections, surjections, bijections + -- Lemmas for the Schroeder-Bernstein Theorem +*) + +Perm = ZF + +consts + O :: "[i,i]=>i" (infixr 60) + id :: "i=>i" + inj,surj,bij:: "[i,i]=>i" + +rules + + (*composition of relations and functions; NOT Suppes's relative product*) + comp_def "r O s == {xz : domain(s)*range(r) . \ +\ EX x y z. xz= & :s & :r}" + + (*the identity function for A*) + id_def "id(A) == (lam x:A. x)" + + (*one-to-one functions from A to B*) + inj_def "inj(A,B) == { f: A->B. ALL w:A. ALL x:A. f`w=f`x --> w=x}" + + (*onto functions from A to B*) + surj_def "surj(A,B) == { f: A->B . ALL y:B. EX x:A. f`x=y}" + + (*one-to-one and onto functions*) + bij_def "bij(A,B) == inj(A,B) Int surj(A,B)" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/qpair.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/qpair.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,299 @@ +(* Title: ZF/qpair.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For qpair.thy. + +Quine-inspired ordered pairs and disjoint sums, for non-well-founded data +structures in ZF. Does not precisely follow Quine's construction. Thanks +to Thomas Forster for suggesting this approach! + +W. V. Quine, On Ordered Pairs and Relations, in Selected Logic Papers, +1966. + +Many proofs are borrowed from pair.ML and sum.ML + +Do we EVER have rank(a) < rank() ? Perhaps if the latter rank + is not a limit ordinal? +*) + + +open QPair; + +(**** Quine ordered pairing ****) + +(** Lemmas for showing that uniquely determines a and b **) + +val QPair_iff = prove_goalw QPair.thy [QPair_def] + " = <-> a=c & b=d" + (fn _=> [rtac sum_equal_iff 1]); + +val QPair_inject = standard (QPair_iff RS iffD1 RS conjE); + +val QPair_inject1 = prove_goal QPair.thy " = ==> a=c" + (fn [major]=> + [ (rtac (major RS QPair_inject) 1), (assume_tac 1) ]); + +val QPair_inject2 = prove_goal QPair.thy " = ==> b=d" + (fn [major]=> + [ (rtac (major RS QPair_inject) 1), (assume_tac 1) ]); + + +(*** QSigma: Disjoint union of a family of sets + Generalizes Cartesian product ***) + +val QSigmaI = prove_goalw QPair.thy [QSigma_def] + "[| a:A; b:B(a) |] ==> : QSigma(A,B)" + (fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]); + +(*The general elimination rule*) +val QSigmaE = prove_goalw QPair.thy [QSigma_def] + "[| c: QSigma(A,B); \ +\ !!x y.[| x:A; y:B(x); c= |] ==> P \ +\ |] ==> P" + (fn major::prems=> + [ (cut_facts_tac [major] 1), + (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]); + +(** Elimination rules for :A*B -- introducing no eigenvariables **) + +val QSigmaE2 = + rule_by_tactic (REPEAT_FIRST (etac QPair_inject ORELSE' bound_hyp_subst_tac) + THEN prune_params_tac) + (read_instantiate [("c","")] QSigmaE); + +val QSigmaD1 = prove_goal QPair.thy " : QSigma(A,B) ==> a : A" + (fn [major]=> + [ (rtac (major RS QSigmaE2) 1), (assume_tac 1) ]); + +val QSigmaD2 = prove_goal QPair.thy " : QSigma(A,B) ==> b : B(a)" + (fn [major]=> + [ (rtac (major RS QSigmaE2) 1), (assume_tac 1) ]); + +val QSigma_cong = prove_goalw QPair.thy [QSigma_def] + "[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> \ +\ QSigma(A,B) = QSigma(A',B')" + (fn prems=> [ (prove_cong_tac (prems@[RepFun_cong]) 1) ]); + +val QSigma_empty1 = prove_goal QPair.thy "QSigma(0,B) = 0" + (fn _ => [ (fast_tac (ZF_cs addIs [equalityI] addSEs [QSigmaE]) 1) ]); + +val QSigma_empty2 = prove_goal QPair.thy "A <*> 0 = 0" + (fn _ => [ (fast_tac (ZF_cs addIs [equalityI] addSEs [QSigmaE]) 1) ]); + + +(*** Eliminator - qsplit ***) + +val qsplit = prove_goalw QPair.thy [qsplit_def] + "qsplit(%x y.c(x,y), ) = c(a,b)" + (fn _ => [ (fast_tac (ZF_cs addIs [the_equality] addEs [QPair_inject]) 1) ]); + +val qsplit_type = prove_goal QPair.thy + "[| p:QSigma(A,B); \ +\ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C() \ +\ |] ==> qsplit(%x y.c(x,y), p) : C(p)" + (fn major::prems=> + [ (rtac (major RS QSigmaE) 1), + (etac ssubst 1), + (REPEAT (ares_tac (prems @ [qsplit RS ssubst]) 1)) ]); + +(*This congruence rule uses NO typing information...*) +val qsplit_cong = prove_goalw QPair.thy [qsplit_def] + "[| p=p'; !!x y.c(x,y) = c'(x,y) |] ==> \ +\ qsplit(%x y.c(x,y), p) = qsplit(%x y.c'(x,y), p')" + (fn prems=> [ (prove_cong_tac (prems@[the_cong]) 1) ]); + + +val qpair_cs = ZF_cs addSIs [QSigmaI] addSEs [QSigmaE2, QSigmaE, QPair_inject]; + +(*** qconverse ***) + +val qconverseI = prove_goalw QPair.thy [qconverse_def] + "!!a b r. :r ==> :qconverse(r)" + (fn _ => [ (fast_tac qpair_cs 1) ]); + +val qconverseD = prove_goalw QPair.thy [qconverse_def] + "!!a b r. : qconverse(r) ==> : r" + (fn _ => [ (fast_tac qpair_cs 1) ]); + +val qconverseE = prove_goalw QPair.thy [qconverse_def] + "[| yx : qconverse(r); \ +\ !!x y. [| yx=; :r |] ==> P \ +\ |] ==> P" + (fn [major,minor]=> + [ (rtac (major RS ReplaceE) 1), + (REPEAT (eresolve_tac [exE, conjE, minor] 1)), + (hyp_subst_tac 1), + (assume_tac 1) ]); + +val qconverse_cs = qpair_cs addSIs [qconverseI] + addSEs [qconverseD,qconverseE]; + +val qconverse_of_qconverse = prove_goal QPair.thy + "!!A B r. r<=QSigma(A,B) ==> qconverse(qconverse(r)) = r" + (fn _ => [ (fast_tac (qconverse_cs addSIs [equalityI]) 1) ]); + +val qconverse_type = prove_goal QPair.thy + "!!A B r. r <= A <*> B ==> qconverse(r) <= B <*> A" + (fn _ => [ (fast_tac qconverse_cs 1) ]); + +val qconverse_of_prod = prove_goal QPair.thy "qconverse(A <*> B) = B <*> A" + (fn _ => [ (fast_tac (qconverse_cs addSIs [equalityI]) 1) ]); + +val qconverse_empty = prove_goal QPair.thy "qconverse(0) = 0" + (fn _ => [ (fast_tac (qconverse_cs addSIs [equalityI]) 1) ]); + + +(*** qsplit for predicates: result type o ***) + +goalw QPair.thy [qfsplit_def] "!!R a b. R(a,b) ==> qfsplit(R, )"; +by (REPEAT (ares_tac [refl,exI,conjI] 1)); +val qfsplitI = result(); + +val major::prems = goalw QPair.thy [qfsplit_def] + "[| qfsplit(R,z); !!x y. [| z = ; R(x,y) |] ==> P |] ==> P"; +by (cut_facts_tac [major] 1); +by (REPEAT (eresolve_tac (prems@[asm_rl,exE,conjE]) 1)); +val qfsplitE = result(); + +goal QPair.thy "!!R a b. qfsplit(R,) ==> R(a,b)"; +by (REPEAT (eresolve_tac [asm_rl,qfsplitE,QPair_inject,ssubst] 1)); +val qfsplitD = result(); + + +(**** The Quine-inspired notion of disjoint sum ****) + +val qsum_defs = [qsum_def,QInl_def,QInr_def,qcase_def]; + +(** Introduction rules for the injections **) + +goalw QPair.thy qsum_defs "!!a A B. a : A ==> QInl(a) : A <+> B"; +by (REPEAT (ares_tac [UnI1,QSigmaI,singletonI] 1)); +val QInlI = result(); + +goalw QPair.thy qsum_defs "!!b A B. b : B ==> QInr(b) : A <+> B"; +by (REPEAT (ares_tac [UnI2,QSigmaI,singletonI] 1)); +val QInrI = result(); + +(** Elimination rules **) + +val major::prems = goalw QPair.thy qsum_defs + "[| u: A <+> B; \ +\ !!x. [| x:A; u=QInl(x) |] ==> P; \ +\ !!y. [| y:B; u=QInr(y) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS UnE) 1); +by (REPEAT (rtac refl 1 + ORELSE eresolve_tac (prems@[QSigmaE,singletonE,ssubst]) 1)); +val qsumE = result(); + +(** QInjection and freeness rules **) + +val [major] = goalw QPair.thy qsum_defs "QInl(a)=QInl(b) ==> a=b"; +by (EVERY1 [rtac (major RS QPair_inject), assume_tac]); +val QInl_inject = result(); + +val [major] = goalw QPair.thy qsum_defs "QInr(a)=QInr(b) ==> a=b"; +by (EVERY1 [rtac (major RS QPair_inject), assume_tac]); +val QInr_inject = result(); + +val [major] = goalw QPair.thy qsum_defs "QInl(a)=QInr(b) ==> P"; +by (rtac (major RS QPair_inject) 1); +by (etac (sym RS one_neq_0) 1); +val QInl_neq_QInr = result(); + +val QInr_neq_QInl = sym RS QInl_neq_QInr; + +(** Injection and freeness equivalences, for rewriting **) + +goal QPair.thy "QInl(a)=QInl(b) <-> a=b"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [QInl_inject,subst_context] 1)); +val QInl_iff = result(); + +goal QPair.thy "QInr(a)=QInr(b) <-> a=b"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [QInr_inject,subst_context] 1)); +val QInr_iff = result(); + +goal QPair.thy "QInl(a)=QInr(b) <-> False"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [QInl_neq_QInr,FalseE] 1)); +val QInl_QInr_iff = result(); + +goal QPair.thy "QInr(b)=QInl(a) <-> False"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [QInr_neq_QInl,FalseE] 1)); +val QInr_QInl_iff = result(); + +val qsum_cs = + ZF_cs addIs [QInlI,QInrI] addSEs [qsumE,QInl_neq_QInr,QInr_neq_QInl] + addSDs [QInl_inject,QInr_inject]; + +(** <+> is itself injective... who cares?? **) + +goal QPair.thy + "u: A <+> B <-> (EX x. x:A & u=QInl(x)) | (EX y. y:B & u=QInr(y))"; +by (fast_tac qsum_cs 1); +val qsum_iff = result(); + +goal QPair.thy "A <+> B <= C <+> D <-> A<=C & B<=D"; +by (fast_tac qsum_cs 1); +val qsum_subset_iff = result(); + +goal QPair.thy "A <+> B = C <+> D <-> A=C & B=D"; +by (SIMP_TAC (ZF_ss addrews [extension,qsum_subset_iff]) 1); +by (fast_tac ZF_cs 1); +val qsum_equal_iff = result(); + +(*** Eliminator -- qcase ***) + +goalw QPair.thy qsum_defs "qcase(c, d, QInl(a)) = c(a)"; +by (rtac (qsplit RS trans) 1); +by (rtac cond_0 1); +val qcase_QInl = result(); + +goalw QPair.thy qsum_defs "qcase(c, d, QInr(b)) = d(b)"; +by (rtac (qsplit RS trans) 1); +by (rtac cond_1 1); +val qcase_QInr = result(); + +val prems = goalw QPair.thy [qcase_def] + "[| u=u'; !!x. c(x)=c'(x); !!y. d(y)=d'(y) |] ==> \ +\ qcase(c,d,u)=qcase(c',d',u')"; +by (REPEAT (resolve_tac ([refl,qsplit_cong,cond_cong] @ prems) 1)); +val qcase_cong = result(); + +val major::prems = goal QPair.thy + "[| u: A <+> B; \ +\ !!x. x: A ==> c(x): C(QInl(x)); \ +\ !!y. y: B ==> d(y): C(QInr(y)) \ +\ |] ==> qcase(c,d,u) : C(u)"; +by (rtac (major RS qsumE) 1); +by (ALLGOALS (etac ssubst)); +by (ALLGOALS (ASM_SIMP_TAC (ZF_ss addrews + (prems@[qcase_QInl,qcase_QInr])))); +val qcase_type = result(); + +(** Rules for the Part primitive **) + +goal QPair.thy "Part(A <+> B,QInl) = {QInl(x). x: A}"; +by (fast_tac (qsum_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_QInl = result(); + +goal QPair.thy "Part(A <+> B,QInr) = {QInr(y). y: B}"; +by (fast_tac (qsum_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_QInr = result(); + +goal QPair.thy "Part(A <+> B, %x.QInr(h(x))) = {QInr(y). y: Part(B,h)}"; +by (fast_tac (qsum_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_QInr2 = result(); + +goal QPair.thy "!!A B C. C <= A <+> B ==> Part(C,QInl) Un Part(C,QInr) = C"; +by (rtac equalityI 1); +by (rtac Un_least 1); +by (rtac Part_subset 1); +by (rtac Part_subset 1); +by (fast_tac (ZF_cs addIs [PartI] addSEs [qsumE]) 1); +val Part_qsum_equality = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/qpair.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/qpair.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,52 @@ +(* Title: ZF/qpair.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Quine-inspired ordered pairs and disjoint sums, for non-well-founded data +structures in ZF. Does not precisely follow Quine's construction. Thanks +to Thomas Forster for suggesting this approach! + +W. V. Quine, On Ordered Pairs and Relations, in Selected Logic Papers, +1966. +*) + +QPair = Sum + +consts + QPair :: "[i, i] => i" ("<(_;/ _)>") + qsplit :: "[[i,i] => i, i] => i" + qfsplit :: "[[i,i] => o, i] => o" + qconverse :: "i => i" + "@QSUM" :: "[idt, i, i] => i" ("(3QSUM _:_./ _)" 10) + " <*>" :: "[i, i] => i" ("(_ <*>/ _)" [81, 80] 80) + QSigma :: "[i, i => i] => i" + + "<+>" :: "[i,i]=>i" (infixr 65) + QInl,QInr :: "i=>i" + qcase :: "[i=>i, i=>i, i]=>i" + +translations + "QSUM x:A. B" => "QSigma(A, %x. B)" + +rules + QPair_def " == a+b" + qsplit_def "qsplit(c,p) == THE y. EX a b. p= & y=c(a,b)" + qfsplit_def "qfsplit(R,z) == EX x y. z= & R(x,y)" + qconverse_def "qconverse(r) == {z. w:r, EX x y. w= & z=}" + QSigma_def "QSigma(A,B) == UN x:A. UN y:B(x). {}" + + qsum_def "A <+> B == QSigma({0}, %x.A) Un QSigma({1}, %x.B)" + QInl_def "QInl(a) == <0;a>" + QInr_def "QInr(b) == <1;b>" + qcase_def "qcase(c,d) == qsplit(%y z. cond(y, d(z), c(z)))" +end + +ML + +(* 'Dependent' type operators *) + +val parse_translation = + [(" <*>", ndependent_tr "QSigma")]; + +val print_translation = + [("QSigma", dependent_tr' ("@QSUM", " <*>"))]; diff -r 000000000000 -r a5a9c433f639 src/ZF/quniv.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/quniv.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,334 @@ +(* Title: ZF/quniv + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For quniv.thy. A small universe for lazy recursive types +*) + +open QUniv; + +(** Introduction and elimination rules avoid tiresome folding/unfolding **) + +goalw QUniv.thy [quniv_def] + "!!X A. X <= univ(eclose(A)) ==> X : quniv(A)"; +be PowI 1; +val qunivI = result(); + +goalw QUniv.thy [quniv_def] + "!!X A. X : quniv(A) ==> X <= univ(eclose(A))"; +be PowD 1; +val qunivD = result(); + +goalw QUniv.thy [quniv_def] "!!A B. A<=B ==> quniv(A) <= quniv(B)"; +by (etac (eclose_mono RS univ_mono RS Pow_mono) 1); +val quniv_mono = result(); + +(*** Closure properties ***) + +goalw QUniv.thy [quniv_def] "univ(eclose(A)) <= quniv(A)"; +by (rtac (Transset_iff_Pow RS iffD1) 1); +by (rtac (Transset_eclose RS Transset_univ) 1); +val univ_eclose_subset_quniv = result(); + +goal QUniv.thy "univ(A) <= quniv(A)"; +by (rtac (arg_subset_eclose RS univ_mono RS subset_trans) 1); +by (rtac univ_eclose_subset_quniv 1); +val univ_subset_quniv = result(); + +val univ_into_quniv = standard (univ_subset_quniv RS subsetD); + +goalw QUniv.thy [quniv_def] "Pow(univ(A)) <= quniv(A)"; +by (rtac (arg_subset_eclose RS univ_mono RS Pow_mono) 1); +val Pow_univ_subset_quniv = result(); + +val univ_subset_into_quniv = standard + (PowI RS (Pow_univ_subset_quniv RS subsetD)); + +val zero_in_quniv = standard (zero_in_univ RS univ_into_quniv); +val one_in_quniv = standard (one_in_univ RS univ_into_quniv); +val two_in_quniv = standard (two_in_univ RS univ_into_quniv); + +val A_subset_quniv = standard + ([A_subset_univ, univ_subset_quniv] MRS subset_trans); + +val A_into_quniv = A_subset_quniv RS subsetD; + +(*** univ(A) closure for Quine-inspired pairs and injections ***) + +(*Quine ordered pairs*) +goalw QUniv.thy [QPair_def] + "!!A a. [| a <= univ(A); b <= univ(A) |] ==> <= univ(A)"; +by (REPEAT (ares_tac [sum_subset_univ] 1)); +val QPair_subset_univ = result(); + +(** Quine disjoint sum **) + +goalw QUniv.thy [QInl_def] "!!A a. a <= univ(A) ==> QInl(a) <= univ(A)"; +by (etac (empty_subsetI RS QPair_subset_univ) 1); +val QInl_subset_univ = result(); + +val naturals_subset_nat = + rewrite_rule [Transset_def] (Ord_nat RS Ord_is_Transset) + RS bspec; + +val naturals_subset_univ = + [naturals_subset_nat, nat_subset_univ] MRS subset_trans; + +goalw QUniv.thy [QInr_def] "!!A a. a <= univ(A) ==> QInr(a) <= univ(A)"; +by (etac (nat_1I RS naturals_subset_univ RS QPair_subset_univ) 1); +val QInr_subset_univ = result(); + +(*** Closure for Quine-inspired products and sums ***) + +(*Quine ordered pairs*) +goalw QUniv.thy [quniv_def,QPair_def] + "!!A a. [| a: quniv(A); b: quniv(A) |] ==> : quniv(A)"; +by (REPEAT (dtac PowD 1)); +by (REPEAT (ares_tac [PowI, sum_subset_univ] 1)); +val QPair_in_quniv = result(); + +goal QUniv.thy "quniv(A) <*> quniv(A) <= quniv(A)"; +by (REPEAT (ares_tac [subsetI, QPair_in_quniv] 1 + ORELSE eresolve_tac [QSigmaE, ssubst] 1)); +val QSigma_quniv = result(); + +val QSigma_subset_quniv = standard + (QSigma_mono RS (QSigma_quniv RSN (2,subset_trans))); + +(*The opposite inclusion*) +goalw QUniv.thy [quniv_def,QPair_def] + "!!A a b. : quniv(A) ==> a: quniv(A) & b: quniv(A)"; +be ([Transset_eclose RS Transset_univ, PowD] MRS + Transset_includes_summands RS conjE) 1; +by (REPEAT (ares_tac [conjI,PowI] 1)); +val quniv_QPair_D = result(); + +val quniv_QPair_E = standard (quniv_QPair_D RS conjE); + +goal QUniv.thy " : quniv(A) <-> a: quniv(A) & b: quniv(A)"; +by (REPEAT (ares_tac [iffI, QPair_in_quniv, quniv_QPair_D] 1 + ORELSE etac conjE 1)); +val quniv_QPair_iff = result(); + + +(** Quine disjoint sum **) + +goalw QUniv.thy [QInl_def] "!!A a. a: quniv(A) ==> QInl(a) : quniv(A)"; +by (REPEAT (ares_tac [zero_in_quniv,QPair_in_quniv] 1)); +val QInl_in_quniv = result(); + +goalw QUniv.thy [QInr_def] "!!A b. b: quniv(A) ==> QInr(b) : quniv(A)"; +by (REPEAT (ares_tac [one_in_quniv, QPair_in_quniv] 1)); +val QInr_in_quniv = result(); + +goal QUniv.thy "quniv(C) <+> quniv(C) <= quniv(C)"; +by (REPEAT (ares_tac [subsetI, QInl_in_quniv, QInr_in_quniv] 1 + ORELSE eresolve_tac [qsumE, ssubst] 1)); +val qsum_quniv = result(); + +val qsum_subset_quniv = standard + (qsum_mono RS (qsum_quniv RSN (2,subset_trans))); + +(*** The natural numbers ***) + +val nat_subset_quniv = standard + ([nat_subset_univ, univ_subset_quniv] MRS subset_trans); + +(* n:nat ==> n:quniv(A) *) +val nat_into_quniv = standard (nat_subset_quniv RS subsetD); + +val bool_subset_quniv = standard + ([bool_subset_univ, univ_subset_quniv] MRS subset_trans); + +val bool_into_quniv = standard (bool_subset_quniv RS subsetD); + + +(**** Properties of Vfrom analogous to the "take-lemma" ****) + +(*** Intersecting a*b with Vfrom... ***) + +(*This version says a, b exist one level down, in the smaller set Vfrom(X,i)*) +goal Univ.thy + "!!X. [| {a,b} : Vfrom(X,succ(i)); Transset(X) |] ==> \ +\ a: Vfrom(X,i) & b: Vfrom(X,i)"; +bd (Transset_Vfrom_succ RS equalityD1 RS subsetD RS PowD) 1; +ba 1; +by (fast_tac ZF_cs 1); +val doubleton_in_Vfrom_D = result(); + +(*This weaker version says a, b exist at the same level*) +val Vfrom_doubleton_D = standard (Transset_Vfrom RS Transset_doubleton_D); + +(** Using only the weaker theorem would prove : Vfrom(X,i) + implies a, b : Vfrom(X,i), which is useless for induction. + Using only the stronger theorem would prove : Vfrom(X,succ(succ(i))) + implies a, b : Vfrom(X,i), leaving the succ(i) case untreated. + The combination gives a reduction by precisely one level, which is + most convenient for proofs. +**) + +goalw Univ.thy [Pair_def] + "!!X. [| : Vfrom(X,succ(i)); Transset(X) |] ==> \ +\ a: Vfrom(X,i) & b: Vfrom(X,i)"; +by (fast_tac (ZF_cs addSDs [doubleton_in_Vfrom_D, Vfrom_doubleton_D]) 1); +val Pair_in_Vfrom_D = result(); + +goal Univ.thy + "!!X. Transset(X) ==> \ +\ (a*b) Int Vfrom(X, succ(i)) <= (a Int Vfrom(X,i)) * (b Int Vfrom(X,i))"; +by (fast_tac (ZF_cs addSDs [Pair_in_Vfrom_D]) 1); +val product_Int_Vfrom_subset = result(); + +(*** Intersecting with Vfrom... ***) + +goalw QUniv.thy [QPair_def,sum_def] + "!!X. Transset(X) ==> \ +\ Int Vfrom(X, succ(i)) <= "; +br (Int_Un_distrib RS ssubst) 1; +br Un_mono 1; +by (REPEAT (ares_tac [product_Int_Vfrom_subset RS subset_trans, + [Int_lower1, subset_refl] MRS Sigma_mono] 1)); +val QPair_Int_Vfrom_succ_subset = result(); + +(** Pairs in quniv -- for handling the base case **) + +goal QUniv.thy "!!X. : quniv(X) ==> : univ(eclose(X))"; +be ([qunivD, Transset_eclose] MRS Transset_Pair_subset_univ) 1; +val Pair_in_quniv_D = result(); + +goal QUniv.thy "a*b Int quniv(A) = a*b Int univ(eclose(A))"; +br equalityI 1; +br ([subset_refl, univ_eclose_subset_quniv] MRS Int_mono) 2; +by (fast_tac (ZF_cs addSEs [Pair_in_quniv_D]) 1); +val product_Int_quniv_eq = result(); + +goalw QUniv.thy [QPair_def,sum_def] + " Int quniv(A) = Int univ(eclose(A))"; +by (SIMP_TAC (ZF_ss addrews [Int_Un_distrib, product_Int_quniv_eq]) 1); +val QPair_Int_quniv_eq = result(); + +(**** "Take-lemma" rules for proving c: quniv(A) ****) + +goalw QUniv.thy [quniv_def] "Transset(quniv(A))"; +br (Transset_eclose RS Transset_univ RS Transset_Pow) 1; +val Transset_quniv = result(); + +val [aprem, iprem] = goal QUniv.thy + "[| a: quniv(quniv(X)); \ +\ !!i. i:nat ==> a Int Vfrom(quniv(X),i) : quniv(A) \ +\ |] ==> a : quniv(A)"; +br (univ_Int_Vfrom_subset RS qunivI) 1; +br (aprem RS qunivD) 1; +by (rtac (Transset_quniv RS Transset_eclose_eq_arg RS ssubst) 1); +be (iprem RS qunivD) 1; +val quniv_Int_Vfrom = result(); + +(** Rules for level 0 **) + +goal QUniv.thy " Int quniv(A) : quniv(A)"; +br (QPair_Int_quniv_eq RS ssubst) 1; +br (Int_lower2 RS qunivI) 1; +val QPair_Int_quniv_in_quniv = result(); + +(*Unused; kept as an example. QInr rule is similar*) +goalw QUniv.thy [QInl_def] "QInl(a) Int quniv(A) : quniv(A)"; +br QPair_Int_quniv_in_quniv 1; +val QInl_Int_quniv_in_quniv = result(); + +goal QUniv.thy "!!a A X. a : quniv(A) ==> a Int X : quniv(A)"; +be ([Int_lower1, qunivD] MRS subset_trans RS qunivI) 1; +val Int_quniv_in_quniv = result(); + +goal QUniv.thy + "!!X. a Int X : quniv(A) ==> a Int Vfrom(X, 0) : quniv(A)"; +by (etac (Vfrom_0 RS ssubst) 1); +val Int_Vfrom_0_in_quniv = result(); + +(** Rules for level succ(i), decreasing it to i **) + +goal QUniv.thy + "!!X. [| a Int Vfrom(X,i) : quniv(A); \ +\ b Int Vfrom(X,i) : quniv(A); \ +\ Transset(X) \ +\ |] ==> Int Vfrom(X, succ(i)) : quniv(A)"; +br (QPair_Int_Vfrom_succ_subset RS subset_trans RS qunivI) 1; +br (QPair_in_quniv RS qunivD) 2; +by (REPEAT (assume_tac 1)); +val QPair_Int_Vfrom_succ_in_quniv = result(); + +val zero_Int_in_quniv = standard + ([Int_lower1,empty_subsetI] MRS subset_trans RS qunivI); + +val one_Int_in_quniv = standard + ([Int_lower1, one_in_quniv RS qunivD] MRS subset_trans RS qunivI); + +(*Unused; kept as an example. QInr rule is similar*) +goalw QUniv.thy [QInl_def] + "!!X. [| a Int Vfrom(X,i) : quniv(A); Transset(X) \ +\ |] ==> QInl(a) Int Vfrom(X, succ(i)) : quniv(A)"; +br QPair_Int_Vfrom_succ_in_quniv 1; +by (REPEAT (ares_tac [zero_Int_in_quniv] 1)); +val QInl_Int_Vfrom_succ_in_quniv = result(); + +(** Rules for level i -- preserving the level, not decreasing it **) + +goalw QUniv.thy [QPair_def] + "!!X. Transset(X) ==> \ +\ Int Vfrom(X,i) <= "; +be (Transset_Vfrom RS Transset_sum_Int_subset) 1; +val QPair_Int_Vfrom_subset = result(); + +goal QUniv.thy + "!!X. [| a Int Vfrom(X,i) : quniv(A); \ +\ b Int Vfrom(X,i) : quniv(A); \ +\ Transset(X) \ +\ |] ==> Int Vfrom(X,i) : quniv(A)"; +br (QPair_Int_Vfrom_subset RS subset_trans RS qunivI) 1; +br (QPair_in_quniv RS qunivD) 2; +by (REPEAT (assume_tac 1)); +val QPair_Int_Vfrom_in_quniv = result(); + + +(**** "Take-lemma" rules for proving a=b by co-induction ****) + +(** Unfortunately, the technique used above does not apply here, since + the base case appears impossible to prove: it involves an intersection + with eclose(X) for arbitrary X. So a=b is proved by transfinite induction + over ALL ordinals, using Vset(i) instead of Vfrom(X,i). +**) + +(*Rule for level 0*) +goal QUniv.thy "a Int Vset(0) <= b"; +by (rtac (Vfrom_0 RS ssubst) 1); +by (fast_tac ZF_cs 1); +val Int_Vset_0_subset = result(); + +(*Rule for level succ(i), decreasing it to i*) +goal QUniv.thy + "!!i. [| a Int Vset(i) <= c; \ +\ b Int Vset(i) <= d \ +\ |] ==> Int Vset(succ(i)) <= "; +br ([Transset_0 RS QPair_Int_Vfrom_succ_subset, QPair_mono] + MRS subset_trans) 1; +by (REPEAT (assume_tac 1)); +val QPair_Int_Vset_succ_subset_trans = result(); + +(*Unused; kept as an example. QInr rule is similar*) +goalw QUniv.thy [QInl_def] + "!!i. a Int Vset(i) <= b ==> QInl(a) Int Vset(succ(i)) <= QInl(b)"; +be (Int_lower1 RS QPair_Int_Vset_succ_subset_trans) 1; +val QInl_Int_Vset_succ_subset_trans = result(); + +(*Rule for level i -- preserving the level, not decreasing it*) +goal QUniv.thy + "!!i. [| a Int Vset(i) <= c; \ +\ b Int Vset(i) <= d \ +\ |] ==> Int Vset(i) <= "; +br ([Transset_0 RS QPair_Int_Vfrom_subset, QPair_mono] + MRS subset_trans) 1; +by (REPEAT (assume_tac 1)); +val QPair_Int_Vset_subset_trans = result(); + + + diff -r 000000000000 -r a5a9c433f639 src/ZF/quniv.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/quniv.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,16 @@ +(* Title: ZF/univ.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +A small universe for lazy recursive types +*) + +QUniv = Univ + QPair + +consts + quniv :: "i=>i" + +rules + quniv_def "quniv(A) == Pow(univ(eclose(A)))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/simpdata.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/simpdata.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,120 @@ +(* Title: ZF/simpdata + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Rewriting for ZF set theory -- based on FOL rewriting +*) + +fun prove_fun s = + (writeln s; prove_goal ZF.thy s + (fn prems => [ (cut_facts_tac prems 1), (fast_tac ZF_cs 1) ])); + +val mem_rews = map prove_fun + [ "a:0 <-> False", + "a : A Un B <-> a:A | a:B", + "a : A Int B <-> a:A & a:B", + "a : A-B <-> a:A & ~a:B", + "a : cons(b,B) <-> a=b | a:B", + "i : succ(j) <-> i=j | i:j", + ": Sigma(A,B) <-> a:A & b:B(a)", + "a : Collect(A,P) <-> a:A & P(a)" ]; + +(** Tactics for type checking -- from CTT **) + +fun is_rigid_elem (Const("Trueprop",_) $ (Const("op :",_) $ a $ _)) = + not (is_Var (head_of a)) + | is_rigid_elem _ = false; + +(*Try solving a:A by assumption provided a is rigid!*) +val test_assume_tac = SUBGOAL(fn (prem,i) => + if is_rigid_elem (Logic.strip_assums_concl prem) + then assume_tac i else no_tac); + +(*Type checking solves a:?A (a rigid, ?A maybe flexible). + match_tac is too strict; would refuse to instantiate ?A*) +fun typechk_step_tac tyrls = + FIRSTGOAL (test_assume_tac ORELSE' filt_resolve_tac tyrls 3); + +fun typechk_tac tyrls = REPEAT (typechk_step_tac tyrls); + +val ZF_typechecks = [if_type,lam_type,SigmaI,apply_type,split_type]; + +(*To instantiate variables in typing conditions; + to perform type checking faster than rewriting can + NOT TERRIBLY USEFUL because it does not simplify conjunctions*) +fun type_auto_tac tyrls hyps = SELECT_GOAL + (DEPTH_SOLVE (typechk_step_tac (tyrls@hyps) + ORELSE ares_tac [TrueI,ballI,allI,conjI,impI] 1)); + +(** New version of mk_rew_rules **) + +(*Should False yield False<->True, or should it solve goals some other way?*) + +(*Analyse a rigid formula*) +val atomize_pairs = + [("Ball", [bspec]), + ("All", [spec]), + ("op -->", [mp]), + ("op &", [conjunct1,conjunct2])]; + +(*Analyse a:b, where b is rigid*) +val atomize_mem_pairs = + [("Collect", [CollectD1,CollectD2]), + ("op -", [DiffD1,DiffD2]), + ("op Int", [IntD1,IntD2])]; + +(*Analyse a theorem to atomic rewrite rules*) +fun atomize th = + let fun tryrules pairs t = + case head_of t of + Const(a,_) => + (case assoc(pairs,a) of + Some rls => flat (map atomize ([th] RL rls)) + | None => [th]) + | _ => [th] + in case concl_of th of (*The operator below is Trueprop*) + _ $ (Const("op :",_) $ a $ b) => tryrules atomize_mem_pairs b + | _ $ (Const("True",_)) => [] (*True is DELETED*) + | _ $ (Const("False",_)) => [] (*should False do something??*) + | _ $ A => tryrules atomize_pairs A + end; + +fun ZF_mk_rew_rules th = map mk_eq (atomize th); + + +fun auto_tac rls hyps = SELECT_GOAL (DEPTH_SOLVE_1 (ares_tac (rls@hyps) 1)); + +structure ZF_SimpData : SIMP_DATA = + struct + val refl_thms = FOL_SimpData.refl_thms + val trans_thms = FOL_SimpData.trans_thms + val red1 = FOL_SimpData.red1 + val red2 = FOL_SimpData.red2 + val mk_rew_rules = ZF_mk_rew_rules + val norm_thms = FOL_SimpData.norm_thms + val subst_thms = FOL_SimpData.subst_thms + val dest_red = FOL_SimpData.dest_red + end; + +structure ZF_Simp = SimpFun(ZF_SimpData); + +open ZF_Simp; + +(*Redeclared because the previous FOL_ss belongs to a different instance + of type simpset*) +val FOL_ss = empty_ss addcongs FOL_congs addrews FOL_rews + setauto auto_tac[TrueI,ballI]; + +(** Basic congruence and rewrite rules for ZF set theory **) + +val ZF_congs = + [ball_cong,bex_cong,Replace_cong,RepFun_cong,Collect_cong,the_cong, + if_cong,Sigma_cong,split_cong,Pi_cong,lam_cong] @ basic_ZF_congs; + +val ZF_rews = [empty_subsetI, ball_rew, if_true, if_false, + beta, eta, restrict, + fst_conv, snd_conv, split]; + +val ZF_ss = FOL_ss addcongs ZF_congs addrews (ZF_rews@mem_rews); + diff -r 000000000000 -r a5a9c433f639 src/ZF/subset.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/subset.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,191 @@ +(* Title: ZF/subset + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Derived rules involving subsets +Union and Intersection as lattice operations +*) + +(*** cons ***) + +val cons_subsetI = prove_goal ZF.thy "[| a:C; B<=C |] ==> cons(a,B) <= C" + (fn prems=> + [ (cut_facts_tac prems 1), + (REPEAT (ares_tac [subsetI] 1 + ORELSE eresolve_tac [consE,ssubst,subsetD] 1)) ]); + +val subset_consI = prove_goal ZF.thy "B <= cons(a,B)" + (fn _=> [ (rtac subsetI 1), (etac consI2 1) ]); + +(*Useful for rewriting!*) +val cons_subset_iff = prove_goal ZF.thy "cons(a,B)<=C <-> a:C & B<=C" + (fn _=> [ (fast_tac upair_cs 1) ]); + +(*A safe special case of subset elimination, adding no new variables + [| cons(a,B) <= C; [| a : C; B <= C |] ==> R |] ==> R *) +val cons_subsetE = standard (cons_subset_iff RS iffD1 RS conjE); + +val subset_empty_iff = prove_goal ZF.thy "A<=0 <-> A=0" + (fn _=> [ (fast_tac (upair_cs addIs [equalityI]) 1) ]); + +val subset_cons_iff = prove_goal ZF.thy + "C<=cons(a,B) <-> C<=B | (a:C & C-{a} <= B)" + (fn _=> [ (fast_tac upair_cs 1) ]); + +(*** succ ***) + +val subset_succI = prove_goal ZF.thy "i <= succ(i)" + (fn _=> [ (rtac subsetI 1), (etac succI2 1) ]); + +(*But if j is an ordinal or is transitive, then i:j implies i<=j! + See ordinal/Ord_succ_subsetI*) +val succ_subsetI = prove_goalw ZF.thy [succ_def] + "[| i:j; i<=j |] ==> succ(i)<=j" + (fn prems=> + [ (REPEAT (ares_tac (prems@[cons_subsetI]) 1)) ]); + +val succ_subsetE = prove_goalw ZF.thy [succ_def] + "[| succ(i) <= j; [| i:j; i<=j |] ==> P \ +\ |] ==> P" + (fn major::prems=> + [ (rtac (major RS cons_subsetE) 1), + (REPEAT (ares_tac prems 1)) ]); + +(*** singletons ***) + +val singleton_subsetI = prove_goal ZF.thy + "a:C ==> {a} <= C" + (fn prems=> + [ (REPEAT (resolve_tac (prems@[cons_subsetI,empty_subsetI]) 1)) ]); + +val singleton_subsetD = prove_goal ZF.thy + "{a} <= C ==> a:C" + (fn prems=> [ (REPEAT (ares_tac (prems@[cons_subsetE]) 1)) ]); + +(*** Big Union -- least upper bound of a set ***) + +val Union_subset_iff = prove_goal ZF.thy "Union(A) <= C <-> (ALL x:A. x <= C)" + (fn _ => [ fast_tac upair_cs 1 ]); + +val Union_upper = prove_goal ZF.thy + "B:A ==> B <= Union(A)" + (fn prems=> [ (REPEAT (ares_tac (prems@[subsetI,UnionI]) 1)) ]); + +val Union_least = prove_goal ZF.thy + "[| !!x. x:A ==> x<=C |] ==> Union(A) <= C" + (fn [prem]=> + [ (rtac (ballI RS (Union_subset_iff RS iffD2)) 1), + (etac prem 1) ]); + +(*** Union of a family of sets ***) + +goal ZF.thy "A <= (UN i:I. B(i)) <-> A = (UN i:I. A Int B(i))"; +by (fast_tac (upair_cs addSIs [equalityI] addSEs [equalityE]) 1); +val subset_UN_iff_eq = result(); + +val UN_subset_iff = prove_goal ZF.thy + "(UN x:A.B(x)) <= C <-> (ALL x:A. B(x) <= C)" + (fn _ => [ fast_tac upair_cs 1 ]); + +val UN_upper = prove_goal ZF.thy + "!!x A. x:A ==> B(x) <= (UN x:A.B(x))" + (fn _ => [ etac (RepFunI RS Union_upper) 1 ]); + +val UN_least = prove_goal ZF.thy + "[| !!x. x:A ==> B(x)<=C |] ==> (UN x:A.B(x)) <= C" + (fn [prem]=> + [ (rtac (ballI RS (UN_subset_iff RS iffD2)) 1), + (etac prem 1) ]); + + +(*** Big Intersection -- greatest lower bound of a nonempty set ***) + +val Inter_subset_iff = prove_goal ZF.thy + "!!a A. a: A ==> C <= Inter(A) <-> (ALL x:A. C <= x)" + (fn _ => [ fast_tac upair_cs 1 ]); + +val Inter_lower = prove_goal ZF.thy "B:A ==> Inter(A) <= B" + (fn prems=> + [ (REPEAT (resolve_tac (prems@[subsetI]) 1 + ORELSE etac InterD 1)) ]); + +val Inter_greatest = prove_goal ZF.thy + "[| a:A; !!x. x:A ==> C<=x |] ==> C <= Inter(A)" + (fn [prem1,prem2]=> + [ (rtac ([prem1, ballI] MRS (Inter_subset_iff RS iffD2)) 1), + (etac prem2 1) ]); + +(*** Intersection of a family of sets ***) + +val INT_lower = prove_goal ZF.thy + "x:A ==> (INT x:A.B(x)) <= B(x)" + (fn [prem] => + [ rtac (prem RS RepFunI RS Inter_lower) 1 ]); + +val INT_greatest = prove_goal ZF.thy + "[| a:A; !!x. x:A ==> C<=B(x) |] ==> C <= (INT x:A.B(x))" + (fn [nonempty,prem] => + [ rtac (nonempty RS RepFunI RS Inter_greatest) 1, + REPEAT (eresolve_tac [RepFunE, prem, ssubst] 1) ]); + + +(*** Finite Union -- the least upper bound of 2 sets ***) + +val Un_subset_iff = prove_goal ZF.thy "A Un B <= C <-> A <= C & B <= C" + (fn _ => [ fast_tac upair_cs 1 ]); + +val Un_upper1 = prove_goal ZF.thy "A <= A Un B" + (fn _ => [ (REPEAT (ares_tac [subsetI,UnI1] 1)) ]); + +val Un_upper2 = prove_goal ZF.thy "B <= A Un B" + (fn _ => [ (REPEAT (ares_tac [subsetI,UnI2] 1)) ]); + +val Un_least = prove_goal ZF.thy "!!A B C. [| A<=C; B<=C |] ==> A Un B <= C" + (fn _ => + [ (rtac (Un_subset_iff RS iffD2) 1), + (REPEAT (ares_tac [conjI] 1)) ]); + +(*** Finite Intersection -- the greatest lower bound of 2 sets *) + +val Int_subset_iff = prove_goal ZF.thy "C <= A Int B <-> C <= A & C <= B" + (fn _ => [ fast_tac upair_cs 1 ]); + +val Int_lower1 = prove_goal ZF.thy "A Int B <= A" + (fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)) ]); + +val Int_lower2 = prove_goal ZF.thy "A Int B <= B" + (fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)) ]); + +val Int_greatest = prove_goal ZF.thy + "!!A B C. [| C<=A; C<=B |] ==> C <= A Int B" + (fn prems=> + [ (rtac (Int_subset_iff RS iffD2) 1), + (REPEAT (ares_tac [conjI] 1)) ]); + +(*** Set difference *) + +val Diff_subset = prove_goal ZF.thy "A-B <= A" + (fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac DiffE 1)) ]); + +val Diff_contains = prove_goal ZF.thy + "[| C<=A; C Int B = 0 |] ==> C <= A-B" + (fn prems=> + [ (cut_facts_tac prems 1), + (rtac subsetI 1), + (REPEAT (ares_tac [DiffI,IntI,notI] 1 + ORELSE eresolve_tac [subsetD,equals0D] 1)) ]); + +(** Collect **) + +val Collect_subset = prove_goal ZF.thy "Collect(A,P) <= A" + (fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac CollectD1 1)) ]); + +(** RepFun **) + +val prems = goal ZF.thy "[| !!x. x:A ==> f(x): B |] ==> {f(x). x:A} <= B"; +by (rtac subsetI 1); +by (etac RepFunE 1); +by (etac ssubst 1); +by (eresolve_tac prems 1); +val RepFun_subset = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/sum.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/sum.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,170 @@ +(* Title: ZF/sum + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Disjoint sums in Zermelo-Fraenkel Set Theory +*) + +open Sum; + +val sum_defs = [sum_def,Inl_def,Inr_def,case_def]; + +(** Introduction rules for the injections **) + +goalw Sum.thy sum_defs "!!a A B. a : A ==> Inl(a) : A+B"; +by (REPEAT (ares_tac [UnI1,SigmaI,singletonI] 1)); +val InlI = result(); + +goalw Sum.thy sum_defs "!!b A B. b : B ==> Inr(b) : A+B"; +by (REPEAT (ares_tac [UnI2,SigmaI,singletonI] 1)); +val InrI = result(); + +(** Elimination rules **) + +val major::prems = goalw Sum.thy sum_defs + "[| u: A+B; \ +\ !!x. [| x:A; u=Inl(x) |] ==> P; \ +\ !!y. [| y:B; u=Inr(y) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS UnE) 1); +by (REPEAT (rtac refl 1 + ORELSE eresolve_tac (prems@[SigmaE,singletonE,ssubst]) 1)); +val sumE = result(); + +(** Injection and freeness rules **) + +val [major] = goalw Sum.thy sum_defs "Inl(a)=Inl(b) ==> a=b"; +by (EVERY1 [rtac (major RS Pair_inject), assume_tac]); +val Inl_inject = result(); + +val [major] = goalw Sum.thy sum_defs "Inr(a)=Inr(b) ==> a=b"; +by (EVERY1 [rtac (major RS Pair_inject), assume_tac]); +val Inr_inject = result(); + +val [major] = goalw Sum.thy sum_defs "Inl(a)=Inr(b) ==> P"; +by (rtac (major RS Pair_inject) 1); +by (etac (sym RS one_neq_0) 1); +val Inl_neq_Inr = result(); + +val Inr_neq_Inl = sym RS Inl_neq_Inr; + +(** Injection and freeness equivalences, for rewriting **) + +goal Sum.thy "Inl(a)=Inl(b) <-> a=b"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [Inl_inject,subst_context] 1)); +val Inl_iff = result(); + +goal Sum.thy "Inr(a)=Inr(b) <-> a=b"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [Inr_inject,subst_context] 1)); +val Inr_iff = result(); + +goal Sum.thy "Inl(a)=Inr(b) <-> False"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [Inl_neq_Inr,FalseE] 1)); +val Inl_Inr_iff = result(); + +goal Sum.thy "Inr(b)=Inl(a) <-> False"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [Inr_neq_Inl,FalseE] 1)); +val Inr_Inl_iff = result(); + +val sum_cs = ZF_cs addIs [InlI,InrI] addSEs [sumE,Inl_neq_Inr,Inr_neq_Inl] + addSDs [Inl_inject,Inr_inject]; + +goal Sum.thy "u: A+B <-> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))"; +by (fast_tac sum_cs 1); +val sum_iff = result(); + +goal Sum.thy "A+B <= C+D <-> A<=C & B<=D"; +by (fast_tac sum_cs 1); +val sum_subset_iff = result(); + +goal Sum.thy "A+B = C+D <-> A=C & B=D"; +by (SIMP_TAC (ZF_ss addrews [extension,sum_subset_iff]) 1); +by (fast_tac ZF_cs 1); +val sum_equal_iff = result(); + +(*** Eliminator -- case ***) + +goalw Sum.thy sum_defs "case(c, d, Inl(a)) = c(a)"; +by (rtac (split RS trans) 1); +by (rtac cond_0 1); +val case_Inl = result(); + +goalw Sum.thy sum_defs "case(c, d, Inr(b)) = d(b)"; +by (rtac (split RS trans) 1); +by (rtac cond_1 1); +val case_Inr = result(); + +val prems = goalw Sum.thy [case_def] + "[| u=u'; !!x. c(x)=c'(x); !!y. d(y)=d'(y) |] ==> \ +\ case(c,d,u)=case(c',d',u')"; +by (REPEAT (resolve_tac ([refl,split_cong,cond_cong] @ prems) 1)); +val case_cong = result(); + +val major::prems = goal Sum.thy + "[| u: A+B; \ +\ !!x. x: A ==> c(x): C(Inl(x)); \ +\ !!y. y: B ==> d(y): C(Inr(y)) \ +\ |] ==> case(c,d,u) : C(u)"; +by (rtac (major RS sumE) 1); +by (ALLGOALS (etac ssubst)); +by (ALLGOALS (ASM_SIMP_TAC (ZF_ss addrews + (prems@[case_Inl,case_Inr])))); +val case_type = result(); + +(** Rules for the Part primitive **) + +goalw Sum.thy [Part_def] + "!!a b A h. [| a : A; a=h(b) |] ==> a : Part(A,h)"; +by (REPEAT (ares_tac [exI,CollectI] 1)); +val Part_eqI = result(); + +val PartI = refl RSN (2,Part_eqI); + +val major::prems = goalw Sum.thy [Part_def] + "[| a : Part(A,h); !!z. [| a : A; a=h(z) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS CollectE) 1); +by (etac exE 1); +by (REPEAT (ares_tac prems 1)); +val PartE = result(); + +goalw Sum.thy [Part_def] "Part(A,h) <= A"; +by (rtac Collect_subset 1); +val Part_subset = result(); + +goal Sum.thy "!!A B h. A<=B ==> Part(A,h)<=Part(B,h)"; +by (fast_tac (ZF_cs addIs [PartI] addSEs [PartE]) 1); +val Part_mono = result(); + +goal Sum.thy "Part(A+B,Inl) = {Inl(x). x: A}"; +by (fast_tac (sum_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_Inl = result(); + +goal Sum.thy "Part(A+B,Inr) = {Inr(y). y: B}"; +by (fast_tac (sum_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_Inr = result(); + +goalw Sum.thy [Part_def] "!!a. a : Part(A,h) ==> a : A"; +by (etac CollectD1 1); +val PartD1 = result(); + +goal Sum.thy "Part(A,%x.x) = A"; +by (fast_tac (ZF_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_id = result(); + +goal Sum.thy "Part(A+B, %x.Inr(h(x))) = {Inr(y). y: Part(B,h)}"; +by (fast_tac (sum_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +val Part_Inr2 = result(); + +goal Sum.thy "!!A B C. C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C"; +by (rtac equalityI 1); +by (rtac Un_least 1); +by (rtac Part_subset 1); +by (rtac Part_subset 1); +by (fast_tac (ZF_cs addIs [PartI] addSEs [sumE]) 1); +val Part_sum_equality = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/sum.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/sum.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,25 @@ +(* Title: ZF/sum.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Disjoint sums in Zermelo-Fraenkel Set Theory +"Part" primitive for simultaneous recursive type definitions +*) + +Sum = Bool + +consts + "+" :: "[i,i]=>i" (infixr 65) + Inl,Inr :: "i=>i" + case :: "[i=>i, i=>i, i]=>i" + Part :: "[i,i=>i] => i" + +rules + sum_def "A+B == {0}*A Un {1}*B" + Inl_def "Inl(a) == <0,a>" + Inr_def "Inr(b) == <1,b>" + case_def "case(c,d) == split(%y z. cond(y, d(z), c(z)))" + + (*operator for selecting out the various summands*) + Part_def "Part(A,h) == {x: A. EX z. x = h(z)}" +end diff -r 000000000000 -r a5a9c433f639 src/ZF/trancl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/trancl.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,193 @@ +(* Title: ZF/trancl.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For trancl.thy. Transitive closure of a relation +*) + +open Trancl; + +val major::prems = goalw Trancl.thy [trans_def] + "[| trans(r); :r; :r |] ==> :r"; +by (rtac (major RS spec RS spec RS spec RS mp RS mp) 1); +by (REPEAT (resolve_tac prems 1)); +val transD = result(); + +goal Trancl.thy "bnd_mono(field(r)*field(r), %s. id(field(r)) Un (r O s))"; +by (rtac bnd_monoI 1); +by (REPEAT (ares_tac [subset_refl, Un_mono, comp_mono] 2)); +by (fast_tac comp_cs 1); +val rtrancl_bnd_mono = result(); + +val [prem] = goalw Trancl.thy [rtrancl_def] "r<=s ==> r^* <= s^*"; +by (rtac lfp_mono 1); +by (REPEAT (resolve_tac [rtrancl_bnd_mono, prem, subset_refl, id_mono, + comp_mono, Un_mono, field_mono, Sigma_mono] 1)); +val rtrancl_mono = result(); + +(* r^* = id(field(r)) Un ( r O r^* ) *) +val rtrancl_unfold = rtrancl_bnd_mono RS (rtrancl_def RS def_lfp_Tarski); + +(** The relation rtrancl **) + +val rtrancl_type = standard (rtrancl_def RS def_lfp_subset); + +(*Reflexivity of rtrancl*) +val [prem] = goal Trancl.thy "[| a: field(r) |] ==> : r^*"; +by (resolve_tac [rtrancl_unfold RS ssubst] 1); +by (rtac (prem RS idI RS UnI1) 1); +val rtrancl_refl = result(); + +(*Closure under composition with r *) +val prems = goal Trancl.thy + "[| : r^*; : r |] ==> : r^*"; +by (resolve_tac [rtrancl_unfold RS ssubst] 1); +by (rtac (compI RS UnI2) 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +val rtrancl_into_rtrancl = result(); + +(*rtrancl of r contains all pairs in r *) +val prems = goal Trancl.thy " : r ==> : r^*"; +by (resolve_tac [rtrancl_refl RS rtrancl_into_rtrancl] 1); +by (REPEAT (resolve_tac (prems@[fieldI1]) 1)); +val r_into_rtrancl = result(); + +(*The premise ensures that r consists entirely of pairs*) +val prems = goal Trancl.thy "r <= Sigma(A,B) ==> r <= r^*"; +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addIs [r_into_rtrancl]) 1); +val r_subset_rtrancl = result(); + +goal Trancl.thy "field(r^*) = field(r)"; +by (fast_tac (eq_cs addIs [r_into_rtrancl] + addSDs [rtrancl_type RS subsetD]) 1); +val rtrancl_field = result(); + + +(** standard induction rule **) + +val major::prems = goal Trancl.thy + "[| : r^*; \ +\ !!x. x: field(r) ==> P(); \ +\ !!x y z.[| P(); : r^*; : r |] ==> P() |] \ +\ ==> P()"; +by (rtac ([rtrancl_def, rtrancl_bnd_mono, major] MRS def_induct) 1); +by (fast_tac (ZF_cs addIs prems addSEs [idE,compE]) 1); +val rtrancl_full_induct = result(); + +(*nice induction rule. + Tried adding the typing hypotheses y,z:field(r), but these + caused expensive case splits!*) +val major::prems = goal Trancl.thy + "[| : r^*; \ +\ P(a); \ +\ !!y z.[| : r^*; : r; P(y) |] ==> P(z) \ +\ |] ==> P(b)"; +(*by induction on this formula*) +by (subgoal_tac "ALL y. = --> P(y)" 1); +(*now solve first subgoal: this formula is sufficient*) +by (EVERY1 [etac (spec RS mp), rtac refl]); +(*now do the induction*) +by (resolve_tac [major RS rtrancl_full_induct] 1); +by (ALLGOALS (fast_tac (ZF_cs addIs prems))); +val rtrancl_induct = result(); + +(*transitivity of transitive closure!! -- by induction.*) +goalw Trancl.thy [trans_def] "trans(r^*)"; +by (REPEAT (resolve_tac [allI,impI] 1)); +by (eres_inst_tac [("b","z")] rtrancl_induct 1); +by (DEPTH_SOLVE (eresolve_tac [asm_rl, rtrancl_into_rtrancl] 1)); +val trans_rtrancl = result(); + +(*elimination of rtrancl -- by induction on a special formula*) +val major::prems = goal Trancl.thy + "[| : r^*; (a=b) ==> P; \ +\ !!y.[| : r^*; : r |] ==> P |] \ +\ ==> P"; +by (subgoal_tac "a = b | (EX y. : r^* & : r)" 1); +(*see HOL/trancl*) +by (rtac (major RS rtrancl_induct) 2); +by (ALLGOALS (fast_tac (ZF_cs addSEs prems))); +val rtranclE = result(); + + +(**** The relation trancl ****) + +(*Transitivity of r^+ is proved by transitivity of r^* *) +goalw Trancl.thy [trans_def,trancl_def] "trans(r^+)"; +by (safe_tac comp_cs); +by (rtac (rtrancl_into_rtrancl RS (trans_rtrancl RS transD RS compI)) 1); +by (REPEAT (assume_tac 1)); +val trans_trancl = result(); + +(** Conversions between trancl and rtrancl **) + +val [major] = goalw Trancl.thy [trancl_def] " : r^+ ==> : r^*"; +by (resolve_tac [major RS compEpair] 1); +by (REPEAT (ares_tac [rtrancl_into_rtrancl] 1)); +val trancl_into_rtrancl = result(); + +(*r^+ contains all pairs in r *) +val [prem] = goalw Trancl.thy [trancl_def] " : r ==> : r^+"; +by (REPEAT (ares_tac [prem,compI,rtrancl_refl,fieldI1] 1)); +val r_into_trancl = result(); + +(*The premise ensures that r consists entirely of pairs*) +val prems = goal Trancl.thy "r <= Sigma(A,B) ==> r <= r^+"; +by (cut_facts_tac prems 1); +by (fast_tac (ZF_cs addIs [r_into_trancl]) 1); +val r_subset_trancl = result(); + +(*intro rule by definition: from r^* and r *) +val prems = goalw Trancl.thy [trancl_def] + "[| : r^*; : r |] ==> : r^+"; +by (REPEAT (resolve_tac ([compI]@prems) 1)); +val rtrancl_into_trancl1 = result(); + +(*intro rule from r and r^* *) +val prems = goal Trancl.thy + "[| : r; : r^* |] ==> : r^+"; +by (resolve_tac (prems RL [rtrancl_induct]) 1); +by (resolve_tac (prems RL [r_into_trancl]) 1); +by (etac (trans_trancl RS transD) 1); +by (etac r_into_trancl 1); +val rtrancl_into_trancl2 = result(); + +(*Nice induction rule for trancl*) +val major::prems = goal Trancl.thy + "[| : r^+; \ +\ !!y. [| : r |] ==> P(y); \ +\ !!y z.[| : r^+; : r; P(y) |] ==> P(z) \ +\ |] ==> P(b)"; +by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); +(*by induction on this formula*) +by (subgoal_tac "ALL z. : r --> P(z)" 1); +(*now solve first subgoal: this formula is sufficient*) +by (fast_tac ZF_cs 1); +by (etac rtrancl_induct 1); +by (ALLGOALS (fast_tac (ZF_cs addIs (rtrancl_into_trancl1::prems)))); +val trancl_induct = result(); + +(*elimination of r^+ -- NOT an induction rule*) +val major::prems = goal Trancl.thy + "[| : r^+; \ +\ : r ==> P; \ +\ !!y.[| : r^+; : r |] ==> P \ +\ |] ==> P"; +by (subgoal_tac " : r | (EX y. : r^+ & : r)" 1); +by (fast_tac (ZF_cs addIs prems) 1); +by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); +by (etac rtranclE 1); +by (ALLGOALS (fast_tac (ZF_cs addIs [rtrancl_into_trancl1]))); +val tranclE = result(); + +goalw Trancl.thy [trancl_def] "r^+ <= field(r)*field(r)"; +by (fast_tac (ZF_cs addEs [compE, rtrancl_type RS subsetD RS SigmaE2]) 1); +val trancl_type = result(); + +val [prem] = goalw Trancl.thy [trancl_def] "r<=s ==> r^+ <= s^+"; +by (REPEAT (resolve_tac [prem, comp_mono, rtrancl_mono] 1)); +val trancl_mono = result(); + diff -r 000000000000 -r a5a9c433f639 src/ZF/trancl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/trancl.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,21 @@ +(* Title: ZF/trancl.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Transitive closure of a relation +*) + +Trancl = Fixedpt + Perm + +consts + "rtrancl" :: "i=>i" ("(_^*)" [100] 100) (*refl/transitive closure*) + "trancl" :: "i=>i" ("(_^+)" [100] 100) (*transitive closure*) + "trans" :: "i=>o" (*transitivity predicate*) + +rules + trans_def "trans(r) == ALL x y z. : r --> : r --> : r" + + rtrancl_def "r^* == lfp(field(r)*field(r), %s. id(field(r)) Un (r O s))" + + trancl_def "r^+ == r O r^*" +end diff -r 000000000000 -r a5a9c433f639 src/ZF/univ.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/univ.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,642 @@ +(* Title: ZF/univ + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +The cumulative hierarchy and a small universe for recursive types +*) + +open Univ; + +(*NOT SUITABLE FOR REWRITING -- RECURSIVE!*) +goal Univ.thy "Vfrom(A,i) = A Un (UN j:i. Pow(Vfrom(A,j)))"; +by (rtac (Vfrom_def RS def_transrec RS ssubst) 1); +by (SIMP_TAC ZF_ss 1); +val Vfrom = result(); + +(** Monotonicity **) + +goal Univ.thy "!!A B. A<=B ==> ALL j. i<=j --> Vfrom(A,i) <= Vfrom(B,j)"; +by (eps_ind_tac "i" 1); +by (rtac (impI RS allI) 1); +by (rtac (Vfrom RS ssubst) 1); +by (rtac (Vfrom RS ssubst) 1); +by (etac Un_mono 1); +by (rtac UN_mono 1); +by (assume_tac 1); +by (rtac Pow_mono 1); +by (etac (bspec RS spec RS mp) 1); +by (assume_tac 1); +by (rtac subset_refl 1); +val Vfrom_mono_lemma = result(); + +(* [| A<=B; i<=x |] ==> Vfrom(A,i) <= Vfrom(B,x) *) +val Vfrom_mono = standard (Vfrom_mono_lemma RS spec RS mp); + + +(** A fundamental equality: Vfrom does not require ordinals! **) + +goal Univ.thy "Vfrom(A,x) <= Vfrom(A,rank(x))"; +by (eps_ind_tac "x" 1); +by (rtac (Vfrom RS ssubst) 1); +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac (ZF_cs addSIs [rank_lt]) 1); +val Vfrom_rank_subset1 = result(); + +goal Univ.thy "Vfrom(A,rank(x)) <= Vfrom(A,x)"; +by (eps_ind_tac "x" 1); +by (rtac (Vfrom RS ssubst) 1); +by (rtac (Vfrom RS ssubst) 1); +br (subset_refl RS Un_mono) 1; +br UN_least 1; +by (etac (rank_implies_mem RS bexE) 1); +br subset_trans 1; +be UN_upper 2; +by (etac (subset_refl RS Vfrom_mono RS subset_trans RS Pow_mono) 1); +by (etac bspec 1); +by (assume_tac 1); +val Vfrom_rank_subset2 = result(); + +goal Univ.thy "Vfrom(A,rank(x)) = Vfrom(A,x)"; +by (rtac equalityI 1); +by (rtac Vfrom_rank_subset2 1); +by (rtac Vfrom_rank_subset1 1); +val Vfrom_rank_eq = result(); + + +(*** Basic closure properties ***) + +goal Univ.thy "!!x y. y:x ==> 0 : Vfrom(A,x)"; +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac ZF_cs 1); +val zero_in_Vfrom = result(); + +goal Univ.thy "i <= Vfrom(A,i)"; +by (eps_ind_tac "i" 1); +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac ZF_cs 1); +val i_subset_Vfrom = result(); + +goal Univ.thy "A <= Vfrom(A,i)"; +by (rtac (Vfrom RS ssubst) 1); +by (rtac Un_upper1 1); +val A_subset_Vfrom = result(); + +goal Univ.thy "!!A a i. a <= Vfrom(A,i) ==> a: Vfrom(A,succ(i))"; +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac ZF_cs 1); +val subset_mem_Vfrom = result(); + +(** Finite sets and ordered pairs **) + +goal Univ.thy "!!a. a: Vfrom(A,i) ==> {a} : Vfrom(A,succ(i))"; +by (rtac subset_mem_Vfrom 1); +by (safe_tac ZF_cs); +val singleton_in_Vfrom = result(); + +goal Univ.thy + "!!A. [| a: Vfrom(A,i); b: Vfrom(A,i) |] ==> {a,b} : Vfrom(A,succ(i))"; +by (rtac subset_mem_Vfrom 1); +by (safe_tac ZF_cs); +val doubleton_in_Vfrom = result(); + +goalw Univ.thy [Pair_def] + "!!A. [| a: Vfrom(A,i); b: Vfrom(A,i) |] ==> \ +\ : Vfrom(A,succ(succ(i)))"; +by (REPEAT (ares_tac [doubleton_in_Vfrom] 1)); +val Pair_in_Vfrom = result(); + +val [prem] = goal Univ.thy + "a<=Vfrom(A,i) ==> succ(a) : Vfrom(A,succ(succ(i)))"; +by (REPEAT (resolve_tac [subset_mem_Vfrom, succ_subsetI] 1)); +by (rtac (Vfrom_mono RSN (2,subset_trans)) 2); +by (REPEAT (resolve_tac [prem, subset_refl, subset_succI] 1)); +val succ_in_Vfrom = result(); + +(*** 0, successor and limit equations fof Vfrom ***) + +goal Univ.thy "Vfrom(A,0) = A"; +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac eq_cs 1); +val Vfrom_0 = result(); + +goal Univ.thy "!!i. Ord(i) ==> Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"; +by (rtac (Vfrom RS trans) 1); +brs ([refl] RL ZF_congs) 1; +by (rtac equalityI 1); +by (rtac (succI1 RS RepFunI RS Union_upper) 2); +by (rtac UN_least 1); +by (rtac (subset_refl RS Vfrom_mono RS Pow_mono) 1); +by (etac member_succD 1); +by (assume_tac 1); +val Vfrom_succ_lemma = result(); + +goal Univ.thy "Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"; +by (res_inst_tac [("x1", "succ(i)")] (Vfrom_rank_eq RS subst) 1); +by (res_inst_tac [("x1", "i")] (Vfrom_rank_eq RS subst) 1); +by (rtac (rank_succ RS ssubst) 1); +by (rtac (Ord_rank RS Vfrom_succ_lemma) 1); +val Vfrom_succ = result(); + +(*The premise distinguishes this from Vfrom(A,0); allowing X=0 forces + the conclusion to be Vfrom(A,Union(X)) = A Un (UN y:X. Vfrom(A,y)) *) +val [prem] = goal Univ.thy "y:X ==> Vfrom(A,Union(X)) = (UN y:X. Vfrom(A,y))"; +by (rtac (Vfrom RS ssubst) 1); +by (rtac equalityI 1); +(*first inclusion*) +by (rtac Un_least 1); +br (A_subset_Vfrom RS subset_trans) 1; +br (prem RS UN_upper) 1; +br UN_least 1; +be UnionE 1; +br subset_trans 1; +be UN_upper 2; +by (rtac (Vfrom RS ssubst) 1); +be ([UN_upper, Un_upper2] MRS subset_trans) 1; +(*opposite inclusion*) +br UN_least 1; +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac ZF_cs 1); +val Vfrom_Union = result(); + +(*** Limit ordinals -- general properties ***) + +goalw Univ.thy [Limit_def] "!!i. Limit(i) ==> Union(i) = i"; +by (fast_tac (eq_cs addEs [Ord_trans]) 1); +val Limit_Union_eq = result(); + +goalw Univ.thy [Limit_def] "!!i. Limit(i) ==> Ord(i)"; +by (etac conjunct1 1); +val Limit_is_Ord = result(); + +goalw Univ.thy [Limit_def] "!!i. Limit(i) ==> 0 : i"; +by (fast_tac ZF_cs 1); +val Limit_has_0 = result(); + +goalw Univ.thy [Limit_def] "!!i. [| Limit(i); j:i |] ==> succ(j) : i"; +by (fast_tac ZF_cs 1); +val Limit_has_succ = result(); + +goalw Univ.thy [Limit_def] "Limit(nat)"; +by (REPEAT (ares_tac [conjI, ballI, nat_0I, nat_succI, Ord_nat] 1)); +val Limit_nat = result(); + +goalw Univ.thy [Limit_def] + "!!i. [| Ord(i); 0:i; ALL y. ~ succ(y)=i |] ==> Limit(i)"; +by (safe_tac subset_cs); +br Ord_member 1; +by (REPEAT_FIRST (eresolve_tac [asm_rl, Ord_in_Ord RS Ord_succ] + ORELSE' dresolve_tac [Ord_succ_subsetI])); +by (fast_tac (subset_cs addSIs [equalityI]) 1); +val non_succ_LimitI = result(); + +goal Univ.thy "!!i. Ord(i) ==> i=0 | (EX j. i=succ(j)) | Limit(i)"; +by (fast_tac (ZF_cs addSIs [non_succ_LimitI, Ord_0_member_iff RS iffD2]) 1); +val Ord_cases_lemma = result(); + +val major::prems = goal Univ.thy + "[| Ord(i); \ +\ i=0 ==> P; \ +\ !!j. i=succ(j) ==> P; \ +\ Limit(i) ==> P \ +\ |] ==> P"; +by (cut_facts_tac [major RS Ord_cases_lemma] 1); +by (REPEAT (eresolve_tac (prems@[disjE, exE]) 1)); +val Ord_cases = result(); + + +(*** Vfrom applied to Limit ordinals ***) + +(*NB. limit ordinals are non-empty; + Vfrom(A,0) = A = A Un (UN y:0. Vfrom(A,y)) *) +val [limiti] = goal Univ.thy + "Limit(i) ==> Vfrom(A,i) = (UN y:i. Vfrom(A,y))"; +by (rtac (limiti RS Limit_has_0 RS Vfrom_Union RS subst) 1); +by (rtac (limiti RS Limit_Union_eq RS ssubst) 1); +by (rtac refl 1); +val Limit_Vfrom_eq = result(); + +val Limit_VfromE = standard (Limit_Vfrom_eq RS equalityD1 RS subsetD RS UN_E); + +val [major,limiti] = goal Univ.thy + "[| a: Vfrom(A,i); Limit(i) |] ==> {a} : Vfrom(A,i)"; +by (rtac (limiti RS Limit_VfromE) 1); +by (rtac major 1); +by (rtac (limiti RS Limit_Vfrom_eq RS ssubst) 1); +by (rtac UN_I 1); +by (etac singleton_in_Vfrom 2); +by (etac (limiti RS Limit_has_succ) 1); +val singleton_in_Vfrom_limit = result(); + +val Vfrom_UnI1 = Un_upper1 RS (subset_refl RS Vfrom_mono RS subsetD) +and Vfrom_UnI2 = Un_upper2 RS (subset_refl RS Vfrom_mono RS subsetD); + +(*Hard work is finding a single j:i such that {a,b}<=Vfrom(A,j)*) +val [aprem,bprem,limiti] = goal Univ.thy + "[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i) |] ==> \ +\ {a,b} : Vfrom(A,i)"; +by (rtac (aprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (bprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (limiti RS Limit_Vfrom_eq RS ssubst) 1); +by (rtac UN_I 1); +by (rtac doubleton_in_Vfrom 2); +by (etac Vfrom_UnI1 2); +by (etac Vfrom_UnI2 2); +by (REPEAT (ares_tac[limiti, Limit_has_succ, Ord_member_UnI, Limit_is_Ord] 1)); +val doubleton_in_Vfrom_limit = result(); + +val [aprem,bprem,limiti] = goal Univ.thy + "[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i) |] ==> \ +\ : Vfrom(A,i)"; +(*Infer that a, b occur at ordinals x,xa < i.*) +by (rtac (aprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (bprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (limiti RS Limit_Vfrom_eq RS ssubst) 1); +by (rtac UN_I 1); +by (rtac Pair_in_Vfrom 2); +(*Infer that succ(succ(x Un xa)) < i *) +by (etac Vfrom_UnI1 2); +by (etac Vfrom_UnI2 2); +by (REPEAT (ares_tac[limiti, Limit_has_succ, Ord_member_UnI, Limit_is_Ord] 1)); +val Pair_in_Vfrom_limit = result(); + + +(*** Properties assuming Transset(A) ***) + +goal Univ.thy "!!i A. Transset(A) ==> Transset(Vfrom(A,i))"; +by (eps_ind_tac "i" 1); +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac (ZF_cs addSIs [Transset_Union_family, Transset_Un, + Transset_Pow]) 1); +val Transset_Vfrom = result(); + +goal Univ.thy "!!A i. Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))"; +by (rtac (Vfrom_succ RS trans) 1); +br (Un_upper2 RSN (2,equalityI)) 1;; +br (subset_refl RSN (2,Un_least)) 1;; +br (A_subset_Vfrom RS subset_trans) 1; +be (Transset_Vfrom RS (Transset_iff_Pow RS iffD1)) 1; +val Transset_Vfrom_succ = result(); + +goalw Ord.thy [Pair_def,Transset_def] + "!!C. [| <= C; Transset(C) |] ==> a: C & b: C"; +by (fast_tac ZF_cs 1); +val Transset_Pair_subset = result(); + +goal Univ.thy + "!!a b.[| <= Vfrom(A,i); Transset(A); Limit(i) |] ==> \ +\ : Vfrom(A,i)"; +be (Transset_Pair_subset RS conjE) 1; +be Transset_Vfrom 1; +by (REPEAT (ares_tac [Pair_in_Vfrom_limit] 1)); +val Transset_Pair_subset_Vfrom_limit = result(); + + +(*** Closure under product/sum applied to elements -- thus Vfrom(A,i) + is a model of simple type theory provided A is a transitive set + and i is a limit ordinal +***) + +(*There are three nearly identical proofs below -- needs a general theorem + for proving ...a...b : Vfrom(A,i) where i is a limit ordinal*) + +(** products **) + +goal Univ.thy + "!!A. [| a: Vfrom(A,i); b: Vfrom(A,i); Transset(A) |] ==> \ +\ a*b : Vfrom(A, succ(succ(succ(i))))"; +by (dtac Transset_Vfrom 1); +by (rtac subset_mem_Vfrom 1); +by (rewtac Transset_def); +by (fast_tac (ZF_cs addIs [Pair_in_Vfrom]) 1); +val prod_in_Vfrom = result(); + +val [aprem,bprem,limiti,transset] = goal Univ.thy + "[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); Transset(A) |] ==> \ +\ a*b : Vfrom(A,i)"; +(*Infer that a, b occur at ordinals x,xa < i.*) +by (rtac (aprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (bprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (limiti RS Limit_Vfrom_eq RS ssubst) 1); +by (rtac UN_I 1); +by (rtac prod_in_Vfrom 2); +(*Infer that succ(succ(succ(x Un xa))) < i *) +by (etac (Vfrom_UnI1 RS Vfrom_UnI2) 2); +by (etac (Vfrom_UnI2 RS Vfrom_UnI2) 2); +by (REPEAT (ares_tac [limiti RS Limit_has_succ, + Ord_member_UnI, limiti RS Limit_is_Ord, transset] 1)); +val prod_in_Vfrom_limit = result(); + +(** Disjoint sums, aka Quine ordered pairs **) + +goalw Univ.thy [sum_def] + "!!A. [| a: Vfrom(A,i); b: Vfrom(A,i); Transset(A); 1:i |] ==> \ +\ a+b : Vfrom(A, succ(succ(succ(i))))"; +by (dtac Transset_Vfrom 1); +by (rtac subset_mem_Vfrom 1); +by (rewtac Transset_def); +by (fast_tac (ZF_cs addIs [zero_in_Vfrom, Pair_in_Vfrom, + i_subset_Vfrom RS subsetD]) 1); +val sum_in_Vfrom = result(); + +val [aprem,bprem,limiti,transset] = goal Univ.thy + "[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); Transset(A) |] ==> \ +\ a+b : Vfrom(A,i)"; +(*Infer that a, b occur at ordinals x,xa < i.*) +by (rtac (aprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (bprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (limiti RS Limit_Vfrom_eq RS ssubst) 1); +by (rtac UN_I 1); +by (rtac (rewrite_rule [one_def] sum_in_Vfrom) 2); +by (rtac (succI1 RS UnI1) 5); +(*Infer that succ(succ(succ(x Un xa))) < i *) +by (etac (Vfrom_UnI1 RS Vfrom_UnI2) 2); +by (etac (Vfrom_UnI2 RS Vfrom_UnI2) 2); +by (REPEAT (ares_tac [limiti RS Limit_has_0, + limiti RS Limit_has_succ, + Ord_member_UnI, limiti RS Limit_is_Ord, transset] 1)); +val sum_in_Vfrom_limit = result(); + +(** function space! **) + +goalw Univ.thy [Pi_def] + "!!A. [| a: Vfrom(A,i); b: Vfrom(A,i); Transset(A) |] ==> \ +\ a->b : Vfrom(A, succ(succ(succ(succ(i)))))"; +by (dtac Transset_Vfrom 1); +by (rtac subset_mem_Vfrom 1); +by (rtac (Collect_subset RS subset_trans) 1); +by (rtac (Vfrom RS ssubst) 1); +by (rtac (subset_trans RS subset_trans) 1); +by (rtac Un_upper2 3); +by (rtac (succI1 RS UN_upper) 2); +by (rtac Pow_mono 1); +by (rewtac Transset_def); +by (fast_tac (ZF_cs addIs [Pair_in_Vfrom]) 1); +val fun_in_Vfrom = result(); + +val [aprem,bprem,limiti,transset] = goal Univ.thy + "[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); Transset(A) |] ==> \ +\ a->b : Vfrom(A,i)"; +(*Infer that a, b occur at ordinals x,xa < i.*) +by (rtac (aprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (bprem RS (limiti RS Limit_VfromE)) 1); +by (rtac (limiti RS Limit_Vfrom_eq RS ssubst) 1); +by (rtac UN_I 1); +by (rtac fun_in_Vfrom 2); +(*Infer that succ(succ(succ(x Un xa))) < i *) +by (etac (Vfrom_UnI1 RS Vfrom_UnI2) 2); +by (etac (Vfrom_UnI2 RS Vfrom_UnI2) 2); +by (REPEAT (ares_tac [limiti RS Limit_has_succ, + Ord_member_UnI, limiti RS Limit_is_Ord, transset] 1)); +val fun_in_Vfrom_limit = result(); + + +(*** The set Vset(i) ***) + +goal Univ.thy "Vset(i) = (UN j:i. Pow(Vset(j)))"; +by (rtac (Vfrom RS ssubst) 1); +by (fast_tac eq_cs 1); +val Vset = result(); + +val Vset_succ = Transset_0 RS Transset_Vfrom_succ; + +val Transset_Vset = Transset_0 RS Transset_Vfrom; + +(** Characterisation of the elements of Vset(i) **) + +val [ordi] = goal Univ.thy "Ord(i) ==> ALL b. b : Vset(i) --> rank(b) : i"; +by (rtac (ordi RS trans_induct) 1); +by (rtac (Vset RS ssubst) 1); +by (safe_tac ZF_cs); +by (rtac (rank RS ssubst) 1); +by (rtac sup_least2 1); +by (assume_tac 1); +by (assume_tac 1); +by (fast_tac ZF_cs 1); +val Vset_rank_imp1 = result(); + +(* [| Ord(i); x : Vset(i) |] ==> rank(x) : i *) +val Vset_D = standard (Vset_rank_imp1 RS spec RS mp); + +val [ordi] = goal Univ.thy "Ord(i) ==> ALL b. rank(b) : i --> b : Vset(i)"; +by (rtac (ordi RS trans_induct) 1); +by (rtac allI 1); +by (rtac (Vset RS ssubst) 1); +by (fast_tac (ZF_cs addSIs [rank_lt]) 1); +val Vset_rank_imp2 = result(); + +(* [| Ord(i); rank(x) : i |] ==> x : Vset(i) *) +val VsetI = standard (Vset_rank_imp2 RS spec RS mp); + +val [ordi] = goal Univ.thy "Ord(i) ==> b : Vset(i) <-> rank(b) : i"; +by (rtac iffI 1); +by (etac (ordi RS Vset_D) 1); +by (etac (ordi RS VsetI) 1); +val Vset_Ord_rank_iff = result(); + +goal Univ.thy "b : Vset(a) <-> rank(b) : rank(a)"; +by (rtac (Vfrom_rank_eq RS subst) 1); +by (rtac (Ord_rank RS Vset_Ord_rank_iff) 1); +val Vset_rank_iff = result(); + +goal Univ.thy "!!i. Ord(i) ==> rank(Vset(i)) = i"; +by (rtac (rank RS ssubst) 1); +by (rtac equalityI 1); +by (safe_tac ZF_cs); +by (EVERY' [rtac UN_I, + etac (i_subset_Vfrom RS subsetD), + etac (Ord_in_Ord RS rank_of_Ord RS ssubst), + assume_tac, + rtac succI1] 3); +by (REPEAT (eresolve_tac [asm_rl,Vset_D,Ord_trans] 1)); +val rank_Vset = result(); + +(** Lemmas for reasoning about sets in terms of their elements' ranks **) + +(* rank(x) : rank(a) ==> x : Vset(rank(a)) *) +val Vset_rankI = Ord_rank RS VsetI; + +goal Univ.thy "a <= Vset(rank(a))"; +br subsetI 1; +be (rank_lt RS Vset_rankI) 1; +val arg_subset_Vset_rank = result(); + +val [iprem] = goal Univ.thy + "[| !!i. Ord(i) ==> a Int Vset(i) <= b |] ==> a <= b"; +br ([subset_refl, arg_subset_Vset_rank] MRS Int_greatest RS subset_trans) 1; +br (Ord_rank RS iprem) 1; +val Int_Vset_subset = result(); + +(** Set up an environment for simplification **) + +val rank_rls = [rank_Inl, rank_Inr, rank_pair1, rank_pair2]; +val rank_trans_rls = rank_rls @ (rank_rls RLN (2, [rank_trans])); + +val rank_ss = ZF_ss + addrews [split, case_Inl, case_Inr, Vset_rankI] + addrews rank_trans_rls; + +(** Recursion over Vset levels! **) + +(*NOT SUITABLE FOR REWRITING: recursive!*) +goalw Univ.thy [Vrec_def] "Vrec(a,H) = H(a, lam x:Vset(rank(a)). Vrec(x,H))"; +by (rtac (transrec RS ssubst) 1); +by (SIMP_TAC (wf_ss addrews [Ord_rank, Ord_succ, Vset_D RS beta, + VsetI RS beta]) 1); +val Vrec = result(); + +(*This form avoids giant explosions in proofs. NOTE USE OF == *) +val rew::prems = goal Univ.thy + "[| !!x. h(x)==Vrec(x,H) |] ==> \ +\ h(a) = H(a, lam x: Vset(rank(a)). h(x))"; +by (rewtac rew); +by (rtac Vrec 1); +val def_Vrec = result(); + +val prems = goalw Univ.thy [Vrec_def] + "[| a=a'; !!x u. H(x,u)=H'(x,u) |] ==> Vrec(a,H)=Vrec(a',H')"; +val Vrec_ss = ZF_ss addcongs ([transrec_cong] @ mk_congs Univ.thy ["rank"]) + addrews (prems RL [sym]); +by (SIMP_TAC Vrec_ss 1); +val Vrec_cong = result(); + + +(*** univ(A) ***) + +goalw Univ.thy [univ_def] "!!A B. A<=B ==> univ(A) <= univ(B)"; +by (etac Vfrom_mono 1); +by (rtac subset_refl 1); +val univ_mono = result(); + +goalw Univ.thy [univ_def] "!!A. Transset(A) ==> Transset(univ(A))"; +by (etac Transset_Vfrom 1); +val Transset_univ = result(); + +(** univ(A) as a limit **) + +goalw Univ.thy [univ_def] "univ(A) = (UN i:nat. Vfrom(A,i))"; +br (Limit_nat RS Limit_Vfrom_eq) 1; +val univ_eq_UN = result(); + +goal Univ.thy "!!c. c <= univ(A) ==> c = (UN i:nat. c Int Vfrom(A,i))"; +br (subset_UN_iff_eq RS iffD1) 1; +be (univ_eq_UN RS subst) 1; +val subset_univ_eq_Int = result(); + +val [aprem, iprem] = goal Univ.thy + "[| a <= univ(X); \ +\ !!i. i:nat ==> a Int Vfrom(X,i) <= b \ +\ |] ==> a <= b"; +br (aprem RS subset_univ_eq_Int RS ssubst) 1; +br UN_least 1; +be iprem 1; +val univ_Int_Vfrom_subset = result(); + +val prems = goal Univ.thy + "[| a <= univ(X); b <= univ(X); \ +\ !!i. i:nat ==> a Int Vfrom(X,i) = b Int Vfrom(X,i) \ +\ |] ==> a = b"; +br equalityI 1; +by (ALLGOALS + (resolve_tac (prems RL [univ_Int_Vfrom_subset]) THEN' + eresolve_tac (prems RL [equalityD1,equalityD2] RL [subset_trans]) THEN' + rtac Int_lower1)); +val univ_Int_Vfrom_eq = result(); + +(** Closure properties **) + +goalw Univ.thy [univ_def] "0 : univ(A)"; +by (rtac (nat_0I RS zero_in_Vfrom) 1); +val zero_in_univ = result(); + +goalw Univ.thy [univ_def] "A <= univ(A)"; +by (rtac A_subset_Vfrom 1); +val A_subset_univ = result(); + +val A_into_univ = A_subset_univ RS subsetD; + +(** Closure under unordered and ordered pairs **) + +goalw Univ.thy [univ_def] "!!A a. a: univ(A) ==> {a} : univ(A)"; +by (rtac singleton_in_Vfrom_limit 1); +by (REPEAT (ares_tac [Ord_nat,Limit_nat] 1)); +val singleton_in_univ = result(); + +goalw Univ.thy [univ_def] + "!!A a. [| a: univ(A); b: univ(A) |] ==> {a,b} : univ(A)"; +by (rtac doubleton_in_Vfrom_limit 1); +by (REPEAT (ares_tac [Ord_nat,Limit_nat] 1)); +val doubleton_in_univ = result(); + +goalw Univ.thy [univ_def] + "!!A a. [| a: univ(A); b: univ(A) |] ==> : univ(A)"; +by (rtac Pair_in_Vfrom_limit 1); +by (REPEAT (ares_tac [Ord_nat,Limit_nat] 1)); +val Pair_in_univ = result(); + +goal Univ.thy "univ(A)*univ(A) <= univ(A)"; +by (REPEAT (ares_tac [subsetI,Pair_in_univ] 1 + ORELSE eresolve_tac [SigmaE, ssubst] 1)); +val product_univ = result(); + +val Sigma_subset_univ = standard + (Sigma_mono RS (product_univ RSN (2,subset_trans))); + +goalw Univ.thy [univ_def] + "!!a b.[| <= univ(A); Transset(A) |] ==> : univ(A)"; +be Transset_Pair_subset_Vfrom_limit 1; +by (REPEAT (ares_tac [Ord_nat,Limit_nat] 1)); +val Transset_Pair_subset_univ = result(); + + +(** The natural numbers **) + +goalw Univ.thy [univ_def] "nat <= univ(A)"; +by (rtac i_subset_Vfrom 1); +val nat_subset_univ = result(); + +(* n:nat ==> n:univ(A) *) +val nat_into_univ = standard (nat_subset_univ RS subsetD); + +(** instances for 1 and 2 **) + +goalw Univ.thy [one_def] "1 : univ(A)"; +by (REPEAT (ares_tac [nat_into_univ, nat_0I, nat_succI] 1)); +val one_in_univ = result(); + +(*unused!*) +goal Univ.thy "succ(succ(0)) : univ(A)"; +by (REPEAT (ares_tac [nat_into_univ, nat_0I, nat_succI] 1)); +val two_in_univ = result(); + +goalw Univ.thy [bool_def] "bool <= univ(A)"; +by (fast_tac (ZF_cs addSIs [zero_in_univ,one_in_univ]) 1); +val bool_subset_univ = result(); + +val bool_into_univ = standard (bool_subset_univ RS subsetD); + + +(** Closure under disjoint union **) + +goalw Univ.thy [Inl_def] "!!A a. a: univ(A) ==> Inl(a) : univ(A)"; +by (REPEAT (ares_tac [zero_in_univ,Pair_in_univ] 1)); +val Inl_in_univ = result(); + +goalw Univ.thy [Inr_def] "!!A b. b: univ(A) ==> Inr(b) : univ(A)"; +by (REPEAT (ares_tac [one_in_univ, Pair_in_univ] 1)); +val Inr_in_univ = result(); + +goal Univ.thy "univ(C)+univ(C) <= univ(C)"; +by (REPEAT (ares_tac [subsetI,Inl_in_univ,Inr_in_univ] 1 + ORELSE eresolve_tac [sumE, ssubst] 1)); +val sum_univ = result(); + +val sum_subset_univ = standard + (sum_mono RS (sum_univ RSN (2,subset_trans))); + + +(** Closure under binary union -- use Un_least **) +(** Closure under Collect -- use (Collect_subset RS subset_trans) **) +(** Closure under RepFun -- use RepFun_subset **) + + diff -r 000000000000 -r a5a9c433f639 src/ZF/univ.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/univ.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,34 @@ +(* Title: ZF/univ.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +The cumulative hierarchy and a small universe for recursive types + +Standard notation for Vset(i) is V(i), but users might want V for a variable +*) + +Univ = Arith + Sum + +consts + Limit :: "i=>o" + Vfrom :: "[i,i]=>i" + Vset :: "i=>i" + Vrec :: "[i, [i,i]=>i] =>i" + univ :: "i=>i" + +translations + (*Apparently a bug prevents using "Vset" == "Vfrom(0)" *) + "Vset(x)" == "Vfrom(0,x)" + +rules + Limit_def "Limit(i) == Ord(i) & 0:i & (ALL y:i. succ(y): i)" + + Vfrom_def "Vfrom(A,i) == transrec(i, %x f. A Un (UN y:x. Pow(f`y)))" + + Vrec_def + "Vrec(a,H) == transrec(rank(a), %x g. lam z: Vset(succ(x)). \ +\ H(z, lam w:Vset(x). g`rank(w)`w)) ` a" + + univ_def "univ(A) == Vfrom(A,nat)" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/upair.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/upair.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,282 @@ +(* Title: ZF/upair + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +UNORDERED pairs in Zermelo-Fraenkel Set Theory + +Observe the order of dependence: + Upair is defined in terms of Replace + Un is defined in terms of Upair and Union (similarly for Int) + cons is defined in terms of Upair and Un + Ordered pairs and descriptions are defined using cons ("set notation") +*) + +(*** Lemmas about power sets ***) + +val Pow_bottom = empty_subsetI RS PowI; (* 0 : Pow(B) *) +val Pow_top = subset_refl RS PowI; (* A : Pow(A) *) +val Pow_neq_0 = Pow_top RSN (2,equals0D); (* Pow(a)=0 ==> P *) + + +(*** Unordered pairs - Upair ***) + +val pairing = prove_goalw ZF.thy [Upair_def] + "c : Upair(a,b) <-> (c=a | c=b)" + (fn _ => [ (fast_tac (lemmas_cs addEs [Pow_neq_0, sym RS Pow_neq_0]) 1) ]); + +val UpairI1 = prove_goal ZF.thy "a : Upair(a,b)" + (fn _ => [ (rtac (refl RS disjI1 RS (pairing RS iffD2)) 1) ]); + +val UpairI2 = prove_goal ZF.thy "b : Upair(a,b)" + (fn _ => [ (rtac (refl RS disjI2 RS (pairing RS iffD2)) 1) ]); + +val UpairE = prove_goal ZF.thy + "[| a : Upair(b,c); a=b ==> P; a=c ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS (pairing RS iffD1 RS disjE)) 1), + (REPEAT (eresolve_tac prems 1)) ]); + +(*** Rules for binary union -- Un -- defined via Upair ***) + +val UnI1 = prove_goalw ZF.thy [Un_def] "c : A ==> c : A Un B" + (fn [prem]=> [ (rtac (prem RS (UpairI1 RS UnionI)) 1) ]); + +val UnI2 = prove_goalw ZF.thy [Un_def] "c : B ==> c : A Un B" + (fn [prem]=> [ (rtac (prem RS (UpairI2 RS UnionI)) 1) ]); + +val UnE = prove_goalw ZF.thy [Un_def] + "[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS UnionE) 1), + (etac UpairE 1), + (REPEAT (EVERY1 [resolve_tac prems, etac subst, assume_tac])) ]); + +val Un_iff = prove_goal ZF.thy "c : A Un B <-> (c:A | c:B)" + (fn _ => [ (fast_tac (lemmas_cs addIs [UnI1,UnI2] addSEs [UnE]) 1) ]); + +(*Classical introduction rule: no commitment to A vs B*) +val UnCI = prove_goal ZF.thy "(~c : B ==> c : A) ==> c : A Un B" + (fn [prem]=> + [ (rtac (disjCI RS (Un_iff RS iffD2)) 1), + (etac prem 1) ]); + + +(*** Rules for small intersection -- Int -- defined via Upair ***) + +val IntI = prove_goalw ZF.thy [Int_def] + "[| c : A; c : B |] ==> c : A Int B" + (fn prems=> + [ (REPEAT (resolve_tac (prems @ [UpairI1,InterI]) 1 + ORELSE eresolve_tac [UpairE, ssubst] 1)) ]); + +val IntD1 = prove_goalw ZF.thy [Int_def] "c : A Int B ==> c : A" + (fn [major]=> + [ (rtac (UpairI1 RS (major RS InterD)) 1) ]); + +val IntD2 = prove_goalw ZF.thy [Int_def] "c : A Int B ==> c : B" + (fn [major]=> + [ (rtac (UpairI2 RS (major RS InterD)) 1) ]); + +val IntE = prove_goal ZF.thy + "[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P" + (fn prems=> + [ (resolve_tac prems 1), + (REPEAT (resolve_tac (prems RL [IntD1,IntD2]) 1)) ]); + +val Int_iff = prove_goal ZF.thy "c : A Int B <-> (c:A & c:B)" + (fn _ => [ (fast_tac (lemmas_cs addSIs [IntI] addSEs [IntE]) 1) ]); + + +(*** Rules for set difference -- defined via Upair ***) + +val DiffI = prove_goalw ZF.thy [Diff_def] + "[| c : A; ~ c : B |] ==> c : A - B" + (fn prems=> [ (REPEAT (resolve_tac (prems @ [CollectI]) 1)) ]); + +val DiffD1 = prove_goalw ZF.thy [Diff_def] + "c : A - B ==> c : A" + (fn [major]=> [ (rtac (major RS CollectD1) 1) ]); + +val DiffD2 = prove_goalw ZF.thy [Diff_def] + "[| c : A - B; c : B |] ==> P" + (fn [major,minor]=> [ (rtac (minor RS (major RS CollectD2 RS notE)) 1) ]); + +val DiffE = prove_goal ZF.thy + "[| c : A - B; [| c:A; ~ c:B |] ==> P |] ==> P" + (fn prems=> + [ (resolve_tac prems 1), + (REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ]); + +val Diff_iff = prove_goal ZF.thy "c : A-B <-> (c:A & ~ c:B)" + (fn _ => [ (fast_tac (lemmas_cs addSIs [DiffI] addSEs [DiffE]) 1) ]); + +(*** Rules for cons -- defined via Un and Upair ***) + +val consI1 = prove_goalw ZF.thy [cons_def] "a : cons(a,B)" + (fn _ => [ (rtac (UpairI1 RS UnI1) 1) ]); + +val consI2 = prove_goalw ZF.thy [cons_def] "a : B ==> a : cons(b,B)" + (fn [prem]=> [ (rtac (prem RS UnI2) 1) ]); + +val consE = prove_goalw ZF.thy [cons_def] + "[| a : cons(b,A); a=b ==> P; a:A ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS UnE) 1), + (REPEAT (eresolve_tac (prems @ [UpairE]) 1)) ]); + +val cons_iff = prove_goal ZF.thy "a : cons(b,A) <-> (a=b | a:A)" + (fn _ => [ (fast_tac (lemmas_cs addIs [consI1,consI2] addSEs [consE]) 1) ]); + +(*Classical introduction rule*) +val consCI = prove_goal ZF.thy "(~ a:B ==> a=b) ==> a: cons(b,B)" + (fn [prem]=> + [ (rtac (disjCI RS (cons_iff RS iffD2)) 1), + (etac prem 1) ]); + +(*** Singletons - using cons ***) + +val singletonI = prove_goal ZF.thy "a : {a}" + (fn _=> [ (rtac consI1 1) ]); + +val singletonE = prove_goal ZF.thy "[| a: {b}; a=b ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS consE) 1), + (REPEAT (eresolve_tac (prems @ [emptyE]) 1)) ]); + + +(*** Rules for Descriptions ***) + +val the_equality = prove_goalw ZF.thy [the_def] + "[| P(a); !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a" + (fn prems=> + [ (fast_tac (lemmas_cs addIs ([equalityI,singletonI]@prems) + addEs (prems RL [subst])) 1) ]); + +(* Only use this if you already know EX!x. P(x) *) +val the_equality2 = prove_goal ZF.thy + "[| EX! x. P(x); P(a) |] ==> (THE x. P(x)) = a" + (fn major::prems=> + [ (rtac the_equality 1), + (rtac (major RS ex1_equalsE) 2), + (REPEAT (ares_tac prems 1)) ]); + +val theI = prove_goal ZF.thy "EX! x. P(x) ==> P(THE x. P(x))" + (fn [major]=> + [ (rtac (major RS ex1E) 1), + (resolve_tac [major RS the_equality2 RS ssubst] 1), + (REPEAT (assume_tac 1)) ]); + +val the_cong = prove_goalw ZF.thy [the_def] + "[| !!y. P(y) <-> Q(y) |] ==> (THE x. P(x)) = (THE x. Q(x))" + (fn prems=> [ (prove_cong_tac (prems@[Replace_cong]) 1) ]); + + +(*** if -- a conditional expression for formulae ***) + +goalw ZF.thy [if_def] "if(True,a,b) = a"; +by (fast_tac (lemmas_cs addIs [the_equality]) 1); +val if_true = result(); + +goalw ZF.thy [if_def] "if(False,a,b) = b"; +by (fast_tac (lemmas_cs addIs [the_equality]) 1); +val if_false = result(); + +val prems = goalw ZF.thy [if_def] + "[| P<->Q; a=c; b=d |] ==> if(P,a,b) = if(Q,c,d)"; +by (REPEAT (resolve_tac (prems@[refl,the_cong]@FOL_congs) 1)); +val if_cong = result(); + +(*Not needed for rewriting, since P would rewrite to True anyway*) +val prems = goalw ZF.thy [if_def] "P ==> if(P,a,b) = a"; +by (cut_facts_tac prems 1); +by (fast_tac (lemmas_cs addSIs [the_equality]) 1); +val if_P = result(); + +(*Not needed for rewriting, since P would rewrite to False anyway*) +val prems = goalw ZF.thy [if_def] "~P ==> if(P,a,b) = b"; +by (cut_facts_tac prems 1); +by (fast_tac (lemmas_cs addSIs [the_equality]) 1); +val if_not_P = result(); + +val if_ss = + FOL_ss addcongs (if_cong :: mk_typed_congs ZF.thy [("P", "i=>o")] + @ basic_ZF_congs) + addrews [if_true,if_false]; + +val expand_if = prove_goal ZF.thy + "P(if(Q,x,y)) <-> ((Q --> P(x)) & (~Q --> P(y)))" + (fn _=> [ (res_inst_tac [("Q","Q")] (excluded_middle RS disjE) 1), + (ASM_SIMP_TAC if_ss 1), + (ASM_SIMP_TAC if_ss 1) ]); + +val prems = goal ZF.thy + "[| P ==> a: A; ~P ==> b: A |] ==> if(P,a,b): A"; +by (res_inst_tac [("Q","P")] (excluded_middle RS disjE) 1); +by (ALLGOALS (ASM_SIMP_TAC (if_ss addrews prems))); +val if_type = result(); + + +(*** Foundation lemmas ***) + +val mem_anti_sym = prove_goal ZF.thy "[| a:b; b:a |] ==> P" + (fn prems=> + [ (rtac disjE 1), + (res_inst_tac [("A","{a,b}")] foundation 1), + (etac equals0D 1), + (rtac consI1 1), + (fast_tac (lemmas_cs addIs (prems@[consI1,consI2]) + addSEs [consE,equalityE]) 1) ]); + +val mem_anti_refl = prove_goal ZF.thy "a:a ==> P" + (fn [major]=> [ (rtac (major RS (major RS mem_anti_sym)) 1) ]); + +val mem_not_refl = prove_goal ZF.thy "~ a:a" + (K [ (rtac notI 1), (etac mem_anti_refl 1) ]); + +(*** Rules for succ ***) + +val succI1 = prove_goalw ZF.thy [succ_def] "i : succ(i)" + (fn _=> [ (rtac consI1 1) ]); + +val succI2 = prove_goalw ZF.thy [succ_def] + "i : j ==> i : succ(j)" + (fn [prem]=> [ (rtac (prem RS consI2) 1) ]); + +(*Classical introduction rule*) +val succCI = prove_goalw ZF.thy [succ_def] + "(~ i:j ==> i=j) ==> i: succ(j)" + (fn prems=> [ (rtac consCI 1), (eresolve_tac prems 1) ]); + +val succE = prove_goalw ZF.thy [succ_def] + "[| i : succ(j); i=j ==> P; i:j ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS consE) 1), + (REPEAT (eresolve_tac prems 1)) ]); + +val succ_neq_0 = prove_goal ZF.thy "succ(n)=0 ==> P" + (fn [major]=> + [ (rtac (major RS equalityD1 RS subsetD RS emptyE) 1), + (rtac succI1 1) ]); + +(*Useful for rewriting*) +val succ_not_0 = prove_goal ZF.thy "~ succ(n)=0" + (fn _=> [ (rtac notI 1), (etac succ_neq_0 1) ]); + +(* succ(c) <= B ==> c : B *) +val succ_subsetD = succI1 RSN (2,subsetD); + +val succ_inject = prove_goal ZF.thy "succ(m) = succ(n) ==> m=n" + (fn [major]=> + [ (rtac (major RS equalityE) 1), + (REPEAT (eresolve_tac [asm_rl, sym, succE, make_elim succ_subsetD, + mem_anti_sym] 1)) ]); + +val succ_inject_iff = prove_goal ZF.thy "succ(m) = succ(n) <-> m=n" + (fn _=> [ (fast_tac (FOL_cs addSEs [succ_inject]) 1) ]); + +(*UpairI1/2 should become UpairCI; mem_anti_refl as a hazE? *) +val upair_cs = lemmas_cs + addSIs [singletonI, DiffI, IntI, UnCI, consCI, succCI, UpairI1,UpairI2] + addSEs [singletonE, DiffE, IntE, UnE, consE, succE, UpairE]; + diff -r 000000000000 -r a5a9c433f639 src/ZF/wf.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/wf.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,262 @@ +(* Title: ZF/wf.ML + ID: $Id$ + Author: Tobias Nipkow and Lawrence C Paulson + Copyright 1992 University of Cambridge + +For wf.thy. Well-founded Recursion + +Derived first for transitive relations, and finally for arbitrary WF relations +via wf_trancl and trans_trancl. + +It is difficult to derive this general case directly, using r^+ instead of +r. In is_recfun, the two occurrences of the relation must have the same +form. Inserting r^+ in the_recfun or wftrec yields a recursion rule with +r^+ -`` {a} instead of r-``{a}. This recursion rule is stronger in +principle, but harder to use, especially to prove wfrec_eclose_eq in +epsilon.ML. Expanding out the definition of wftrec in wfrec would yield +a mess. +*) + +open WF; + +val [H_cong] = mk_typed_congs WF.thy[("H","[i,i]=>i")]; + +val wf_ss = ZF_ss addcongs [H_cong]; + + +(*** Well-founded relations ***) + +(*Are these two theorems at all useful??*) + +(*If every subset of field(r) possesses an r-minimal element then wf(r). + Seems impossible to prove this for domain(r) or range(r) instead... + Consider in particular finite wf relations!*) +val [prem1,prem2] = goalw WF.thy [wf_def] + "[| field(r)<=A; \ +\ !!Z u. [| Z<=A; u:Z; ALL x:Z. EX y:Z. :r |] ==> False |] \ +\ ==> wf(r)"; +by (rtac (equals0I RS disjCI RS allI) 1); +by (rtac prem2 1); +by (res_inst_tac [ ("B1", "Z") ] (prem1 RS (Int_lower1 RS subset_trans)) 1); +by (fast_tac ZF_cs 1); +by (fast_tac ZF_cs 1); +val wfI = result(); + +(*If r allows well-founded induction then wf(r)*) +val [prem1,prem2] = goal WF.thy + "[| field(r)<=A; \ +\ !!B. ALL x:A. (ALL y. : r --> y:B) --> x:B ==> A<=B |] \ +\ ==> wf(r)"; +by (rtac (prem1 RS wfI) 1); +by (res_inst_tac [ ("B", "A-Z") ] (prem2 RS subsetCE) 1); +by (fast_tac ZF_cs 3); +by (fast_tac ZF_cs 2); +by (fast_tac ZF_cs 1); +val wfI2 = result(); + + +(** Well-founded Induction **) + +(*Consider the least z in domain(r) Un {a} such that P(z) does not hold...*) +val major::prems = goalw WF.thy [wf_def] + "[| wf(r); \ +\ !!x.[| ALL y. : r --> P(y) |] ==> P(x) \ +\ |] ==> P(a)"; +by (res_inst_tac [ ("x", "{z:domain(r) Un {a}. ~P(z)}") ] (major RS allE) 1); +by (etac disjE 1); +by (rtac classical 1); +by (etac equals0D 1); +by (etac (singletonI RS UnI2 RS CollectI) 1); +by (etac bexE 1); +by (etac CollectE 1); +by (etac swap 1); +by (resolve_tac prems 1); +by (fast_tac ZF_cs 1); +val wf_induct = result(); + +(*Perform induction on i, then prove the wf(r) subgoal using prems. *) +fun wf_ind_tac a prems i = + EVERY [res_inst_tac [("a",a)] wf_induct i, + rename_last_tac a ["1"] (i+1), + ares_tac prems i]; + +(*The form of this rule is designed to match wfI2*) +val wfr::amem::prems = goal WF.thy + "[| wf(r); a:A; field(r)<=A; \ +\ !!x.[| x: A; ALL y. : r --> P(y) |] ==> P(x) \ +\ |] ==> P(a)"; +by (rtac (amem RS rev_mp) 1); +by (wf_ind_tac "a" [wfr] 1); +by (rtac impI 1); +by (eresolve_tac prems 1); +by (fast_tac (ZF_cs addIs (prems RL [subsetD])) 1); +val wf_induct2 = result(); + +val prems = goal WF.thy "[| wf(r); :r; :r |] ==> False"; +by (subgoal_tac "ALL x. :r --> :r --> False" 1); +by (wf_ind_tac "a" prems 2); +by (fast_tac ZF_cs 2); +by (fast_tac (FOL_cs addIs prems) 1); +val wf_anti_sym = result(); + +(*transitive closure of a WF relation is WF!*) +val [prem] = goal WF.thy "wf(r) ==> wf(r^+)"; +by (rtac (trancl_type RS field_rel_subset RS wfI2) 1); +by (rtac subsetI 1); +(*must retain the universal formula for later use!*) +by (rtac (bspec RS mp) 1 THEN assume_tac 1 THEN assume_tac 1); +by (eres_inst_tac [("a","x")] (prem RS wf_induct2) 1); +by (rtac subset_refl 1); +by (rtac (impI RS allI) 1); +by (etac tranclE 1); +by (etac (bspec RS mp) 1); +by (etac fieldI1 1); +by (fast_tac ZF_cs 1); +by (fast_tac ZF_cs 1); +val wf_trancl = result(); + +(** r-``{a} is the set of everything under a in r **) + +val underI = standard (vimage_singleton_iff RS iffD2); +val underD = standard (vimage_singleton_iff RS iffD1); + +(** is_recfun **) + +val [major] = goalw WF.thy [is_recfun_def] + "is_recfun(r,a,H,f) ==> f: r-``{a} -> range(f)"; +by (rtac (major RS ssubst) 1); +by (rtac (lamI RS rangeI RS lam_type) 1); +by (assume_tac 1); +val is_recfun_type = result(); + +val [isrec,rel] = goalw WF.thy [is_recfun_def] + "[| is_recfun(r,a,H,f); :r |] ==> f`x = H(x, restrict(f,r-``{x}))"; +by (res_inst_tac [("P", "%x.?t(x) = ?u::i")] (isrec RS ssubst) 1); +by (rtac (rel RS underI RS beta) 1); +val apply_recfun = result(); + +(*eresolve_tac transD solves :r using transitivity AT MOST ONCE + spec RS mp instantiates induction hypotheses*) +fun indhyp_tac hyps = + ares_tac (TrueI::hyps) ORELSE' + (cut_facts_tac hyps THEN' + DEPTH_SOLVE_1 o (ares_tac [TrueI, ballI] ORELSE' + eresolve_tac [underD, transD, spec RS mp])); + +(*** NOTE! some simplifications need a different auto_tac!! ***) +val wf_super_ss = wf_ss setauto indhyp_tac; + +val prems = goalw WF.thy [is_recfun_def] + "[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g) |] ==> \ +\ :r --> :r --> f`x=g`x"; +by (cut_facts_tac prems 1); +by (wf_ind_tac "x" prems 1); +by (REPEAT (rtac impI 1 ORELSE etac ssubst 1)); +by (rewtac restrict_def); +by (ASM_SIMP_TAC (wf_super_ss addrews [vimage_singleton_iff]) 1); +val is_recfun_equal_lemma = result(); +val is_recfun_equal = standard (is_recfun_equal_lemma RS mp RS mp); + +val prems as [wfr,transr,recf,recg,_] = goal WF.thy + "[| wf(r); trans(r); \ +\ is_recfun(r,a,H,f); is_recfun(r,b,H,g); :r |] ==> \ +\ restrict(f, r-``{b}) = g"; +by (cut_facts_tac prems 1); +by (rtac (consI1 RS restrict_type RS fun_extension) 1); +by (etac is_recfun_type 1); +by (ALLGOALS + (ASM_SIMP_TAC (wf_super_ss addrews + [ [wfr,transr,recf,recg] MRS is_recfun_equal ]))); +val is_recfun_cut = result(); + +(*** Main Existence Lemma ***) + +val prems = goal WF.thy + "[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,a,H,g) |] ==> f=g"; +by (cut_facts_tac prems 1); +by (rtac fun_extension 1); +by (REPEAT (ares_tac [is_recfun_equal] 1 + ORELSE eresolve_tac [is_recfun_type,underD] 1)); +val is_recfun_functional = result(); + +(*If some f satisfies is_recfun(r,a,H,-) then so does the_recfun(r,a,H) *) +val prems = goalw WF.thy [the_recfun_def] + "[| is_recfun(r,a,H,f); wf(r); trans(r) |] \ +\ ==> is_recfun(r, a, H, the_recfun(r,a,H))"; +by (rtac (ex1I RS theI) 1); +by (REPEAT (ares_tac (prems@[is_recfun_functional]) 1)); +val is_the_recfun = result(); + +val prems = goal WF.thy + "[| wf(r); trans(r) |] ==> is_recfun(r, a, H, the_recfun(r,a,H))"; +by (cut_facts_tac prems 1); +by (wf_ind_tac "a" prems 1); +by (res_inst_tac [("f", "lam y: r-``{a1}. wftrec(r,y,H)")] is_the_recfun 1); +by (REPEAT (assume_tac 2)); +by (rewrite_goals_tac [is_recfun_def, wftrec_def]); +(*Applying the substitution: must keep the quantified assumption!!*) +by (REPEAT (dtac underD 1 ORELSE resolve_tac [refl, lam_cong, H_cong] 1)); +by (fold_tac [is_recfun_def]); +by (rtac (consI1 RS restrict_type RSN (2,fun_extension)) 1); +by (rtac is_recfun_type 1); +by (ALLGOALS + (ASM_SIMP_TAC + (wf_super_ss addrews [underI RS beta, apply_recfun, is_recfun_cut]))); +val unfold_the_recfun = result(); + + +(*** Unfolding wftrec ***) + +val prems = goal WF.thy + "[| wf(r); trans(r); :r |] ==> \ +\ restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)"; +by (REPEAT (ares_tac (prems @ [is_recfun_cut, unfold_the_recfun]) 1)); +val the_recfun_cut = result(); + +(*NOT SUITABLE FOR REWRITING since it is recursive!*) +val prems = goalw WF.thy [wftrec_def] + "[| wf(r); trans(r) |] ==> \ +\ wftrec(r,a,H) = H(a, lam x: r-``{a}. wftrec(r,x,H))"; +by (rtac (rewrite_rule [is_recfun_def] unfold_the_recfun RS ssubst) 1); +by (ALLGOALS (ASM_SIMP_TAC + (wf_ss addrews (prems@[vimage_singleton_iff RS iff_sym, + the_recfun_cut])))); +val wftrec = result(); + +(** Removal of the premise trans(r) **) + +(*NOT SUITABLE FOR REWRITING since it is recursive!*) +val [wfr] = goalw WF.thy [wfrec_def] + "wf(r) ==> wfrec(r,a,H) = H(a, lam x:r-``{a}. wfrec(r,x,H))"; +by (rtac (wfr RS wf_trancl RS wftrec RS ssubst) 1); +by (rtac trans_trancl 1); +by (rtac (refl RS H_cong) 1); +by (rtac (vimage_pair_mono RS restrict_lam_eq) 1); +by (etac r_into_trancl 1); +by (rtac subset_refl 1); +val wfrec = result(); + +(*This form avoids giant explosions in proofs. NOTE USE OF == *) +val rew::prems = goal WF.thy + "[| !!x. h(x)==wfrec(r,x,H); wf(r) |] ==> \ +\ h(a) = H(a, lam x: r-``{a}. h(x))"; +by (rewtac rew); +by (REPEAT (resolve_tac (prems@[wfrec]) 1)); +val def_wfrec = result(); + +val prems = goal WF.thy + "[| wf(r); a:A; field(r)<=A; \ +\ !!x u. [| x: A; u: Pi(r-``{x}, B) |] ==> H(x,u) : B(x) \ +\ |] ==> wfrec(r,a,H) : B(a)"; +by (res_inst_tac [("a","a")] wf_induct2 1); +by (rtac (wfrec RS ssubst) 4); +by (REPEAT (ares_tac (prems@[lam_type]) 1 + ORELSE eresolve_tac [spec RS mp, underD] 1)); +val wfrec_type = result(); + +val prems = goalw WF.thy [wfrec_def,wftrec_def,the_recfun_def,is_recfun_def] + "[| r=r'; !!x u. H(x,u)=H'(x,u); a=a' |] \ +\ ==> wfrec(r,a,H)=wfrec(r',a',H')"; +by (EVERY1 (map rtac (prems RL [subst]))); +by (SIMP_TAC (wf_ss addrews (prems RL [sym])) 1); +val wfrec_cong = result(); diff -r 000000000000 -r a5a9c433f639 src/ZF/wf.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/wf.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,30 @@ +(* Title: ZF/wf.thy + ID: $Id$ + Author: Tobias Nipkow and Lawrence C Paulson + Copyright 1992 University of Cambridge + +Well-founded Recursion +*) + +WF = Trancl + +consts + wf :: "i=>o" + wftrec,wfrec :: "[i, i, [i,i]=>i] =>i" + is_recfun :: "[i, i, [i,i]=>i, i] =>o" + the_recfun :: "[i, i, [i,i]=>i] =>i" + +rules + (*r is a well-founded relation*) + wf_def "wf(r) == ALL Z. Z=0 | (EX x:Z. ALL y. :r --> ~ y:Z)" + + is_recfun_def "is_recfun(r,a,H,f) == \ +\ (f = (lam x: r-``{a}. H(x, restrict(f, r-``{x}))))" + + the_recfun_def "the_recfun(r,a,H) == (THE f.is_recfun(r,a,H,f))" + + wftrec_def "wftrec(r,a,H) == H(a, the_recfun(r,a,H))" + + (*public version. Does not require r to be transitive*) + wfrec_def "wfrec(r,a,H) == wftrec(r^+, a, %x f. H(x, restrict(f,r-``{x})))" + +end diff -r 000000000000 -r a5a9c433f639 src/ZF/zf.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/zf.ML Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,455 @@ +(* Title: ZF/zf.ML + ID: $Id$ + Author: Lawrence C Paulson and Martin D Coen, CU Computer Laboratory + Copyright 1992 University of Cambridge + +Basic introduction and elimination rules for Zermelo-Fraenkel Set Theory +*) + +open ZF; + +signature ZF_LEMMAS = + sig + val ballE : thm + val ballI : thm + val ball_cong : thm + val ball_rew : thm + val ball_tac : int -> tactic + val basic_ZF_congs : thm list + val bexCI : thm + val bexE : thm + val bexI : thm + val bex_cong : thm + val bspec : thm + val CollectD1 : thm + val CollectD2 : thm + val CollectE : thm + val CollectI : thm + val Collect_cong : thm + val emptyE : thm + val empty_subsetI : thm + val equalityCE : thm + val equalityD1 : thm + val equalityD2 : thm + val equalityE : thm + val equalityI : thm + val equality_iffI : thm + val equals0D : thm + val equals0I : thm + val ex1_functional : thm + val InterD : thm + val InterE : thm + val InterI : thm + val INT_E : thm + val INT_I : thm + val lemmas_cs : claset + val PowD : thm + val PowI : thm + val prove_cong_tac : thm list -> int -> tactic + val RepFunE : thm + val RepFunI : thm + val RepFun_eqI : thm + val RepFun_cong : thm + val ReplaceE : thm + val ReplaceI : thm + val Replace_iff : thm + val Replace_cong : thm + val rev_ballE : thm + val rev_bspec : thm + val rev_subsetD : thm + val separation : thm + val setup_induction : thm + val set_mp_tac : int -> tactic + val subsetCE : thm + val subsetD : thm + val subsetI : thm + val subset_refl : thm + val subset_trans : thm + val UnionE : thm + val UnionI : thm + val UN_E : thm + val UN_I : thm + end; + + +structure ZF_Lemmas : ZF_LEMMAS = +struct + +val basic_ZF_congs = mk_congs ZF.thy + ["op `", "op ``", "op Int", "op Un", "op -", "op <=", "op :", + "Pow", "Union", "Inter", "fst", "snd", "succ", "Pair", "Upair", "cons", + "domain", "range", "restrict"]; + +fun prove_cong_tac prems i = + REPEAT (ares_tac (prems@[refl]@FOL_congs@basic_ZF_congs) i); + +(*** Bounded universal quantifier ***) + +val ballI = prove_goalw ZF.thy [Ball_def] + "[| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)" + (fn prems=> [ (REPEAT (ares_tac (prems @ [allI,impI]) 1)) ]); + +val bspec = prove_goalw ZF.thy [Ball_def] + "[| ALL x:A. P(x); x: A |] ==> P(x)" + (fn major::prems=> + [ (rtac (major RS spec RS mp) 1), + (resolve_tac prems 1) ]); + +val ballE = prove_goalw ZF.thy [Ball_def] + "[| ALL x:A. P(x); P(x) ==> Q; ~ x:A ==> Q |] ==> Q" + (fn major::prems=> + [ (rtac (major RS allE) 1), + (REPEAT (eresolve_tac (prems@[asm_rl,impCE]) 1)) ]); + +(*Used in the datatype package*) +val rev_bspec = prove_goal ZF.thy + "!!x A P. [| x: A; ALL x:A. P(x) |] ==> P(x)" + (fn _ => + [ REPEAT (ares_tac [bspec] 1) ]); + +(*Instantiates x first: better for automatic theorem proving?*) +val rev_ballE = prove_goal ZF.thy + "[| ALL x:A. P(x); ~ x:A ==> Q; P(x) ==> Q |] ==> Q" + (fn major::prems=> + [ (rtac (major RS ballE) 1), + (REPEAT (eresolve_tac prems 1)) ]); + +(*Takes assumptions ALL x:A.P(x) and a:A; creates assumption P(a)*) +val ball_tac = dtac bspec THEN' assume_tac; + +(*Trival rewrite rule; (ALL x:A.P)<->P holds only if A is nonempty!*) +val ball_rew = prove_goal ZF.thy "(ALL x:A. True) <-> True" + (fn prems=> [ (REPEAT (ares_tac [TrueI,ballI,iffI] 1)) ]); + +(*Congruence rule for rewriting*) +val ball_cong = prove_goalw ZF.thy [Ball_def] + "[| A=A'; !!x. x:A' ==> P(x) <-> P'(x) \ +\ |] ==> (ALL x:A. P(x)) <-> (ALL x:A'. P'(x))" + (fn prems=> [ (prove_cong_tac prems 1) ]); + +(*** Bounded existential quantifier ***) + +val bexI = prove_goalw ZF.thy [Bex_def] + "[| P(x); x: A |] ==> EX x:A. P(x)" + (fn prems=> [ (REPEAT (ares_tac (prems @ [exI,conjI]) 1)) ]); + +(*Not of the general form for such rules; ~EX has become ALL~ *) +val bexCI = prove_goal ZF.thy + "[| ALL x:A. ~P(x) ==> P(a); a: A |] ==> EX x:A.P(x)" + (fn prems=> + [ (rtac classical 1), + (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ]); + +val bexE = prove_goalw ZF.thy [Bex_def] + "[| EX x:A. P(x); !!x. [| x:A; P(x) |] ==> Q \ +\ |] ==> Q" + (fn major::prems=> + [ (rtac (major RS exE) 1), + (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)) ]); + +(*We do not even have (EX x:A. True) <-> True unless A is nonempty!!*) + +val bex_cong = prove_goalw ZF.thy [Bex_def] + "[| A=A'; !!x. x:A' ==> P(x) <-> P'(x) \ +\ |] ==> (EX x:A. P(x)) <-> (EX x:A'. P'(x))" + (fn prems=> [ (prove_cong_tac prems 1) ]); + +(*** Rules for subsets ***) + +val subsetI = prove_goalw ZF.thy [subset_def] + "(!!x.x:A ==> x:B) ==> A <= B" + (fn prems=> [ (REPEAT (ares_tac (prems @ [ballI]) 1)) ]); + +(*Rule in Modus Ponens style [was called subsetE] *) +val subsetD = prove_goalw ZF.thy [subset_def] "[| A <= B; c:A |] ==> c:B" + (fn major::prems=> + [ (rtac (major RS bspec) 1), + (resolve_tac prems 1) ]); + +(*Classical elimination rule*) +val subsetCE = prove_goalw ZF.thy [subset_def] + "[| A <= B; ~(c:A) ==> P; c:B ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS ballE) 1), + (REPEAT (eresolve_tac prems 1)) ]); + +(*Takes assumptions A<=B; c:A and creates the assumption c:B *) +val set_mp_tac = dtac subsetD THEN' assume_tac; + +(*Sometimes useful with premises in this order*) +val rev_subsetD = prove_goal ZF.thy "!!A B c. [| c:A; A<=B |] ==> c:B" + (fn _=> [REPEAT (ares_tac [subsetD] 1)]); + +val subset_refl = prove_goal ZF.thy "A <= A" + (fn _=> [ (rtac subsetI 1), atac 1 ]); + +val subset_trans = prove_goal ZF.thy "[| A<=B; B<=C |] ==> A<=C" + (fn prems=> [ (REPEAT (ares_tac ([subsetI]@(prems RL [subsetD])) 1)) ]); + + +(*** Rules for equality ***) + +(*Anti-symmetry of the subset relation*) +val equalityI = prove_goal ZF.thy "[| A <= B; B <= A |] ==> A = B" + (fn prems=> [ (REPEAT (resolve_tac (prems@[conjI, extension RS iffD2]) 1)) ]); + +val equality_iffI = prove_goal ZF.thy "(!!x. x:A <-> x:B) ==> A = B" + (fn [prem] => + [ (rtac equalityI 1), + (REPEAT (ares_tac [subsetI, prem RS iffD1, prem RS iffD2] 1)) ]); + +val equalityD1 = prove_goal ZF.thy "A = B ==> A<=B" + (fn prems=> + [ (rtac (extension RS iffD1 RS conjunct1) 1), + (resolve_tac prems 1) ]); + +val equalityD2 = prove_goal ZF.thy "A = B ==> B<=A" + (fn prems=> + [ (rtac (extension RS iffD1 RS conjunct2) 1), + (resolve_tac prems 1) ]); + +val equalityE = prove_goal ZF.thy + "[| A = B; [| A<=B; B<=A |] ==> P |] ==> P" + (fn prems=> + [ (DEPTH_SOLVE (resolve_tac (prems@[equalityD1,equalityD2]) 1)) ]); + +val equalityCE = prove_goal ZF.thy + "[| A = B; [| c:A; c:B |] ==> P; [| ~ c:A; ~ c:B |] ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS equalityE) 1), + (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1)) ]); + +(*Lemma for creating induction formulae -- for "pattern matching" on p + To make the induction hypotheses usable, apply "spec" or "bspec" to + put universal quantifiers over the free variables in p. + Would it be better to do subgoal_tac "ALL z. p = f(z) --> R(z)" ??*) +val setup_induction = prove_goal ZF.thy + "[| p: A; !!z. z: A ==> p=z --> R |] ==> R" + (fn prems=> + [ (rtac mp 1), + (REPEAT (resolve_tac (refl::prems) 1)) ]); + + +(*** Rules for Replace -- the derived form of replacement ***) + +val ex1_functional = prove_goal ZF.thy + "[| EX! z. P(a,z); P(a,b); P(a,c) |] ==> b = c" + (fn prems=> + [ (cut_facts_tac prems 1), + (best_tac FOL_dup_cs 1) ]); + +val Replace_iff = prove_goalw ZF.thy [Replace_def] + "b : {y. x:A, P(x,y)} <-> (EX x:A. P(x,b) & (ALL y. P(x,y) --> y=b))" + (fn _=> + [ (rtac (replacement RS iff_trans) 1), + (REPEAT (ares_tac [refl,bex_cong,iffI,ballI,allI,conjI,impI,ex1I] 1 + ORELSE eresolve_tac [conjE, spec RS mp, ex1_functional] 1)) ]); + +(*Introduction; there must be a unique y such that P(x,y), namely y=b. *) +val ReplaceI = prove_goal ZF.thy + "[| x: A; P(x,b); !!y. P(x,y) ==> y=b |] ==> \ +\ b : {y. x:A, P(x,y)}" + (fn prems=> + [ (rtac (Replace_iff RS iffD2) 1), + (REPEAT (ares_tac (prems@[bexI,conjI,allI,impI]) 1)) ]); + +(*Elimination; may asssume there is a unique y such that P(x,y), namely y=b. *) +val ReplaceE = prove_goal ZF.thy + "[| b : {y. x:A, P(x,y)}; \ +\ !!x. [| x: A; P(x,b); ALL y. P(x,y)-->y=b |] ==> R \ +\ |] ==> R" + (fn prems=> + [ (rtac (Replace_iff RS iffD1 RS bexE) 1), + (etac conjE 2), + (REPEAT (ares_tac prems 1)) ]); + +val Replace_cong = prove_goal ZF.thy + "[| A=B; !!x y. x:B ==> P(x,y) <-> Q(x,y) |] ==> \ +\ {y. x:A, P(x,y)} = {y. x:B, Q(x,y)}" + (fn prems=> + let val substprems = prems RL [subst, ssubst] + and iffprems = prems RL [iffD1,iffD2] + in [ (rtac equalityI 1), + (REPEAT (eresolve_tac (substprems@[asm_rl, ReplaceE, spec RS mp]) 1 + ORELSE resolve_tac [subsetI, ReplaceI] 1 + ORELSE (resolve_tac iffprems 1 THEN assume_tac 2))) ] + end); + + +(*** Rules for RepFun ***) + +val RepFunI = prove_goalw ZF.thy [RepFun_def] + "!!a A. a : A ==> f(a) : {f(x). x:A}" + (fn _ => [ (REPEAT (ares_tac [ReplaceI,refl] 1)) ]); + +(*Useful for co-induction proofs*) +val RepFun_eqI = prove_goal ZF.thy + "!!b a f. [| b=f(a); a : A |] ==> b : {f(x). x:A}" + (fn _ => [ etac ssubst 1, etac RepFunI 1 ]); + +val RepFunE = prove_goalw ZF.thy [RepFun_def] + "[| b : {f(x). x:A}; \ +\ !!x.[| x:A; b=f(x) |] ==> P |] ==> \ +\ P" + (fn major::prems=> + [ (rtac (major RS ReplaceE) 1), + (REPEAT (ares_tac prems 1)) ]); + +val RepFun_cong = prove_goalw ZF.thy [RepFun_def] + "[| A=B; !!x. x:B ==> f(x)=g(x) |] ==> \ +\ {f(x). x:A} = {g(x). x:B}" + (fn prems=> [ (prove_cong_tac (prems@[Replace_cong]) 1) ]); + + +(*** Rules for Collect -- forming a subset by separation ***) + +(*Separation is derivable from Replacement*) +val separation = prove_goalw ZF.thy [Collect_def] + "a : {x:A. P(x)} <-> a:A & P(a)" + (fn _=> [ (fast_tac (FOL_cs addIs [bexI,ReplaceI] + addSEs [bexE,ReplaceE]) 1) ]); + +val CollectI = prove_goal ZF.thy + "[| a:A; P(a) |] ==> a : {x:A. P(x)}" + (fn prems=> + [ (rtac (separation RS iffD2) 1), + (REPEAT (resolve_tac (prems@[conjI]) 1)) ]); + +val CollectE = prove_goal ZF.thy + "[| a : {x:A. P(x)}; [| a:A; P(a) |] ==> R |] ==> R" + (fn prems=> + [ (rtac (separation RS iffD1 RS conjE) 1), + (REPEAT (ares_tac prems 1)) ]); + +val CollectD1 = prove_goal ZF.thy "a : {x:A. P(x)} ==> a:A" + (fn [major]=> + [ (rtac (major RS CollectE) 1), + (assume_tac 1) ]); + +val CollectD2 = prove_goal ZF.thy "a : {x:A. P(x)} ==> P(a)" + (fn [major]=> + [ (rtac (major RS CollectE) 1), + (assume_tac 1) ]); + +val Collect_cong = prove_goalw ZF.thy [Collect_def] + "[| A=B; !!x. x:B ==> P(x) <-> Q(x) |] ==> \ +\ {x:A. P(x)} = {x:B. Q(x)}" + (fn prems=> [ (prove_cong_tac (prems@[Replace_cong]) 1) ]); + +(*** Rules for Unions ***) + +(*The order of the premises presupposes that C is rigid; A may be flexible*) +val UnionI = prove_goal ZF.thy "[| B: C; A: B |] ==> A: Union(C)" + (fn prems=> + [ (resolve_tac [union_iff RS iffD2] 1), + (REPEAT (resolve_tac (prems @ [bexI]) 1)) ]); + +val UnionE = prove_goal ZF.thy + "[| A : Union(C); !!B.[| A: B; B: C |] ==> R |] ==> R" + (fn prems=> + [ (resolve_tac [union_iff RS iffD1 RS bexE] 1), + (REPEAT (ares_tac prems 1)) ]); + +(*** Rules for Inter ***) + +(*Not obviously useful towards proving InterI, InterD, InterE*) +val Inter_iff = prove_goalw ZF.thy [Inter_def,Ball_def] + "A : Inter(C) <-> (ALL x:C. A: x) & (EX x. x:C)" + (fn _=> [ (rtac (separation RS iff_trans) 1), + (fast_tac (FOL_cs addIs [UnionI] addSEs [UnionE]) 1) ]); + +(* Intersection is well-behaved only if the family is non-empty! *) +val InterI = prove_goalw ZF.thy [Inter_def] + "[| !!x. x: C ==> A: x; c:C |] ==> A : Inter(C)" + (fn prems=> + [ (DEPTH_SOLVE (ares_tac ([CollectI,UnionI,ballI] @ prems) 1)) ]); + +(*A "destruct" rule -- every B in C contains A as an element, but + A:B can hold when B:C does not! This rule is analogous to "spec". *) +val InterD = prove_goalw ZF.thy [Inter_def] + "[| A : Inter(C); B : C |] ==> A : B" + (fn [major,minor]=> + [ (rtac (major RS CollectD2 RS bspec) 1), + (rtac minor 1) ]); + +(*"Classical" elimination rule -- does not require exhibiting B:C *) +val InterE = prove_goalw ZF.thy [Inter_def] + "[| A : Inter(C); A:B ==> R; ~ B:C ==> R |] ==> R" + (fn major::prems=> + [ (rtac (major RS CollectD2 RS ballE) 1), + (REPEAT (eresolve_tac prems 1)) ]); + +(*** Rules for Unions of families ***) +(* UN x:A. B(x) abbreviates Union({B(x). x:A}) *) + +(*The order of the premises presupposes that A is rigid; b may be flexible*) +val UN_I = prove_goal ZF.thy "[| a: A; b: B(a) |] ==> b: (UN x:A. B(x))" + (fn prems=> + [ (REPEAT (resolve_tac (prems@[UnionI,RepFunI]) 1)) ]); + +val UN_E = prove_goal ZF.thy + "[| b : (UN x:A. B(x)); !!x.[| x: A; b: B(x) |] ==> R |] ==> R" + (fn major::prems=> + [ (rtac (major RS UnionE) 1), + (REPEAT (eresolve_tac (prems@[asm_rl, RepFunE, subst]) 1)) ]); + + +(*** Rules for Intersections of families ***) +(* INT x:A. B(x) abbreviates Inter({B(x). x:A}) *) + +val INT_I = prove_goal ZF.thy + "[| !!x. x: A ==> b: B(x); a: A |] ==> b: (INT x:A. B(x))" + (fn prems=> + [ (REPEAT (ares_tac (prems@[InterI,RepFunI]) 1 + ORELSE eresolve_tac [RepFunE,ssubst] 1)) ]); + +val INT_E = prove_goal ZF.thy + "[| b : (INT x:A. B(x)); a: A |] ==> b : B(a)" + (fn [major,minor]=> + [ (rtac (major RS InterD) 1), + (rtac (minor RS RepFunI) 1) ]); + + +(*** Rules for Powersets ***) + +val PowI = prove_goal ZF.thy "A <= B ==> A : Pow(B)" + (fn [prem]=> [ (rtac (prem RS (power_set RS iffD2)) 1) ]); + +val PowD = prove_goal ZF.thy "A : Pow(B) ==> A<=B" + (fn [major]=> [ (rtac (major RS (power_set RS iffD1)) 1) ]); + + +(*** Rules for the empty set ***) + +(*The set {x:0.False} is empty; by foundation it equals 0 + See Suppes, page 21.*) +val emptyE = prove_goal ZF.thy "a:0 ==> P" + (fn [major]=> + [ (rtac (foundation RS disjE) 1), + (etac (equalityD2 RS subsetD RS CollectD2 RS FalseE) 1), + (rtac major 1), + (etac bexE 1), + (etac (CollectD2 RS FalseE) 1) ]); + +val empty_subsetI = prove_goal ZF.thy "0 <= A" + (fn _ => [ (REPEAT (ares_tac [equalityI,subsetI,emptyE] 1)) ]); + +val equals0I = prove_goal ZF.thy "[| !!y. y:A ==> False |] ==> A=0" + (fn prems=> + [ (REPEAT (ares_tac (prems@[empty_subsetI,subsetI,equalityI]) 1 + ORELSE eresolve_tac (prems RL [FalseE]) 1)) ]); + +val equals0D = prove_goal ZF.thy "[| A=0; a:A |] ==> P" + (fn [major,minor]=> + [ (rtac (minor RS (major RS equalityD1 RS subsetD RS emptyE)) 1) ]); + +val lemmas_cs = FOL_cs + addSIs [ballI, InterI, CollectI, PowI, subsetI] + addIs [bexI, UnionI, ReplaceI, RepFunI] + addSEs [bexE, make_elim PowD, UnionE, ReplaceE, RepFunE, + CollectE, emptyE] + addEs [rev_ballE, InterD, make_elim InterD, subsetD, subsetCE]; + +end; + +open ZF_Lemmas; diff -r 000000000000 -r a5a9c433f639 src/ZF/zf.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/zf.thy Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,227 @@ +(* Title: ZF/zf.thy + ID: $Id$ + Author: Lawrence C Paulson and Martin D Coen, CU Computer Laboratory + Copyright 1993 University of Cambridge + +Zermelo-Fraenkel Set Theory +*) + +ZF = FOL + + +types + i, is, syntax 0 + +arities + i :: term + + +consts + + "0" :: "i" ("0") (*the empty set*) + Pow :: "i => i" (*power sets*) + Inf :: "i" (*infinite set*) + + (* Bounded Quantifiers *) + + "@Ball" :: "[idt, i, o] => o" ("(3ALL _:_./ _)" 10) + "@Bex" :: "[idt, i, o] => o" ("(3EX _:_./ _)" 10) + Ball :: "[i, i => o] => o" + Bex :: "[i, i => o] => o" + + (* General Union and Intersection *) + + "@INTER" :: "[idt, i, i] => i" ("(3INT _:_./ _)" 10) + "@UNION" :: "[idt, i, i] => i" ("(3UN _:_./ _)" 10) + Union, Inter :: "i => i" + + (* Variations on Replacement *) + + "@Replace" :: "[idt, idt, i, o] => i" ("(1{_ ./ _: _, _})") + "@RepFun" :: "[i, idt, i] => i" ("(1{_ ./ _: _})") + "@Collect" :: "[idt, i, o] => i" ("(1{_: _ ./ _})") + PrimReplace :: "[i, [i, i] => o] => i" + Replace :: "[i, [i, i] => o] => i" + RepFun :: "[i, i => i] => i" + Collect :: "[i, i => o] => i" + + (* Descriptions *) + + "@THE" :: "[idt, o] => i" ("(3THE _./ _)" 10) + The :: "[i => o] => i" + if :: "[o, i, i] => i" + + (* Enumerations of type i *) + + "" :: "i => is" ("_") + "@Enum" :: "[i, is] => is" ("_,/ _") + + (* Finite Sets *) + + "@Finset" :: "is => i" ("{(_)}") + Upair, cons :: "[i, i] => i" + succ :: "i => i" + + (* Ordered Pairing and n-Tuples *) + + "@Tuple" :: "[i, is] => i" ("<(_,/ _)>") + PAIR :: "syntax" + Pair :: "[i, i] => i" + fst, snd :: "i => i" + split :: "[[i,i] => i, i] => i" + fsplit :: "[[i,i] => o, i] => o" + + (* Sigma and Pi Operators *) + + "@PROD" :: "[idt, i, i] => i" ("(3PROD _:_./ _)" 10) + "@SUM" :: "[idt, i, i] => i" ("(3SUM _:_./ _)" 10) + "@lam" :: "[idt, i, i] => i" ("(3lam _:_./ _)" 10) + Pi, Sigma :: "[i, i => i] => i" + + (* Relations and Functions *) + + domain :: "i => i" + range :: "i => i" + field :: "i => i" + converse :: "i => i" + Lambda :: "[i, i => i] => i" + restrict :: "[i, i] => i" + + (* Infixes in order of decreasing precedence *) + + "``" :: "[i, i] => i" (infixl 90) (*image*) + "-``" :: "[i, i] => i" (infixl 90) (*inverse image*) + "`" :: "[i, i] => i" (infixl 90) (*function application*) + + (*Except for their translations, * and -> are right-associating infixes*) + " *" :: "[i, i] => i" ("(_ */ _)" [81, 80] 80) (*Cartesian product*) + "Int" :: "[i, i] => i" (infixl 70) (*binary intersection*) + "Un" :: "[i, i] => i" (infixl 65) (*binary union*) + "-" :: "[i, i] => i" (infixl 65) (*set difference*) + " ->" :: "[i, i] => i" ("(_ ->/ _)" [61, 60] 60) (*function space*) + "<=" :: "[i, i] => o" (infixl 50) (*subset relation*) + ":" :: "[i, i] => o" (infixl 50) (*membership relation*) + + +translations + "{x, xs}" == "cons(x, {xs})" + "{x}" == "cons(x, 0)" + + "PAIR(x, Pair(y, z))" <= "Pair(x, Pair(y, z))" + "PAIR(x, PAIR(y, z))" <= "Pair(x, PAIR(y, z))" + "" <= "PAIR(x, )" + "" == "Pair(x, )" + "" == "Pair(x, y)" + + "{x:A. P}" == "Collect(A, %x. P)" + "{y. x:A, Q}" == "Replace(A, %x y. Q)" + "{f. x:A}" == "RepFun(A, %x. f)" + "INT x:A. B" == "Inter({B. x:A})" + "UN x:A. B" == "Union({B. x:A})" + "PROD x:A. B" => "Pi(A, %x. B)" + "SUM x:A. B" => "Sigma(A, %x. B)" + "THE x. P" == "The(%x. P)" + "lam x:A. f" == "Lambda(A, %x. f)" + "ALL x:A. P" == "Ball(A, %x. P)" + "EX x:A. P" == "Bex(A, %x. P)" + + +rules + + (* Bounded Quantifiers *) +Ball_def "Ball(A,P) == ALL x. x:A --> P(x)" +Bex_def "Bex(A,P) == EX x. x:A & P(x)" +subset_def "A <= B == ALL x:A. x:B" + + (* ZF axioms -- see Suppes p.238 + Axioms for Union, Pow and Replace state existence only, + uniqueness is derivable using extensionality. *) + +extension "A = B <-> A <= B & B <= A" +union_iff "A : Union(C) <-> (EX B:C. A:B)" +power_set "A : Pow(B) <-> A <= B" +succ_def "succ(i) == cons(i,i)" + + (*We may name this set, though it is not uniquely defined. *) +infinity "0:Inf & (ALL y:Inf. succ(y): Inf)" + + (*This formulation facilitates case analysis on A. *) +foundation "A=0 | (EX x:A. ALL y:x. ~ y:A)" + + (* Schema axiom since predicate P is a higher-order variable *) +replacement "(ALL x:A. ALL y z. P(x,y) & P(x,z) --> y=z) ==> \ +\ b : PrimReplace(A,P) <-> (EX x:A. P(x,b))" + + (* Derived form of replacement, restricting P to its functional part. + The resulting set (for functional P) is the same as with + PrimReplace, but the rules are simpler. *) +Replace_def "Replace(A,P) == PrimReplace(A, %x y. (EX!z.P(x,z)) & P(x,y))" + + (* Functional form of replacement -- analgous to ML's map functional *) +RepFun_def "RepFun(A,f) == {y . x:A, y=f(x)}" + + (* Separation and Pairing can be derived from the Replacement + and Powerset Axioms using the following definitions. *) + +Collect_def "Collect(A,P) == {y . x:A, x=y & P(x)}" + + (*Unordered pairs (Upair) express binary union/intersection and cons; + set enumerations translate as {a,...,z} = cons(a,...,cons(z,0)...) *) +Upair_def "Upair(a,b) == {y. x:Pow(Pow(0)), (x=0 & y=a) | (x=Pow(0) & y=b)}" +cons_def "cons(a,A) == Upair(a,a) Un A" + + (* Difference, general intersection, binary union and small intersection *) + +Diff_def "A - B == { x:A . ~(x:B) }" +Inter_def "Inter(A) == { x:Union(A) . ALL y:A. x:y}" +Un_def "A Un B == Union(Upair(A,B))" +Int_def "A Int B == Inter(Upair(A,B))" + + (* Definite descriptions -- via Replace over the set "1" *) + +the_def "The(P) == Union({y . x:{0}, P(y)})" +if_def "if(P,a,b) == THE z. P & z=a | ~P & z=b" + + (* Ordered pairs and disjoint union of a family of sets *) + + (* this "symmetric" definition works better than {{a}, {a,b}} *) +Pair_def " == {{a,a}, {a,b}}" +fst_def "fst == split(%x y.x)" +snd_def "snd == split(%x y.y)" +split_def "split(c,p) == THE y. EX a b. p= & y=c(a,b)" +fsplit_def "fsplit(R,z) == EX x y. z= & R(x,y)" +Sigma_def "Sigma(A,B) == UN x:A. UN y:B(x). {}" + + (* Operations on relations *) + +(*converse of relation r, inverse of function*) +converse_def "converse(r) == {z. w:r, EX x y. w= & z=}" + +domain_def "domain(r) == {x. w:r, EX y. w=}" +range_def "range(r) == domain(converse(r))" +field_def "field(r) == domain(r) Un range(r)" +image_def "r `` A == {y : range(r) . EX x:A. : r}" +vimage_def "r -`` A == converse(r)``A" + + (* Abstraction, application and Cartesian product of a family of sets *) + +lam_def "Lambda(A,b) == { . x:A}" +apply_def "f`a == THE y. : f" +Pi_def "Pi(A,B) == {f: Pow(Sigma(A,B)). ALL x:A. EX! y. : f}" + + (* Restrict the function f to the domain A *) +restrict_def "restrict(f,A) == lam x:A.f`x" + +end + + +ML + +(* 'Dependent' type operators *) + +val parse_translation = + [(" ->", ndependent_tr "Pi"), + (" *", ndependent_tr "Sigma")]; + +val print_translation = + [("Pi", dependent_tr' ("@PROD", " ->")), + ("Sigma", dependent_tr' ("@SUM", " *"))]; diff -r 000000000000 -r a5a9c433f639 teeinput --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/teeinput Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,4 @@ +#! /bin/sh +# teeinput -- start a program and log all inputs to a file +# environment variable $LISTEN specifies the file name +tee -a -i $LISTEN | $* \ No newline at end of file diff -r 000000000000 -r a5a9c433f639 xlisten --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/xlisten Thu Sep 16 12:20:38 1993 +0200 @@ -0,0 +1,10 @@ +#! /bin/sh +# xlisten -- start a program in one window and create a listener window +# environment variable $LISTEN specifies the file name + +#create the file! +date > $LISTEN + +xterm -geo 80x10+0+0 -T Listener -n Listener -e tail -f $LISTEN & +sleep 2 +xterm -geo 80x45+0-0 -e teeinput $* &