| author | huffman | 
| Wed, 18 Apr 2012 10:52:49 +0200 | |
| changeset 47534 | 06cc372a80ed | 
| parent 46950 | d0181abdbdac | 
| child 48357 | 828ace4f75ab | 
| permissions | -rw-r--r-- | 
| 7700 | 1 | (* Title: HOL/Inductive.thy | 
| 10402 | 2 | Author: Markus Wenzel, TU Muenchen | 
| 11688 | 3 | *) | 
| 10727 | 4 | |
| 24915 | 5 | header {* Knaster-Tarski Fixpoint Theorem and inductive definitions *}
 | 
| 1187 | 6 | |
| 15131 | 7 | theory Inductive | 
| 44860 
56101fa00193
renamed theory Complete_Lattice to Complete_Lattices, in accordance with Lattices, Orderings etc.
 haftmann parents: 
43580diff
changeset | 8 | imports Complete_Lattices | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46947diff
changeset | 9 | keywords | 
| 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46947diff
changeset | 10 | "inductive" "coinductive" :: thy_decl and | 
| 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46947diff
changeset | 11 | "inductive_cases" "inductive_simps" :: thy_script and "monos" and | 
| 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46947diff
changeset | 12 | "rep_datatype" :: thy_goal and | 
| 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46947diff
changeset | 13 | "primrec" :: thy_decl | 
| 16417 | 14 | uses | 
| 45897 
65cef0298158
clarified modules that contribute to datatype package;
 wenzelm parents: 
45891diff
changeset | 15 | "Tools/dseq.ML" | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31604diff
changeset | 16 |   ("Tools/inductive.ML")
 | 
| 45897 
65cef0298158
clarified modules that contribute to datatype package;
 wenzelm parents: 
45891diff
changeset | 17 |   ("Tools/Datatype/datatype_aux.ML")
 | 
| 
65cef0298158
clarified modules that contribute to datatype package;
 wenzelm parents: 
45891diff
changeset | 18 |   ("Tools/Datatype/datatype_prop.ML")
 | 
| 33963 
977b94b64905
renamed former datatype.ML to datatype_data.ML; datatype.ML provides uniform view on datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33959diff
changeset | 19 |   ("Tools/Datatype/datatype_data.ML")
 | 
| 45891 
d73605c829cc
clarified module dependencies: Datatype_Data, Datatype_Case, Rep_Datatype;
 wenzelm parents: 
45890diff
changeset | 20 |   ("Tools/Datatype/datatype_case.ML")
 | 
| 45890 | 21 |   ("Tools/Datatype/rep_datatype.ML")
 | 
| 31775 | 22 |   ("Tools/Datatype/datatype_codegen.ML")
 | 
| 45897 
65cef0298158
clarified modules that contribute to datatype package;
 wenzelm parents: 
45891diff
changeset | 23 |   ("Tools/Datatype/primrec.ML")
 | 
| 15131 | 24 | begin | 
| 10727 | 25 | |
| 24915 | 26 | subsection {* Least and greatest fixed points *}
 | 
| 27 | ||
| 26013 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 28 | context complete_lattice | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 29 | begin | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 30 | |
| 24915 | 31 | definition | 
| 26013 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 32 |   lfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" where
 | 
| 24915 | 33 |   "lfp f = Inf {u. f u \<le> u}"    --{*least fixed point*}
 | 
| 34 | ||
| 35 | definition | |
| 26013 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 36 |   gfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" where
 | 
| 24915 | 37 |   "gfp f = Sup {u. u \<le> f u}"    --{*greatest fixed point*}
 | 
| 38 | ||
| 39 | ||
| 40 | subsection{* Proof of Knaster-Tarski Theorem using @{term lfp} *}
 | |
| 41 | ||
| 42 | text{*@{term "lfp f"} is the least upper bound of 
 | |
| 43 |       the set @{term "{u. f(u) \<le> u}"} *}
 | |
| 44 | ||
| 45 | lemma lfp_lowerbound: "f A \<le> A ==> lfp f \<le> A" | |
| 46 | by (auto simp add: lfp_def intro: Inf_lower) | |
| 47 | ||
| 48 | lemma lfp_greatest: "(!!u. f u \<le> u ==> A \<le> u) ==> A \<le> lfp f" | |
| 49 | by (auto simp add: lfp_def intro: Inf_greatest) | |
| 50 | ||
| 26013 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 51 | end | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 52 | |
| 24915 | 53 | lemma lfp_lemma2: "mono f ==> f (lfp f) \<le> lfp f" | 
| 54 | by (iprover intro: lfp_greatest order_trans monoD lfp_lowerbound) | |
| 55 | ||
| 56 | lemma lfp_lemma3: "mono f ==> lfp f \<le> f (lfp f)" | |
| 57 | by (iprover intro: lfp_lemma2 monoD lfp_lowerbound) | |
| 58 | ||
| 59 | lemma lfp_unfold: "mono f ==> lfp f = f (lfp f)" | |
| 60 | by (iprover intro: order_antisym lfp_lemma2 lfp_lemma3) | |
| 61 | ||
| 62 | lemma lfp_const: "lfp (\<lambda>x. t) = t" | |
| 63 | by (rule lfp_unfold) (simp add:mono_def) | |
| 64 | ||
| 65 | ||
| 66 | subsection {* General induction rules for least fixed points *}
 | |
| 67 | ||
| 68 | theorem lfp_induct: | |
| 69 | assumes mono: "mono f" and ind: "f (inf (lfp f) P) <= P" | |
| 70 | shows "lfp f <= P" | |
| 71 | proof - | |
| 72 | have "inf (lfp f) P <= lfp f" by (rule inf_le1) | |
| 73 | with mono have "f (inf (lfp f) P) <= f (lfp f)" .. | |
| 74 | also from mono have "f (lfp f) = lfp f" by (rule lfp_unfold [symmetric]) | |
| 75 | finally have "f (inf (lfp f) P) <= lfp f" . | |
| 76 | from this and ind have "f (inf (lfp f) P) <= inf (lfp f) P" by (rule le_infI) | |
| 77 | hence "lfp f <= inf (lfp f) P" by (rule lfp_lowerbound) | |
| 78 | also have "inf (lfp f) P <= P" by (rule inf_le2) | |
| 79 | finally show ?thesis . | |
| 80 | qed | |
| 81 | ||
| 82 | lemma lfp_induct_set: | |
| 83 | assumes lfp: "a: lfp(f)" | |
| 84 | and mono: "mono(f)" | |
| 85 |       and indhyp: "!!x. [| x: f(lfp(f) Int {x. P(x)}) |] ==> P(x)"
 | |
| 86 | shows "P(a)" | |
| 87 | by (rule lfp_induct [THEN subsetD, THEN CollectD, OF mono _ lfp]) | |
| 32683 
7c1fe854ca6a
inter and union are mere abbreviations for inf and sup
 haftmann parents: 
32587diff
changeset | 88 | (auto simp: intro: indhyp) | 
| 24915 | 89 | |
| 26013 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 90 | lemma lfp_ordinal_induct: | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 91 | fixes f :: "'a\<Colon>complete_lattice \<Rightarrow> 'a" | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 92 | assumes mono: "mono f" | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 93 | and P_f: "\<And>S. P S \<Longrightarrow> P (f S)" | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 94 | and P_Union: "\<And>M. \<forall>S\<in>M. P S \<Longrightarrow> P (Sup M)" | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 95 | shows "P (lfp f)" | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 96 | proof - | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 97 |   let ?M = "{S. S \<le> lfp f \<and> P S}"
 | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 98 | have "P (Sup ?M)" using P_Union by simp | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 99 | also have "Sup ?M = lfp f" | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 100 | proof (rule antisym) | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 101 | show "Sup ?M \<le> lfp f" by (blast intro: Sup_least) | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 102 | hence "f (Sup ?M) \<le> f (lfp f)" by (rule mono [THEN monoD]) | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 103 | hence "f (Sup ?M) \<le> lfp f" using mono [THEN lfp_unfold] by simp | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 104 | hence "f (Sup ?M) \<in> ?M" using P_f P_Union by simp | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 105 | hence "f (Sup ?M) \<le> Sup ?M" by (rule Sup_upper) | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 106 | thus "lfp f \<le> Sup ?M" by (rule lfp_lowerbound) | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 107 | qed | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 108 | finally show ?thesis . | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 109 | qed | 
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 110 | |
| 
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
 haftmann parents: 
25557diff
changeset | 111 | lemma lfp_ordinal_induct_set: | 
| 24915 | 112 | assumes mono: "mono f" | 
| 113 | and P_f: "!!S. P S ==> P(f S)" | |
| 114 | and P_Union: "!!M. !S:M. P S ==> P(Union M)" | |
| 115 | shows "P(lfp f)" | |
| 46008 
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
 wenzelm parents: 
45907diff
changeset | 116 | using assms by (rule lfp_ordinal_induct) | 
| 24915 | 117 | |
| 118 | ||
| 119 | text{*Definition forms of @{text lfp_unfold} and @{text lfp_induct}, 
 | |
| 120 | to control unfolding*} | |
| 121 | ||
| 122 | lemma def_lfp_unfold: "[| h==lfp(f); mono(f) |] ==> h = f(h)" | |
| 45899 | 123 | by (auto intro!: lfp_unfold) | 
| 24915 | 124 | |
| 125 | lemma def_lfp_induct: | |
| 126 | "[| A == lfp(f); mono(f); | |
| 127 | f (inf A P) \<le> P | |
| 128 | |] ==> A \<le> P" | |
| 129 | by (blast intro: lfp_induct) | |
| 130 | ||
| 131 | lemma def_lfp_induct_set: | |
| 132 | "[| A == lfp(f); mono(f); a:A; | |
| 133 |         !!x. [| x: f(A Int {x. P(x)}) |] ==> P(x)         
 | |
| 134 | |] ==> P(a)" | |
| 135 | by (blast intro: lfp_induct_set) | |
| 136 | ||
| 137 | (*Monotonicity of lfp!*) | |
| 138 | lemma lfp_mono: "(!!Z. f Z \<le> g Z) ==> lfp f \<le> lfp g" | |
| 139 | by (rule lfp_lowerbound [THEN lfp_greatest], blast intro: order_trans) | |
| 140 | ||
| 141 | ||
| 142 | subsection {* Proof of Knaster-Tarski Theorem using @{term gfp} *}
 | |
| 143 | ||
| 144 | text{*@{term "gfp f"} is the greatest lower bound of 
 | |
| 145 |       the set @{term "{u. u \<le> f(u)}"} *}
 | |
| 146 | ||
| 147 | lemma gfp_upperbound: "X \<le> f X ==> X \<le> gfp f" | |
| 148 | by (auto simp add: gfp_def intro: Sup_upper) | |
| 149 | ||
| 150 | lemma gfp_least: "(!!u. u \<le> f u ==> u \<le> X) ==> gfp f \<le> X" | |
| 151 | by (auto simp add: gfp_def intro: Sup_least) | |
| 152 | ||
| 153 | lemma gfp_lemma2: "mono f ==> gfp f \<le> f (gfp f)" | |
| 154 | by (iprover intro: gfp_least order_trans monoD gfp_upperbound) | |
| 155 | ||
| 156 | lemma gfp_lemma3: "mono f ==> f (gfp f) \<le> gfp f" | |
| 157 | by (iprover intro: gfp_lemma2 monoD gfp_upperbound) | |
| 158 | ||
| 159 | lemma gfp_unfold: "mono f ==> gfp f = f (gfp f)" | |
| 160 | by (iprover intro: order_antisym gfp_lemma2 gfp_lemma3) | |
| 161 | ||
| 162 | ||
| 163 | subsection {* Coinduction rules for greatest fixed points *}
 | |
| 164 | ||
| 165 | text{*weak version*}
 | |
| 166 | lemma weak_coinduct: "[| a: X; X \<subseteq> f(X) |] ==> a : gfp(f)" | |
| 45899 | 167 | by (rule gfp_upperbound [THEN subsetD]) auto | 
| 24915 | 168 | |
| 169 | lemma weak_coinduct_image: "!!X. [| a : X; g`X \<subseteq> f (g`X) |] ==> g a : gfp f" | |
| 45899 | 170 | apply (erule gfp_upperbound [THEN subsetD]) | 
| 171 | apply (erule imageI) | |
| 172 | done | |
| 24915 | 173 | |
| 174 | lemma coinduct_lemma: | |
| 175 | "[| X \<le> f (sup X (gfp f)); mono f |] ==> sup X (gfp f) \<le> f (sup X (gfp f))" | |
| 176 | apply (frule gfp_lemma2) | |
| 177 | apply (drule mono_sup) | |
| 178 | apply (rule le_supI) | |
| 179 | apply assumption | |
| 180 | apply (rule order_trans) | |
| 181 | apply (rule order_trans) | |
| 182 | apply assumption | |
| 183 | apply (rule sup_ge2) | |
| 184 | apply assumption | |
| 185 | done | |
| 186 | ||
| 187 | text{*strong version, thanks to Coen and Frost*}
 | |
| 188 | lemma coinduct_set: "[| mono(f); a: X; X \<subseteq> f(X Un gfp(f)) |] ==> a : gfp(f)" | |
| 45899 | 189 | by (blast intro: weak_coinduct [OF _ coinduct_lemma]) | 
| 24915 | 190 | |
| 191 | lemma coinduct: "[| mono(f); X \<le> f (sup X (gfp f)) |] ==> X \<le> gfp(f)" | |
| 192 | apply (rule order_trans) | |
| 193 | apply (rule sup_ge1) | |
| 194 | apply (erule gfp_upperbound [OF coinduct_lemma]) | |
| 195 | apply assumption | |
| 196 | done | |
| 197 | ||
| 198 | lemma gfp_fun_UnI2: "[| mono(f); a: gfp(f) |] ==> a: f(X Un gfp(f))" | |
| 45899 | 199 | by (blast dest: gfp_lemma2 mono_Un) | 
| 24915 | 200 | |
| 201 | ||
| 202 | subsection {* Even Stronger Coinduction Rule, by Martin Coen *}
 | |
| 203 | ||
| 204 | text{* Weakens the condition @{term "X \<subseteq> f(X)"} to one expressed using both
 | |
| 205 |   @{term lfp} and @{term gfp}*}
 | |
| 206 | ||
| 207 | lemma coinduct3_mono_lemma: "mono(f) ==> mono(%x. f(x) Un X Un B)" | |
| 208 | by (iprover intro: subset_refl monoI Un_mono monoD) | |
| 209 | ||
| 210 | lemma coinduct3_lemma: | |
| 211 | "[| X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f))); mono(f) |] | |
| 212 | ==> lfp(%x. f(x) Un X Un gfp(f)) \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)))" | |
| 213 | apply (rule subset_trans) | |
| 214 | apply (erule coinduct3_mono_lemma [THEN lfp_lemma3]) | |
| 215 | apply (rule Un_least [THEN Un_least]) | |
| 216 | apply (rule subset_refl, assumption) | |
| 217 | apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption) | |
| 46008 
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
 wenzelm parents: 
45907diff
changeset | 218 | apply (rule monoD, assumption) | 
| 24915 | 219 | apply (subst coinduct3_mono_lemma [THEN lfp_unfold], auto) | 
| 220 | done | |
| 221 | ||
| 222 | lemma coinduct3: | |
| 223 | "[| mono(f); a:X; X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f))) |] ==> a : gfp(f)" | |
| 224 | apply (rule coinduct3_lemma [THEN [2] weak_coinduct]) | |
| 41081 | 225 | apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst]) | 
| 226 | apply (simp_all) | |
| 24915 | 227 | done | 
| 228 | ||
| 229 | ||
| 230 | text{*Definition forms of @{text gfp_unfold} and @{text coinduct}, 
 | |
| 231 | to control unfolding*} | |
| 232 | ||
| 233 | lemma def_gfp_unfold: "[| A==gfp(f); mono(f) |] ==> A = f(A)" | |
| 45899 | 234 | by (auto intro!: gfp_unfold) | 
| 24915 | 235 | |
| 236 | lemma def_coinduct: | |
| 237 | "[| A==gfp(f); mono(f); X \<le> f(sup X A) |] ==> X \<le> A" | |
| 45899 | 238 | by (iprover intro!: coinduct) | 
| 24915 | 239 | |
| 240 | lemma def_coinduct_set: | |
| 241 | "[| A==gfp(f); mono(f); a:X; X \<subseteq> f(X Un A) |] ==> a: A" | |
| 45899 | 242 | by (auto intro!: coinduct_set) | 
| 24915 | 243 | |
| 244 | (*The version used in the induction/coinduction package*) | |
| 245 | lemma def_Collect_coinduct: | |
| 246 | "[| A == gfp(%w. Collect(P(w))); mono(%w. Collect(P(w))); | |
| 247 | a: X; !!z. z: X ==> P (X Un A) z |] ==> | |
| 248 | a : A" | |
| 45899 | 249 | by (erule def_coinduct_set) auto | 
| 24915 | 250 | |
| 251 | lemma def_coinduct3: | |
| 252 | "[| A==gfp(f); mono(f); a:X; X \<subseteq> f(lfp(%x. f(x) Un X Un A)) |] ==> a: A" | |
| 45899 | 253 | by (auto intro!: coinduct3) | 
| 24915 | 254 | |
| 255 | text{*Monotonicity of @{term gfp}!*}
 | |
| 256 | lemma gfp_mono: "(!!Z. f Z \<le> g Z) ==> gfp f \<le> gfp g" | |
| 257 | by (rule gfp_upperbound [THEN gfp_least], blast intro: order_trans) | |
| 258 | ||
| 259 | ||
| 23734 | 260 | subsection {* Inductive predicates and sets *}
 | 
| 11688 | 261 | |
| 262 | text {* Package setup. *}
 | |
| 10402 | 263 | |
| 23734 | 264 | theorems basic_monos = | 
| 22218 | 265 | subset_refl imp_refl disj_mono conj_mono ex_mono all_mono if_bool_eq_conj | 
| 11688 | 266 | Collect_mono in_mono vimage_mono | 
| 267 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31604diff
changeset | 268 | use "Tools/inductive.ML" | 
| 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31604diff
changeset | 269 | setup Inductive.setup | 
| 21018 
e6b8d6784792
Added new package for inductive definitions, moved old package
 berghofe parents: 
20604diff
changeset | 270 | |
| 23734 | 271 | theorems [mono] = | 
| 22218 | 272 | imp_refl disj_mono conj_mono ex_mono all_mono if_bool_eq_conj | 
| 33934 
25d6a8982e37
Streamlined setup for monotonicity rules (no longer requires classical rules).
 berghofe parents: 
32701diff
changeset | 273 | imp_mono not_mono | 
| 21018 
e6b8d6784792
Added new package for inductive definitions, moved old package
 berghofe parents: 
20604diff
changeset | 274 | Ball_def Bex_def | 
| 
e6b8d6784792
Added new package for inductive definitions, moved old package
 berghofe parents: 
20604diff
changeset | 275 | induct_rulify_fallback | 
| 
e6b8d6784792
Added new package for inductive definitions, moved old package
 berghofe parents: 
20604diff
changeset | 276 | |
| 11688 | 277 | |
| 12023 | 278 | subsection {* Inductive datatypes and primitive recursion *}
 | 
| 11688 | 279 | |
| 11825 | 280 | text {* Package setup. *}
 | 
| 281 | ||
| 45897 
65cef0298158
clarified modules that contribute to datatype package;
 wenzelm parents: 
45891diff
changeset | 282 | use "Tools/Datatype/datatype_aux.ML" | 
| 
65cef0298158
clarified modules that contribute to datatype package;
 wenzelm parents: 
45891diff
changeset | 283 | use "Tools/Datatype/datatype_prop.ML" | 
| 45891 
d73605c829cc
clarified module dependencies: Datatype_Data, Datatype_Case, Rep_Datatype;
 wenzelm parents: 
45890diff
changeset | 284 | use "Tools/Datatype/datatype_data.ML" setup Datatype_Data.setup | 
| 
d73605c829cc
clarified module dependencies: Datatype_Data, Datatype_Case, Rep_Datatype;
 wenzelm parents: 
45890diff
changeset | 285 | use "Tools/Datatype/datatype_case.ML" setup Datatype_Case.setup | 
| 45890 | 286 | use "Tools/Datatype/rep_datatype.ML" | 
| 45897 
65cef0298158
clarified modules that contribute to datatype package;
 wenzelm parents: 
45891diff
changeset | 287 | use "Tools/Datatype/datatype_codegen.ML" setup Datatype_Codegen.setup | 
| 
65cef0298158
clarified modules that contribute to datatype package;
 wenzelm parents: 
45891diff
changeset | 288 | use "Tools/Datatype/primrec.ML" | 
| 12437 
6d4e02b6dd43
Moved code generator setup from Recdef to Inductive.
 berghofe parents: 
12023diff
changeset | 289 | |
| 23526 | 290 | text{* Lambda-abstractions with pattern matching: *}
 | 
| 291 | ||
| 292 | syntax | |
| 23529 | 293 |   "_lam_pats_syntax" :: "cases_syn => 'a => 'b"               ("(%_)" 10)
 | 
| 23526 | 294 | syntax (xsymbols) | 
| 23529 | 295 |   "_lam_pats_syntax" :: "cases_syn => 'a => 'b"               ("(\<lambda>_)" 10)
 | 
| 23526 | 296 | |
| 23529 | 297 | parse_translation (advanced) {*
 | 
| 298 | let | |
| 299 | fun fun_tr ctxt [cs] = | |
| 300 | let | |
| 45899 | 301 | val x = Syntax.free (fst (Name.variant "x" (Term.declare_term_frees cs Name.context))); | 
| 45891 
d73605c829cc
clarified module dependencies: Datatype_Data, Datatype_Case, Rep_Datatype;
 wenzelm parents: 
45890diff
changeset | 302 | val ft = Datatype_Case.case_tr true ctxt [x, cs]; | 
| 23529 | 303 | in lambda x ft end | 
| 35115 | 304 | in [(@{syntax_const "_lam_pats_syntax"}, fun_tr)] end
 | 
| 23526 | 305 | *} | 
| 306 | ||
| 307 | end |