1(*  Title:      HOL/BNF_Least_Fixpoint.thy
2    Author:     Dmitriy Traytel, TU Muenchen
3    Author:     Lorenz Panny, TU Muenchen
4    Author:     Jasmin Blanchette, TU Muenchen
5    Copyright   2012, 2013, 2014
6
7Least fixpoint (datatype) operation on bounded natural functors.
8*)
9
10section \<open>Least Fixpoint (Datatype) Operation on Bounded Natural Functors\<close>
11
12theory BNF_Least_Fixpoint
13imports BNF_Fixpoint_Base
14keywords
15  "datatype" :: thy_decl and
16  "datatype_compat" :: thy_decl
17begin
18
19lemma subset_emptyI: "(\<And>x. x \<in> A \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> {}"
20  by blast
21
22lemma image_Collect_subsetI: "(\<And>x. P x \<Longrightarrow> f x \<in> B) \<Longrightarrow> f ` {x. P x} \<subseteq> B"
23  by blast
24
25lemma Collect_restrict: "{x. x \<in> X \<and> P x} \<subseteq> X"
26  by auto
27
28lemma prop_restrict: "\<lbrakk>x \<in> Z; Z \<subseteq> {x. x \<in> X \<and> P x}\<rbrakk> \<Longrightarrow> P x"
29  by auto
30
31lemma underS_I: "\<lbrakk>i \<noteq> j; (i, j) \<in> R\<rbrakk> \<Longrightarrow> i \<in> underS R j"
32  unfolding underS_def by simp
33
34lemma underS_E: "i \<in> underS R j \<Longrightarrow> i \<noteq> j \<and> (i, j) \<in> R"
35  unfolding underS_def by simp
36
37lemma underS_Field: "i \<in> underS R j \<Longrightarrow> i \<in> Field R"
38  unfolding underS_def Field_def by auto
39
40lemma bij_betwE: "bij_betw f A B \<Longrightarrow> \<forall>a\<in>A. f a \<in> B"
41  unfolding bij_betw_def by auto
42
43lemma f_the_inv_into_f_bij_betw:
44  "bij_betw f A B \<Longrightarrow> (bij_betw f A B \<Longrightarrow> x \<in> B) \<Longrightarrow> f (the_inv_into A f x) = x"
45  unfolding bij_betw_def by (blast intro: f_the_inv_into_f)
46
47lemma ex_bij_betw: "|A| \<le>o (r :: 'b rel) \<Longrightarrow> \<exists>f B :: 'b set. bij_betw f B A"
48  by (subst (asm) internalize_card_of_ordLeq) (auto dest!: iffD2[OF card_of_ordIso ordIso_symmetric])
49
50lemma bij_betwI':
51  "\<lbrakk>\<And>x y. \<lbrakk>x \<in> X; y \<in> X\<rbrakk> \<Longrightarrow> (f x = f y) = (x = y);
52    \<And>x. x \<in> X \<Longrightarrow> f x \<in> Y;
53    \<And>y. y \<in> Y \<Longrightarrow> \<exists>x \<in> X. y = f x\<rbrakk> \<Longrightarrow> bij_betw f X Y"
54  unfolding bij_betw_def inj_on_def by blast
55
56lemma surj_fun_eq:
57  assumes surj_on: "f ` X = UNIV" and eq_on: "\<forall>x \<in> X. (g1 \<circ> f) x = (g2 \<circ> f) x"
58  shows "g1 = g2"
59proof (rule ext)
60  fix y
61  from surj_on obtain x where "x \<in> X" and "y = f x" by blast
62  thus "g1 y = g2 y" using eq_on by simp
63qed
64
65lemma Card_order_wo_rel: "Card_order r \<Longrightarrow> wo_rel r"
66  unfolding wo_rel_def card_order_on_def by blast
67
68lemma Cinfinite_limit: "\<lbrakk>x \<in> Field r; Cinfinite r\<rbrakk> \<Longrightarrow> \<exists>y \<in> Field r. x \<noteq> y \<and> (x, y) \<in> r"
69  unfolding cinfinite_def by (auto simp add: infinite_Card_order_limit)
70
71lemma Card_order_trans:
72  "\<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"
73  unfolding card_order_on_def well_order_on_def linear_order_on_def
74    partial_order_on_def preorder_on_def trans_def antisym_def by blast
75
76lemma Cinfinite_limit2:
77  assumes x1: "x1 \<in> Field r" and x2: "x2 \<in> Field r" and r: "Cinfinite r"
78  shows "\<exists>y \<in> Field r. (x1 \<noteq> y \<and> (x1, y) \<in> r) \<and> (x2 \<noteq> y \<and> (x2, y) \<in> r)"
79proof -
80  from r have trans: "trans r" and total: "Total r" and antisym: "antisym r"
81    unfolding card_order_on_def well_order_on_def linear_order_on_def
82      partial_order_on_def preorder_on_def by auto
83  obtain y1 where y1: "y1 \<in> Field r" "x1 \<noteq> y1" "(x1, y1) \<in> r"
84    using Cinfinite_limit[OF x1 r] by blast
85  obtain y2 where y2: "y2 \<in> Field r" "x2 \<noteq> y2" "(x2, y2) \<in> r"
86    using Cinfinite_limit[OF x2 r] by blast
87  show ?thesis
88  proof (cases "y1 = y2")
89    case True with y1 y2 show ?thesis by blast
90  next
91    case False
92    with y1(1) y2(1) total have "(y1, y2) \<in> r \<or> (y2, y1) \<in> r"
93      unfolding total_on_def by auto
94    thus ?thesis
95    proof
96      assume *: "(y1, y2) \<in> r"
97      with trans y1(3) have "(x1, y2) \<in> r" unfolding trans_def by blast
98      with False y1 y2 * antisym show ?thesis by (cases "x1 = y2") (auto simp: antisym_def)
99    next
100      assume *: "(y2, y1) \<in> r"
101      with trans y2(3) have "(x2, y1) \<in> r" unfolding trans_def by blast
102      with False y1 y2 * antisym show ?thesis by (cases "x2 = y1") (auto simp: antisym_def)
103    qed
104  qed
105qed
106
107lemma Cinfinite_limit_finite:
108  "\<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)"
109proof (induct X rule: finite_induct)
110  case empty thus ?case unfolding cinfinite_def using ex_in_conv[of "Field r"] finite.emptyI by auto
111next
112  case (insert x X)
113  then obtain y where y: "y \<in> Field r" "\<forall>x \<in> X. (x \<noteq> y \<and> (x, y) \<in> r)" by blast
114  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"
115    using Cinfinite_limit2[OF _ y(1) insert(5), of x] insert(4) by blast
116  show ?case
117    apply (intro bexI ballI)
118    apply (erule insertE)
119    apply hypsubst
120    apply (rule z(2))
121    using Card_order_trans[OF insert(5)[THEN conjunct2]] y(2) z(3)
122    apply blast
123    apply (rule z(1))
124    done
125qed
126
127lemma insert_subsetI: "\<lbrakk>x \<in> A; X \<subseteq> A\<rbrakk> \<Longrightarrow> insert x X \<subseteq> A"
128  by auto
129
130lemmas well_order_induct_imp = wo_rel.well_order_induct[of r "\<lambda>x. x \<in> Field r \<longrightarrow> P x" for r P]
131
132lemma meta_spec2:
133  assumes "(\<And>x y. PROP P x y)"
134  shows "PROP P x y"
135  by (rule assms)
136
137lemma nchotomy_relcomppE:
138  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"
139  shows P
140proof (rule relcompp.cases[OF assms(2)], hypsubst)
141  fix b assume "r a b" "s b c"
142  moreover from assms(1) obtain b' where "b = f b'" by blast
143  ultimately show P by (blast intro: assms(3))
144qed
145
146lemma predicate2D_vimage2p: "\<lbrakk>R \<le> vimage2p f g S; R x y\<rbrakk> \<Longrightarrow> S (f x) (g y)"
147  unfolding vimage2p_def by auto
148
149lemma ssubst_Pair_rhs: "\<lbrakk>(r, s) \<in> R; s' = s\<rbrakk> \<Longrightarrow> (r, s') \<in> R"
150  by (rule ssubst)
151
152lemma all_mem_range1:
153  "(\<And>y. y \<in> range f \<Longrightarrow> P y) \<equiv> (\<And>x. P (f x)) "
154  by (rule equal_intr_rule) fast+
155
156lemma all_mem_range2:
157  "(\<And>fa y. fa \<in> range f \<Longrightarrow> y \<in> range fa \<Longrightarrow> P y) \<equiv> (\<And>x xa. P (f x xa))"
158  by (rule equal_intr_rule) fast+
159
160lemma all_mem_range3:
161  "(\<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))"
162  by (rule equal_intr_rule) fast+
163
164lemma all_mem_range4:
165  "(\<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>
166   (\<And>x xa xb xc. P (f x xa xb xc))"
167  by (rule equal_intr_rule) fast+
168
169lemma all_mem_range5:
170  "(\<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>
171     y \<in> range fd \<Longrightarrow> P y) \<equiv>
172   (\<And>x xa xb xc xd. P (f x xa xb xc xd))"
173  by (rule equal_intr_rule) fast+
174
175lemma all_mem_range6:
176  "(\<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>
177     fe \<in> range fd \<Longrightarrow> ff \<in> range fe \<Longrightarrow> y \<in> range ff \<Longrightarrow> P y) \<equiv>
178   (\<And>x xa xb xc xd xe xf. P (f x xa xb xc xd xe xf))"
179  by (rule equal_intr_rule) (fastforce, fast)
180
181lemma all_mem_range7:
182  "(\<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>
183     fe \<in> range fd \<Longrightarrow> ff \<in> range fe \<Longrightarrow> fg \<in> range ff \<Longrightarrow> y \<in> range fg \<Longrightarrow> P y) \<equiv>
184   (\<And>x xa xb xc xd xe xf xg. P (f x xa xb xc xd xe xf xg))"
185  by (rule equal_intr_rule) (fastforce, fast)
186
187lemma all_mem_range8:
188  "(\<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>
189     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>
190   (\<And>x xa xb xc xd xe xf xg xh. P (f x xa xb xc xd xe xf xg xh))"
191  by (rule equal_intr_rule) (fastforce, fast)
192
193lemmas all_mem_range = all_mem_range1 all_mem_range2 all_mem_range3 all_mem_range4 all_mem_range5
194  all_mem_range6 all_mem_range7 all_mem_range8
195
196lemma pred_fun_True_id: "NO_MATCH id p \<Longrightarrow> pred_fun (\<lambda>x. True) p f = pred_fun (\<lambda>x. True) id (p \<circ> f)"
197  unfolding fun.pred_map unfolding comp_def id_def ..
198
199ML_file "Tools/BNF/bnf_lfp_util.ML"
200ML_file "Tools/BNF/bnf_lfp_tactics.ML"
201ML_file "Tools/BNF/bnf_lfp.ML"
202ML_file "Tools/BNF/bnf_lfp_compat.ML"
203ML_file "Tools/BNF/bnf_lfp_rec_sugar_more.ML"
204ML_file "Tools/BNF/bnf_lfp_size.ML"
205
206end
207