author | blanchet |
Tue, 17 Feb 2015 17:22:45 +0100 | |
changeset 59552 | ae50c9b82444 |
parent 58916 | 229765cc3414 |
child 60758 | d8d85a8172b5 |
permissions | -rw-r--r-- |
58128 | 1 |
(* Title: HOL/BNF_Least_Fixpoint.thy |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
2 |
Author: Dmitriy Traytel, TU Muenchen |
53305 | 3 |
Author: Lorenz Panny, TU Muenchen |
4 |
Author: Jasmin Blanchette, TU Muenchen |
|
57698 | 5 |
Copyright 2012, 2013, 2014 |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
6 |
|
58352
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
58314
diff
changeset
|
7 |
Least fixpoint (datatype) operation on bounded natural functors. |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
8 |
*) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
9 |
|
58889 | 10 |
section {* Least Fixpoint (Datatype) Operation on Bounded Natural Functors *} |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
11 |
|
58128 | 12 |
theory BNF_Least_Fixpoint |
13 |
imports BNF_Fixpoint_Base |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
14 |
keywords |
58305
57752a91eec4
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents:
58276
diff
changeset
|
15 |
"datatype" :: thy_decl and |
55575
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
blanchet
parents:
55571
diff
changeset
|
16 |
"datatype_compat" :: thy_decl |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
17 |
begin |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
18 |
|
49312 | 19 |
lemma subset_emptyI: "(\<And>x. x \<in> A \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> {}" |
57987 | 20 |
by blast |
49312 | 21 |
|
56346 | 22 |
lemma image_Collect_subsetI: "(\<And>x. P x \<Longrightarrow> f x \<in> B) \<Longrightarrow> f ` {x. P x} \<subseteq> B" |
57987 | 23 |
by blast |
49312 | 24 |
|
25 |
lemma Collect_restrict: "{x. x \<in> X \<and> P x} \<subseteq> X" |
|
57987 | 26 |
by auto |
49312 | 27 |
|
28 |
lemma prop_restrict: "\<lbrakk>x \<in> Z; Z \<subseteq> {x. x \<in> X \<and> P x}\<rbrakk> \<Longrightarrow> P x" |
|
57987 | 29 |
by auto |
49312 | 30 |
|
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54841
diff
changeset
|
31 |
lemma underS_I: "\<lbrakk>i \<noteq> j; (i, j) \<in> R\<rbrakk> \<Longrightarrow> i \<in> underS R j" |
57987 | 32 |
unfolding underS_def by simp |
49312 | 33 |
|
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54841
diff
changeset
|
34 |
lemma underS_E: "i \<in> underS R j \<Longrightarrow> i \<noteq> j \<and> (i, j) \<in> R" |
57987 | 35 |
unfolding underS_def by simp |
49312 | 36 |
|
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54841
diff
changeset
|
37 |
lemma underS_Field: "i \<in> underS R j \<Longrightarrow> i \<in> Field R" |
57987 | 38 |
unfolding underS_def Field_def by auto |
49312 | 39 |
|
40 |
lemma FieldI2: "(i, j) \<in> R \<Longrightarrow> j \<in> Field R" |
|
57987 | 41 |
unfolding Field_def by auto |
49312 | 42 |
|
57641
dc59f147b27d
more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.thy);
wenzelm
parents:
57493
diff
changeset
|
43 |
lemma fst_convol': "fst (\<langle>f, g\<rangle> x) = f x" |
57987 | 44 |
using fst_convol unfolding convol_def by simp |
49312 | 45 |
|
57641
dc59f147b27d
more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.thy);
wenzelm
parents:
57493
diff
changeset
|
46 |
lemma snd_convol': "snd (\<langle>f, g\<rangle> x) = g x" |
57987 | 47 |
using snd_convol unfolding convol_def by simp |
49312 | 48 |
|
57641
dc59f147b27d
more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.thy);
wenzelm
parents:
57493
diff
changeset
|
49 |
lemma convol_expand_snd: "fst o f = g \<Longrightarrow> \<langle>g, snd o f\<rangle> = f" |
57987 | 50 |
unfolding convol_def by auto |
49312 | 51 |
|
55811 | 52 |
lemma convol_expand_snd': |
53 |
assumes "(fst o f = g)" |
|
57641
dc59f147b27d
more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.thy);
wenzelm
parents:
57493
diff
changeset
|
54 |
shows "h = snd o f \<longleftrightarrow> \<langle>g, h\<rangle> = f" |
55811 | 55 |
proof - |
57641
dc59f147b27d
more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.thy);
wenzelm
parents:
57493
diff
changeset
|
56 |
from assms have *: "\<langle>g, snd o f\<rangle> = f" by (rule convol_expand_snd) |
dc59f147b27d
more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.thy);
wenzelm
parents:
57493
diff
changeset
|
57 |
then have "h = snd o f \<longleftrightarrow> h = snd o \<langle>g, snd o f\<rangle>" by simp |
55811 | 58 |
moreover have "\<dots> \<longleftrightarrow> h = snd o f" by (simp add: snd_convol) |
57641
dc59f147b27d
more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.thy);
wenzelm
parents:
57493
diff
changeset
|
59 |
moreover have "\<dots> \<longleftrightarrow> \<langle>g, h\<rangle> = f" by (subst (2) *[symmetric]) (auto simp: convol_def fun_eq_iff) |
55811 | 60 |
ultimately show ?thesis by simp |
61 |
qed |
|
57987 | 62 |
|
49312 | 63 |
lemma bij_betwE: "bij_betw f A B \<Longrightarrow> \<forall>a\<in>A. f a \<in> B" |
57987 | 64 |
unfolding bij_betw_def by auto |
49312 | 65 |
|
66 |
lemma bij_betw_imageE: "bij_betw f A B \<Longrightarrow> f ` A = B" |
|
57987 | 67 |
unfolding bij_betw_def by auto |
49312 | 68 |
|
58159 | 69 |
lemma f_the_inv_into_f_bij_betw: |
70 |
"bij_betw f A B \<Longrightarrow> (bij_betw f A B \<Longrightarrow> x \<in> B) \<Longrightarrow> f (the_inv_into A f x) = x" |
|
56237 | 71 |
unfolding bij_betw_def by (blast intro: f_the_inv_into_f) |
49312 | 72 |
|
56237 | 73 |
lemma ex_bij_betw: "|A| \<le>o (r :: 'b rel) \<Longrightarrow> \<exists>f B :: 'b set. bij_betw f B A" |
58159 | 74 |
by (subst (asm) internalize_card_of_ordLeq) (auto dest!: iffD2[OF card_of_ordIso ordIso_symmetric]) |
49312 | 75 |
|
76 |
lemma bij_betwI': |
|
77 |
"\<lbrakk>\<And>x y. \<lbrakk>x \<in> X; y \<in> X\<rbrakk> \<Longrightarrow> (f x = f y) = (x = y); |
|
78 |
\<And>x. x \<in> X \<Longrightarrow> f x \<in> Y; |
|
79 |
\<And>y. y \<in> Y \<Longrightarrow> \<exists>x \<in> X. y = f x\<rbrakk> \<Longrightarrow> bij_betw f X Y" |
|
57987 | 80 |
unfolding bij_betw_def inj_on_def by blast |
49312 | 81 |
|
82 |
lemma surj_fun_eq: |
|
83 |
assumes surj_on: "f ` X = UNIV" and eq_on: "\<forall>x \<in> X. (g1 o f) x = (g2 o f) x" |
|
84 |
shows "g1 = g2" |
|
85 |
proof (rule ext) |
|
86 |
fix y |
|
87 |
from surj_on obtain x where "x \<in> X" and "y = f x" by blast |
|
88 |
thus "g1 y = g2 y" using eq_on by simp |
|
89 |
qed |
|
90 |
||
91 |
lemma Card_order_wo_rel: "Card_order r \<Longrightarrow> wo_rel r" |
|
58147 | 92 |
unfolding wo_rel_def card_order_on_def by blast |
49312 | 93 |
|
58147 | 94 |
lemma Cinfinite_limit: "\<lbrakk>x \<in> Field r; Cinfinite r\<rbrakk> \<Longrightarrow> \<exists>y \<in> Field r. x \<noteq> y \<and> (x, y) \<in> r" |
95 |
unfolding cinfinite_def by (auto simp add: infinite_Card_order_limit) |
|
49312 | 96 |
|
97 |
lemma Card_order_trans: |
|
98 |
"\<lbrakk>Card_order r; x \<noteq> y; (x, y) \<in> r; y \<noteq> z; (y, z) \<in> r\<rbrakk> \<Longrightarrow> x \<noteq> z \<and> (x, z) \<in> r" |
|
58147 | 99 |
unfolding card_order_on_def well_order_on_def linear_order_on_def |
100 |
partial_order_on_def preorder_on_def trans_def antisym_def by blast |
|
49312 | 101 |
|
102 |
lemma Cinfinite_limit2: |
|
58147 | 103 |
assumes x1: "x1 \<in> Field r" and x2: "x2 \<in> Field r" and r: "Cinfinite r" |
104 |
shows "\<exists>y \<in> Field r. (x1 \<noteq> y \<and> (x1, y) \<in> r) \<and> (x2 \<noteq> y \<and> (x2, y) \<in> r)" |
|
49312 | 105 |
proof - |
106 |
from r have trans: "trans r" and total: "Total r" and antisym: "antisym r" |
|
107 |
unfolding card_order_on_def well_order_on_def linear_order_on_def |
|
108 |
partial_order_on_def preorder_on_def by auto |
|
109 |
obtain y1 where y1: "y1 \<in> Field r" "x1 \<noteq> y1" "(x1, y1) \<in> r" |
|
110 |
using Cinfinite_limit[OF x1 r] by blast |
|
111 |
obtain y2 where y2: "y2 \<in> Field r" "x2 \<noteq> y2" "(x2, y2) \<in> r" |
|
112 |
using Cinfinite_limit[OF x2 r] by blast |
|
113 |
show ?thesis |
|
114 |
proof (cases "y1 = y2") |
|
115 |
case True with y1 y2 show ?thesis by blast |
|
116 |
next |
|
117 |
case False |
|
118 |
with y1(1) y2(1) total have "(y1, y2) \<in> r \<or> (y2, y1) \<in> r" |
|
119 |
unfolding total_on_def by auto |
|
120 |
thus ?thesis |
|
121 |
proof |
|
122 |
assume *: "(y1, y2) \<in> r" |
|
123 |
with trans y1(3) have "(x1, y2) \<in> r" unfolding trans_def by blast |
|
124 |
with False y1 y2 * antisym show ?thesis by (cases "x1 = y2") (auto simp: antisym_def) |
|
125 |
next |
|
126 |
assume *: "(y2, y1) \<in> r" |
|
127 |
with trans y2(3) have "(x2, y1) \<in> r" unfolding trans_def by blast |
|
128 |
with False y1 y2 * antisym show ?thesis by (cases "x2 = y1") (auto simp: antisym_def) |
|
129 |
qed |
|
130 |
qed |
|
131 |
qed |
|
132 |
||
58147 | 133 |
lemma Cinfinite_limit_finite: |
134 |
"\<lbrakk>finite X; X \<subseteq> Field r; Cinfinite r\<rbrakk> \<Longrightarrow> \<exists>y \<in> Field r. \<forall>x \<in> X. (x \<noteq> y \<and> (x, y) \<in> r)" |
|
49312 | 135 |
proof (induct X rule: finite_induct) |
136 |
case empty thus ?case unfolding cinfinite_def using ex_in_conv[of "Field r"] finite.emptyI by auto |
|
137 |
next |
|
138 |
case (insert x X) |
|
139 |
then obtain y where y: "y \<in> Field r" "\<forall>x \<in> X. (x \<noteq> y \<and> (x, y) \<in> r)" by blast |
|
140 |
then obtain z where z: "z \<in> Field r" "x \<noteq> z \<and> (x, z) \<in> r" "y \<noteq> z \<and> (y, z) \<in> r" |
|
141 |
using Cinfinite_limit2[OF _ y(1) insert(5), of x] insert(4) by blast |
|
49326 | 142 |
show ?case |
143 |
apply (intro bexI ballI) |
|
144 |
apply (erule insertE) |
|
145 |
apply hypsubst |
|
146 |
apply (rule z(2)) |
|
147 |
using Card_order_trans[OF insert(5)[THEN conjunct2]] y(2) z(3) |
|
148 |
apply blast |
|
149 |
apply (rule z(1)) |
|
150 |
done |
|
49312 | 151 |
qed |
152 |
||
153 |
lemma insert_subsetI: "\<lbrakk>x \<in> A; X \<subseteq> A\<rbrakk> \<Longrightarrow> insert x X \<subseteq> A" |
|
58147 | 154 |
by auto |
49312 | 155 |
|
58136 | 156 |
lemmas well_order_induct_imp = wo_rel.well_order_induct[of r "\<lambda>x. x \<in> Field r \<longrightarrow> P x" for r P] |
49312 | 157 |
|
158 |
lemma meta_spec2: |
|
159 |
assumes "(\<And>x y. PROP P x y)" |
|
160 |
shows "PROP P x y" |
|
58136 | 161 |
by (rule assms) |
49312 | 162 |
|
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54246
diff
changeset
|
163 |
lemma nchotomy_relcomppE: |
55811 | 164 |
assumes "\<And>y. \<exists>x. y = f x" "(r OO s) a c" "\<And>b. r a (f b) \<Longrightarrow> s (f b) c \<Longrightarrow> P" |
165 |
shows P |
|
166 |
proof (rule relcompp.cases[OF assms(2)], hypsubst) |
|
167 |
fix b assume "r a b" "s b c" |
|
168 |
moreover from assms(1) obtain b' where "b = f b'" by blast |
|
169 |
ultimately show P by (blast intro: assms(3)) |
|
170 |
qed |
|
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54246
diff
changeset
|
171 |
|
52731 | 172 |
lemma predicate2D_vimage2p: "\<lbrakk>R \<le> vimage2p f g S; R x y\<rbrakk> \<Longrightarrow> S (f x) (g y)" |
173 |
unfolding vimage2p_def by auto |
|
174 |
||
55945 | 175 |
lemma id_transfer: "rel_fun A A id id" |
176 |
unfolding rel_fun_def by simp |
|
55084 | 177 |
|
58444 | 178 |
lemma fst_transfer: "rel_fun (rel_prod A B) A fst fst" |
58916 | 179 |
unfolding rel_fun_def by simp |
58444 | 180 |
|
181 |
lemma snd_transfer: "rel_fun (rel_prod A B) B snd snd" |
|
58916 | 182 |
unfolding rel_fun_def by simp |
58444 | 183 |
|
184 |
lemma convol_transfer: |
|
185 |
"rel_fun (rel_fun R S) (rel_fun (rel_fun R T) (rel_fun R (rel_prod S T))) BNF_Def.convol BNF_Def.convol" |
|
58916 | 186 |
unfolding rel_fun_def convol_def by auto |
58444 | 187 |
|
55770
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
traytel
parents:
55575
diff
changeset
|
188 |
lemma ssubst_Pair_rhs: "\<lbrakk>(r, s) \<in> R; s' = s\<rbrakk> \<Longrightarrow> (r, s') \<in> R" |
55851
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
blanchet
parents:
55811
diff
changeset
|
189 |
by (rule ssubst) |
55770
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
traytel
parents:
55575
diff
changeset
|
190 |
|
58211
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
191 |
lemma all_mem_range1: |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
192 |
"(\<And>y. y \<in> range f \<Longrightarrow> P y) \<equiv> (\<And>x. P (f x)) " |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
193 |
by (rule equal_intr_rule) fast+ |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
194 |
|
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
195 |
lemma all_mem_range2: |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
196 |
"(\<And>fa y. fa \<in> range f \<Longrightarrow> y \<in> range fa \<Longrightarrow> P y) \<equiv> (\<And>x xa. P (f x xa))" |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
197 |
by (rule equal_intr_rule) fast+ |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
198 |
|
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
199 |
lemma all_mem_range3: |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
200 |
"(\<And>fa fb y. fa \<in> range f \<Longrightarrow> fb \<in> range fa \<Longrightarrow> y \<in> range fb \<Longrightarrow> P y) \<equiv> (\<And>x xa xb. P (f x xa xb))" |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
201 |
by (rule equal_intr_rule) fast+ |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
202 |
|
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
203 |
lemma all_mem_range4: |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
204 |
"(\<And>fa fb fc y. fa \<in> range f \<Longrightarrow> fb \<in> range fa \<Longrightarrow> fc \<in> range fb \<Longrightarrow> y \<in> range fc \<Longrightarrow> P y) \<equiv> |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
205 |
(\<And>x xa xb xc. P (f x xa xb xc))" |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
206 |
by (rule equal_intr_rule) fast+ |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
207 |
|
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
208 |
lemma all_mem_range5: |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
209 |
"(\<And>fa fb fc fd y. fa \<in> range f \<Longrightarrow> fb \<in> range fa \<Longrightarrow> fc \<in> range fb \<Longrightarrow> fd \<in> range fc \<Longrightarrow> |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
210 |
y \<in> range fd \<Longrightarrow> P y) \<equiv> |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
211 |
(\<And>x xa xb xc xd. P (f x xa xb xc xd))" |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
212 |
by (rule equal_intr_rule) fast+ |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
213 |
|
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
214 |
lemma all_mem_range6: |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
215 |
"(\<And>fa fb fc fd fe ff y. fa \<in> range f \<Longrightarrow> fb \<in> range fa \<Longrightarrow> fc \<in> range fb \<Longrightarrow> fd \<in> range fc \<Longrightarrow> |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
216 |
fe \<in> range fd \<Longrightarrow> ff \<in> range fe \<Longrightarrow> y \<in> range ff \<Longrightarrow> P y) \<equiv> |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
217 |
(\<And>x xa xb xc xd xe xf. P (f x xa xb xc xd xe xf))" |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
218 |
by (rule equal_intr_rule) (fastforce, fast) |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
219 |
|
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
220 |
lemma all_mem_range7: |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
221 |
"(\<And>fa fb fc fd fe ff fg y. fa \<in> range f \<Longrightarrow> fb \<in> range fa \<Longrightarrow> fc \<in> range fb \<Longrightarrow> fd \<in> range fc \<Longrightarrow> |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
222 |
fe \<in> range fd \<Longrightarrow> ff \<in> range fe \<Longrightarrow> fg \<in> range ff \<Longrightarrow> y \<in> range fg \<Longrightarrow> P y) \<equiv> |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
223 |
(\<And>x xa xb xc xd xe xf xg. P (f x xa xb xc xd xe xf xg))" |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
224 |
by (rule equal_intr_rule) (fastforce, fast) |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
225 |
|
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
226 |
lemma all_mem_range8: |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
227 |
"(\<And>fa fb fc fd fe ff fg fh y. fa \<in> range f \<Longrightarrow> fb \<in> range fa \<Longrightarrow> fc \<in> range fb \<Longrightarrow> fd \<in> range fc \<Longrightarrow> |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
228 |
fe \<in> range fd \<Longrightarrow> ff \<in> range fe \<Longrightarrow> fg \<in> range ff \<Longrightarrow> fh \<in> range fg \<Longrightarrow> y \<in> range fh \<Longrightarrow> P y) \<equiv> |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
229 |
(\<And>x xa xb xc xd xe xf xg xh. P (f x xa xb xc xd xe xf xg xh))" |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
230 |
by (rule equal_intr_rule) (fastforce, fast) |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
231 |
|
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
232 |
lemmas all_mem_range = all_mem_range1 all_mem_range2 all_mem_range3 all_mem_range4 all_mem_range5 |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
233 |
all_mem_range6 all_mem_range7 all_mem_range8 |
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58182
diff
changeset
|
234 |
|
55062 | 235 |
ML_file "Tools/BNF/bnf_lfp_util.ML" |
236 |
ML_file "Tools/BNF/bnf_lfp_tactics.ML" |
|
237 |
ML_file "Tools/BNF/bnf_lfp.ML" |
|
238 |
ML_file "Tools/BNF/bnf_lfp_compat.ML" |
|
55571 | 239 |
ML_file "Tools/BNF/bnf_lfp_rec_sugar_more.ML" |
58179 | 240 |
ML_file "Tools/BNF/bnf_lfp_size.ML" |
241 |
||
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
242 |
end |