| author | wenzelm | 
| Sun, 01 Oct 2017 20:50:26 +0200 | |
| changeset 66744 | fec1504e5f03 | 
| parent 63435 | 7743df69a6b4 | 
| permissions | -rw-r--r-- | 
| 26189 | 1 | (* Title: ZF/Inductive_ZF.thy | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 3 | Copyright 1993 University of Cambridge | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 4 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 5 | Inductive definitions use least fixedpoints with standard products and sums | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 6 | Coinductive definitions use greatest fixedpoints with Quine products and sums | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 7 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 8 | Sums are used only for mutual recursion; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 9 | Products are used only to derive "streamlined" induction rules for relations | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 10 | *) | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 11 | |
| 60770 | 12 | section\<open>Inductive and Coinductive Definitions\<close> | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 13 | |
| 26189 | 14 | theory Inductive_ZF | 
| 15 | imports Fixedpt QPair Nat_ZF | |
| 46947 | 16 | keywords | 
| 56146 
8453d35e4684
discontinued somewhat pointless "thy_script" keyword kind;
 wenzelm parents: 
48891diff
changeset | 17 | "inductive" "coinductive" "inductive_cases" "rep_datatype" "primrec" :: thy_decl and | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46947diff
changeset | 18 | "domains" "intros" "monos" "con_defs" "type_intros" "type_elims" | 
| 63435 | 19 | "elimination" "induction" "case_eqns" "recursor_eqns" :: quasi_command | 
| 26189 | 20 | begin | 
| 21 | ||
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
38514diff
changeset | 22 | lemma def_swap_iff: "a == b ==> a = c \<longleftrightarrow> c = b" | 
| 26189 | 23 | by blast | 
| 24 | ||
| 25 | lemma def_trans: "f == g ==> g(a) = b ==> f(a) = b" | |
| 26 | by simp | |
| 27 | ||
| 28 | lemma refl_thin: "!!P. a = a ==> P ==> P" . | |
| 29 | ||
| 48891 | 30 | ML_file "ind_syntax.ML" | 
| 31 | ML_file "Tools/ind_cases.ML" | |
| 32 | ML_file "Tools/cartprod.ML" | |
| 33 | ML_file "Tools/inductive_package.ML" | |
| 34 | ML_file "Tools/induct_tacs.ML" | |
| 35 | ML_file "Tools/primrec_package.ML" | |
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 36 | |
| 60770 | 37 | ML \<open> | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 38 | structure Lfp = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 39 | struct | 
| 26189 | 40 |   val oper      = @{const lfp}
 | 
| 41 |   val bnd_mono  = @{const bnd_mono}
 | |
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 42 |   val bnd_monoI = @{thm bnd_monoI}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 43 |   val subs      = @{thm def_lfp_subset}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 44 |   val Tarski    = @{thm def_lfp_unfold}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 45 |   val induct    = @{thm def_induct}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 46 | end; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 47 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 48 | structure Standard_Prod = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 49 | struct | 
| 26189 | 50 |   val sigma     = @{const Sigma}
 | 
| 51 |   val pair      = @{const Pair}
 | |
| 52 |   val split_name = @{const_name split}
 | |
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 53 |   val pair_iff  = @{thm Pair_iff}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 54 |   val split_eq  = @{thm split}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 55 |   val fsplitI   = @{thm splitI}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 56 |   val fsplitD   = @{thm splitD}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 57 |   val fsplitE   = @{thm splitE}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 58 | end; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 59 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 60 | structure Standard_CP = CartProd_Fun (Standard_Prod); | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 61 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 62 | structure Standard_Sum = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 63 | struct | 
| 26189 | 64 |   val sum       = @{const sum}
 | 
| 65 |   val inl       = @{const Inl}
 | |
| 66 |   val inr       = @{const Inr}
 | |
| 67 |   val elim      = @{const case}
 | |
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 68 |   val case_inl  = @{thm case_Inl}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 69 |   val case_inr  = @{thm case_Inr}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 70 |   val inl_iff   = @{thm Inl_iff}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 71 |   val inr_iff   = @{thm Inr_iff}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 72 |   val distinct  = @{thm Inl_Inr_iff}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 73 |   val distinct' = @{thm Inr_Inl_iff}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 74 | val free_SEs = Ind_Syntax.mk_free_SEs | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 75 | [distinct, distinct', inl_iff, inr_iff, Standard_Prod.pair_iff] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 76 | end; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 77 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 78 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 79 | structure Ind_Package = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 80 | Add_inductive_def_Fun | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 81 | (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 82 | and Su=Standard_Sum val coind = false); | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 83 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 84 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 85 | structure Gfp = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 86 | struct | 
| 26189 | 87 |   val oper      = @{const gfp}
 | 
| 88 |   val bnd_mono  = @{const bnd_mono}
 | |
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 89 |   val bnd_monoI = @{thm bnd_monoI}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 90 |   val subs      = @{thm def_gfp_subset}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 91 |   val Tarski    = @{thm def_gfp_unfold}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 92 |   val induct    = @{thm def_Collect_coinduct}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 93 | end; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 94 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 95 | structure Quine_Prod = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 96 | struct | 
| 26189 | 97 |   val sigma     = @{const QSigma}
 | 
| 98 |   val pair      = @{const QPair}
 | |
| 99 |   val split_name = @{const_name qsplit}
 | |
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 100 |   val pair_iff  = @{thm QPair_iff}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 101 |   val split_eq  = @{thm qsplit}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 102 |   val fsplitI   = @{thm qsplitI}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 103 |   val fsplitD   = @{thm qsplitD}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 104 |   val fsplitE   = @{thm qsplitE}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 105 | end; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 106 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 107 | structure Quine_CP = CartProd_Fun (Quine_Prod); | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 108 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 109 | structure Quine_Sum = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 110 | struct | 
| 26189 | 111 |   val sum       = @{const qsum}
 | 
| 112 |   val inl       = @{const QInl}
 | |
| 113 |   val inr       = @{const QInr}
 | |
| 114 |   val elim      = @{const qcase}
 | |
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 115 |   val case_inl  = @{thm qcase_QInl}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 116 |   val case_inr  = @{thm qcase_QInr}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 117 |   val inl_iff   = @{thm QInl_iff}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 118 |   val inr_iff   = @{thm QInr_iff}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 119 |   val distinct  = @{thm QInl_QInr_iff}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 120 |   val distinct' = @{thm QInr_QInl_iff}
 | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 121 | val free_SEs = Ind_Syntax.mk_free_SEs | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 122 | [distinct, distinct', inl_iff, inr_iff, Quine_Prod.pair_iff] | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 123 | end; | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 124 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 125 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 126 | structure CoInd_Package = | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 127 | Add_inductive_def_Fun(structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 128 | and Su=Quine_Sum val coind = true); | 
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 129 | |
| 60770 | 130 | \<close> | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 131 | |
| 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: diff
changeset | 132 | end |