| author | Andreas Lochbihler |
| Fri, 29 Jul 2016 09:49:23 +0200 | |
| changeset 63561 | fba08009ff3e |
| parent 63540 | f8652d0534fa |
| child 63588 | d0e2bad67bd4 |
| permissions | -rw-r--r-- |
| 7700 | 1 |
(* Title: HOL/Inductive.thy |
| 10402 | 2 |
Author: Markus Wenzel, TU Muenchen |
| 11688 | 3 |
*) |
| 10727 | 4 |
|
| 60758 | 5 |
section \<open>Knaster-Tarski Fixpoint Theorem and inductive definitions\<close> |
| 1187 | 6 |
|
|
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52143
diff
changeset
|
7 |
theory Inductive |
|
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52143
diff
changeset
|
8 |
imports Complete_Lattices Ctr_Sugar |
|
46950
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents:
46947
diff
changeset
|
9 |
keywords |
|
56146
8453d35e4684
discontinued somewhat pointless "thy_script" keyword kind;
wenzelm
parents:
55604
diff
changeset
|
10 |
"inductive" "coinductive" "inductive_cases" "inductive_simps" :: thy_decl and |
|
8453d35e4684
discontinued somewhat pointless "thy_script" keyword kind;
wenzelm
parents:
55604
diff
changeset
|
11 |
"monos" and |
|
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52143
diff
changeset
|
12 |
"print_inductives" :: diag and |
|
58306
117ba6cbe414
renamed 'rep_datatype' to 'old_rep_datatype' (HOL)
blanchet
parents:
58187
diff
changeset
|
13 |
"old_rep_datatype" :: thy_goal and |
|
55575
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
blanchet
parents:
55534
diff
changeset
|
14 |
"primrec" :: thy_decl |
| 15131 | 15 |
begin |
| 10727 | 16 |
|
| 60758 | 17 |
subsection \<open>Least and greatest fixed points\<close> |
| 24915 | 18 |
|
|
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
19 |
context complete_lattice |
|
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
20 |
begin |
|
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
21 |
|
| 63400 | 22 |
definition lfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" \<comment> \<open>least fixed point\<close>
|
23 |
where "lfp f = Inf {u. f u \<le> u}"
|
|
| 24915 | 24 |
|
| 63400 | 25 |
definition gfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" \<comment> \<open>greatest fixed point\<close>
|
26 |
where "gfp f = Sup {u. u \<le> f u}"
|
|
| 24915 | 27 |
|
28 |
||
| 63400 | 29 |
subsection \<open>Proof of Knaster-Tarski Theorem using @{term lfp}\<close>
|
| 24915 | 30 |
|
| 63400 | 31 |
text \<open>@{term "lfp f"} is the least upper bound of the set @{term "{u. f u \<le> u}"}\<close>
|
| 24915 | 32 |
|
| 63400 | 33 |
lemma lfp_lowerbound: "f A \<le> A \<Longrightarrow> lfp f \<le> A" |
| 24915 | 34 |
by (auto simp add: lfp_def intro: Inf_lower) |
35 |
||
| 63400 | 36 |
lemma lfp_greatest: "(\<And>u. f u \<le> u \<Longrightarrow> A \<le> u) \<Longrightarrow> A \<le> lfp f" |
| 24915 | 37 |
by (auto simp add: lfp_def intro: Inf_greatest) |
38 |
||
|
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
39 |
end |
|
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
40 |
|
| 63400 | 41 |
lemma lfp_lemma2: "mono f \<Longrightarrow> f (lfp f) \<le> lfp f" |
| 24915 | 42 |
by (iprover intro: lfp_greatest order_trans monoD lfp_lowerbound) |
43 |
||
| 63400 | 44 |
lemma lfp_lemma3: "mono f \<Longrightarrow> lfp f \<le> f (lfp f)" |
| 24915 | 45 |
by (iprover intro: lfp_lemma2 monoD lfp_lowerbound) |
46 |
||
| 63400 | 47 |
lemma lfp_unfold: "mono f \<Longrightarrow> lfp f = f (lfp f)" |
| 24915 | 48 |
by (iprover intro: order_antisym lfp_lemma2 lfp_lemma3) |
49 |
||
50 |
lemma lfp_const: "lfp (\<lambda>x. t) = t" |
|
| 63400 | 51 |
by (rule lfp_unfold) (simp add: mono_def) |
| 24915 | 52 |
|
|
63561
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
53 |
lemma lfp_eqI: "\<lbrakk> mono F; F x = x; \<And>z. F z = z \<Longrightarrow> x \<le> z \<rbrakk> \<Longrightarrow> lfp F = x" |
|
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
54 |
by (rule antisym) (simp_all add: lfp_lowerbound lfp_unfold[symmetric]) |
|
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
55 |
|
| 24915 | 56 |
|
| 60758 | 57 |
subsection \<open>General induction rules for least fixed points\<close> |
| 24915 | 58 |
|
| 63400 | 59 |
lemma lfp_ordinal_induct [case_names mono step union]: |
| 61076 | 60 |
fixes f :: "'a::complete_lattice \<Rightarrow> 'a" |
|
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
61 |
assumes mono: "mono f" |
| 63400 | 62 |
and P_f: "\<And>S. P S \<Longrightarrow> S \<le> lfp f \<Longrightarrow> P (f S)" |
63 |
and P_Union: "\<And>M. \<forall>S\<in>M. P S \<Longrightarrow> P (Sup M)" |
|
|
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
64 |
shows "P (lfp f)" |
|
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
65 |
proof - |
|
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
66 |
let ?M = "{S. S \<le> lfp f \<and> P S}"
|
|
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
67 |
have "P (Sup ?M)" using P_Union by simp |
|
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
68 |
also have "Sup ?M = lfp f" |
|
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
69 |
proof (rule antisym) |
|
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
70 |
show "Sup ?M \<le> lfp f" by (blast intro: Sup_least) |
| 63400 | 71 |
then have "f (Sup ?M) \<le> f (lfp f)" |
72 |
by (rule mono [THEN monoD]) |
|
73 |
then have "f (Sup ?M) \<le> lfp f" |
|
74 |
using mono [THEN lfp_unfold] by simp |
|
75 |
then have "f (Sup ?M) \<in> ?M" |
|
76 |
using P_Union by simp (intro P_f Sup_least, auto) |
|
77 |
then have "f (Sup ?M) \<le> Sup ?M" |
|
78 |
by (rule Sup_upper) |
|
79 |
then show "lfp f \<le> Sup ?M" |
|
80 |
by (rule lfp_lowerbound) |
|
|
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
81 |
qed |
|
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
82 |
finally show ?thesis . |
| 63400 | 83 |
qed |
|
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
84 |
|
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
85 |
theorem lfp_induct: |
| 63400 | 86 |
assumes mono: "mono f" |
87 |
and ind: "f (inf (lfp f) P) \<le> P" |
|
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
88 |
shows "lfp f \<le> P" |
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
89 |
proof (induction rule: lfp_ordinal_induct) |
| 63400 | 90 |
case (step S) |
91 |
then show ?case |
|
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
92 |
by (intro order_trans[OF _ ind] monoD[OF mono]) auto |
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
93 |
qed (auto intro: mono Sup_least) |
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
94 |
|
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
95 |
lemma lfp_induct_set: |
| 63400 | 96 |
assumes lfp: "a \<in> lfp f" |
97 |
and mono: "mono f" |
|
98 |
and hyp: "\<And>x. x \<in> f (lfp f \<inter> {x. P x}) \<Longrightarrow> P x"
|
|
99 |
shows "P a" |
|
100 |
by (rule lfp_induct [THEN subsetD, THEN CollectD, OF mono _ lfp]) (auto intro: hyp) |
|
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
101 |
|
| 63400 | 102 |
lemma lfp_ordinal_induct_set: |
| 24915 | 103 |
assumes mono: "mono f" |
| 63400 | 104 |
and P_f: "\<And>S. P S \<Longrightarrow> P (f S)" |
105 |
and P_Union: "\<And>M. \<forall>S\<in>M. P S \<Longrightarrow> P (\<Union>M)" |
|
106 |
shows "P (lfp f)" |
|
|
46008
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents:
45907
diff
changeset
|
107 |
using assms by (rule lfp_ordinal_induct) |
| 24915 | 108 |
|
109 |
||
| 63400 | 110 |
text \<open>Definition forms of \<open>lfp_unfold\<close> and \<open>lfp_induct\<close>, to control unfolding.\<close> |
| 24915 | 111 |
|
| 63400 | 112 |
lemma def_lfp_unfold: "h \<equiv> lfp f \<Longrightarrow> mono f \<Longrightarrow> h = f h" |
| 45899 | 113 |
by (auto intro!: lfp_unfold) |
| 24915 | 114 |
|
| 63400 | 115 |
lemma def_lfp_induct: "A \<equiv> lfp f \<Longrightarrow> mono f \<Longrightarrow> f (inf A P) \<le> P \<Longrightarrow> A \<le> P" |
| 24915 | 116 |
by (blast intro: lfp_induct) |
117 |
||
| 63400 | 118 |
lemma def_lfp_induct_set: |
119 |
"A \<equiv> lfp f \<Longrightarrow> mono f \<Longrightarrow> a \<in> A \<Longrightarrow> (\<And>x. x \<in> f (A \<inter> {x. P x}) \<Longrightarrow> P x) \<Longrightarrow> P a"
|
|
| 24915 | 120 |
by (blast intro: lfp_induct_set) |
121 |
||
| 63400 | 122 |
text \<open>Monotonicity of \<open>lfp\<close>!\<close> |
123 |
lemma lfp_mono: "(\<And>Z. f Z \<le> g Z) \<Longrightarrow> lfp f \<le> lfp g" |
|
124 |
by (rule lfp_lowerbound [THEN lfp_greatest]) (blast intro: order_trans) |
|
| 24915 | 125 |
|
126 |
||
| 63400 | 127 |
subsection \<open>Proof of Knaster-Tarski Theorem using \<open>gfp\<close>\<close> |
| 24915 | 128 |
|
| 63400 | 129 |
text \<open>@{term "gfp f"} is the greatest lower bound of the set @{term "{u. u \<le> f u}"}\<close>
|
| 24915 | 130 |
|
| 63400 | 131 |
lemma gfp_upperbound: "X \<le> f X \<Longrightarrow> X \<le> gfp f" |
| 24915 | 132 |
by (auto simp add: gfp_def intro: Sup_upper) |
133 |
||
| 63400 | 134 |
lemma gfp_least: "(\<And>u. u \<le> f u \<Longrightarrow> u \<le> X) \<Longrightarrow> gfp f \<le> X" |
| 24915 | 135 |
by (auto simp add: gfp_def intro: Sup_least) |
136 |
||
| 63400 | 137 |
lemma gfp_lemma2: "mono f \<Longrightarrow> gfp f \<le> f (gfp f)" |
| 24915 | 138 |
by (iprover intro: gfp_least order_trans monoD gfp_upperbound) |
139 |
||
| 63400 | 140 |
lemma gfp_lemma3: "mono f \<Longrightarrow> f (gfp f) \<le> gfp f" |
| 24915 | 141 |
by (iprover intro: gfp_lemma2 monoD gfp_upperbound) |
142 |
||
| 63400 | 143 |
lemma gfp_unfold: "mono f \<Longrightarrow> gfp f = f (gfp f)" |
| 24915 | 144 |
by (iprover intro: order_antisym gfp_lemma2 gfp_lemma3) |
145 |
||
|
63561
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
146 |
lemma gfp_const: "gfp (\<lambda>x. t) = t" |
|
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
147 |
by (rule gfp_unfold) (simp add: mono_def) |
|
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
148 |
|
|
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
149 |
lemma gfp_eqI: "\<lbrakk> mono F; F x = x; \<And>z. F z = z \<Longrightarrow> z \<le> x \<rbrakk> \<Longrightarrow> gfp F = x" |
|
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
150 |
by (rule antisym) (simp_all add: gfp_upperbound gfp_unfold[symmetric]) |
|
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
151 |
|
| 24915 | 152 |
|
| 60758 | 153 |
subsection \<open>Coinduction rules for greatest fixed points\<close> |
| 24915 | 154 |
|
| 63400 | 155 |
text \<open>Weak version.\<close> |
156 |
lemma weak_coinduct: "a \<in> X \<Longrightarrow> X \<subseteq> f X \<Longrightarrow> a \<in> gfp f" |
|
| 45899 | 157 |
by (rule gfp_upperbound [THEN subsetD]) auto |
| 24915 | 158 |
|
| 63400 | 159 |
lemma weak_coinduct_image: "a \<in> X \<Longrightarrow> g`X \<subseteq> f (g`X) \<Longrightarrow> g a \<in> gfp f" |
| 45899 | 160 |
apply (erule gfp_upperbound [THEN subsetD]) |
161 |
apply (erule imageI) |
|
162 |
done |
|
| 24915 | 163 |
|
| 63400 | 164 |
lemma coinduct_lemma: "X \<le> f (sup X (gfp f)) \<Longrightarrow> mono f \<Longrightarrow> sup X (gfp f) \<le> f (sup X (gfp f))" |
| 24915 | 165 |
apply (frule gfp_lemma2) |
166 |
apply (drule mono_sup) |
|
167 |
apply (rule le_supI) |
|
168 |
apply assumption |
|
169 |
apply (rule order_trans) |
|
170 |
apply (rule order_trans) |
|
171 |
apply assumption |
|
172 |
apply (rule sup_ge2) |
|
173 |
apply assumption |
|
174 |
done |
|
175 |
||
| 63400 | 176 |
text \<open>Strong version, thanks to Coen and Frost.\<close> |
177 |
lemma coinduct_set: "mono f \<Longrightarrow> a \<in> X \<Longrightarrow> X \<subseteq> f (X \<union> gfp f) \<Longrightarrow> a \<in> gfp f" |
|
| 55604 | 178 |
by (rule weak_coinduct[rotated], rule coinduct_lemma) blast+ |
| 24915 | 179 |
|
| 63400 | 180 |
lemma gfp_fun_UnI2: "mono f \<Longrightarrow> a \<in> gfp f \<Longrightarrow> a \<in> f (X \<union> gfp f)" |
| 45899 | 181 |
by (blast dest: gfp_lemma2 mono_Un) |
| 24915 | 182 |
|
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
183 |
lemma gfp_ordinal_induct[case_names mono step union]: |
| 61076 | 184 |
fixes f :: "'a::complete_lattice \<Rightarrow> 'a" |
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
185 |
assumes mono: "mono f" |
| 63400 | 186 |
and P_f: "\<And>S. P S \<Longrightarrow> gfp f \<le> S \<Longrightarrow> P (f S)" |
187 |
and P_Union: "\<And>M. \<forall>S\<in>M. P S \<Longrightarrow> P (Inf M)" |
|
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
188 |
shows "P (gfp f)" |
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
189 |
proof - |
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
190 |
let ?M = "{S. gfp f \<le> S \<and> P S}"
|
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
191 |
have "P (Inf ?M)" using P_Union by simp |
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
192 |
also have "Inf ?M = gfp f" |
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
193 |
proof (rule antisym) |
| 63400 | 194 |
show "gfp f \<le> Inf ?M" |
195 |
by (blast intro: Inf_greatest) |
|
196 |
then have "f (gfp f) \<le> f (Inf ?M)" |
|
197 |
by (rule mono [THEN monoD]) |
|
198 |
then have "gfp f \<le> f (Inf ?M)" |
|
199 |
using mono [THEN gfp_unfold] by simp |
|
200 |
then have "f (Inf ?M) \<in> ?M" |
|
201 |
using P_Union by simp (intro P_f Inf_greatest, auto) |
|
202 |
then have "Inf ?M \<le> f (Inf ?M)" |
|
203 |
by (rule Inf_lower) |
|
204 |
then show "Inf ?M \<le> gfp f" |
|
205 |
by (rule gfp_upperbound) |
|
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
206 |
qed |
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
207 |
finally show ?thesis . |
| 63400 | 208 |
qed |
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
209 |
|
| 63400 | 210 |
lemma coinduct: |
211 |
assumes mono: "mono f" |
|
212 |
and ind: "X \<le> f (sup X (gfp f))" |
|
213 |
shows "X \<le> gfp f" |
|
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
214 |
proof (induction rule: gfp_ordinal_induct) |
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
215 |
case (step S) then show ?case |
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
216 |
by (intro order_trans[OF ind _] monoD[OF mono]) auto |
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
217 |
qed (auto intro: mono Inf_greatest) |
| 24915 | 218 |
|
| 63400 | 219 |
|
| 60758 | 220 |
subsection \<open>Even Stronger Coinduction Rule, by Martin Coen\<close> |
| 24915 | 221 |
|
| 63400 | 222 |
text \<open>Weakens the condition @{term "X \<subseteq> f X"} to one expressed using both
|
| 60758 | 223 |
@{term lfp} and @{term gfp}\<close>
|
| 63400 | 224 |
lemma coinduct3_mono_lemma: "mono f \<Longrightarrow> mono (\<lambda>x. f x \<union> X \<union> B)" |
225 |
by (iprover intro: subset_refl monoI Un_mono monoD) |
|
| 24915 | 226 |
|
227 |
lemma coinduct3_lemma: |
|
| 63400 | 228 |
"X \<subseteq> f (lfp (\<lambda>x. f x \<union> X \<union> gfp f)) \<Longrightarrow> mono f \<Longrightarrow> |
229 |
lfp (\<lambda>x. f x \<union> X \<union> gfp f) \<subseteq> f (lfp (\<lambda>x. f x \<union> X \<union> gfp f))" |
|
230 |
apply (rule subset_trans) |
|
231 |
apply (erule coinduct3_mono_lemma [THEN lfp_lemma3]) |
|
232 |
apply (rule Un_least [THEN Un_least]) |
|
233 |
apply (rule subset_refl, assumption) |
|
234 |
apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption) |
|
235 |
apply (rule monoD, assumption) |
|
236 |
apply (subst coinduct3_mono_lemma [THEN lfp_unfold], auto) |
|
237 |
done |
|
| 24915 | 238 |
|
| 63400 | 239 |
lemma coinduct3: "mono f \<Longrightarrow> a \<in> X \<Longrightarrow> X \<subseteq> f (lfp (\<lambda>x. f x \<union> X \<union> gfp f)) \<Longrightarrow> a \<in> gfp f" |
240 |
apply (rule coinduct3_lemma [THEN [2] weak_coinduct]) |
|
241 |
apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst]) |
|
242 |
apply simp_all |
|
243 |
done |
|
| 24915 | 244 |
|
| 63400 | 245 |
text \<open>Definition forms of \<open>gfp_unfold\<close> and \<open>coinduct\<close>, to control unfolding.\<close> |
| 24915 | 246 |
|
| 63400 | 247 |
lemma def_gfp_unfold: "A \<equiv> gfp f \<Longrightarrow> mono f \<Longrightarrow> A = f A" |
| 45899 | 248 |
by (auto intro!: gfp_unfold) |
| 24915 | 249 |
|
| 63400 | 250 |
lemma def_coinduct: "A \<equiv> gfp f \<Longrightarrow> mono f \<Longrightarrow> X \<le> f (sup X A) \<Longrightarrow> X \<le> A" |
| 45899 | 251 |
by (iprover intro!: coinduct) |
| 24915 | 252 |
|
| 63400 | 253 |
lemma def_coinduct_set: "A \<equiv> gfp f \<Longrightarrow> mono f \<Longrightarrow> a \<in> X \<Longrightarrow> X \<subseteq> f (X \<union> A) \<Longrightarrow> a \<in> A" |
| 45899 | 254 |
by (auto intro!: coinduct_set) |
| 24915 | 255 |
|
256 |
lemma def_Collect_coinduct: |
|
| 63400 | 257 |
"A \<equiv> gfp (\<lambda>w. Collect (P w)) \<Longrightarrow> mono (\<lambda>w. Collect (P w)) \<Longrightarrow> a \<in> X \<Longrightarrow> |
258 |
(\<And>z. z \<in> X \<Longrightarrow> P (X \<union> A) z) \<Longrightarrow> a \<in> A" |
|
| 45899 | 259 |
by (erule def_coinduct_set) auto |
| 24915 | 260 |
|
| 63400 | 261 |
lemma def_coinduct3: "A \<equiv> gfp f \<Longrightarrow> mono f \<Longrightarrow> a \<in> X \<Longrightarrow> X \<subseteq> f (lfp (\<lambda>x. f x \<union> X \<union> A)) \<Longrightarrow> a \<in> A" |
| 45899 | 262 |
by (auto intro!: coinduct3) |
| 24915 | 263 |
|
| 63400 | 264 |
text \<open>Monotonicity of @{term gfp}!\<close>
|
265 |
lemma gfp_mono: "(\<And>Z. f Z \<le> g Z) \<Longrightarrow> gfp f \<le> gfp g" |
|
266 |
by (rule gfp_upperbound [THEN gfp_least]) (blast intro: order_trans) |
|
267 |
||
| 24915 | 268 |
|
| 60758 | 269 |
subsection \<open>Rules for fixed point calculus\<close> |
| 60173 | 270 |
|
271 |
lemma lfp_rolling: |
|
272 |
assumes "mono g" "mono f" |
|
273 |
shows "g (lfp (\<lambda>x. f (g x))) = lfp (\<lambda>x. g (f x))" |
|
274 |
proof (rule antisym) |
|
275 |
have *: "mono (\<lambda>x. f (g x))" |
|
276 |
using assms by (auto simp: mono_def) |
|
277 |
show "lfp (\<lambda>x. g (f x)) \<le> g (lfp (\<lambda>x. f (g x)))" |
|
278 |
by (rule lfp_lowerbound) (simp add: lfp_unfold[OF *, symmetric]) |
|
279 |
show "g (lfp (\<lambda>x. f (g x))) \<le> lfp (\<lambda>x. g (f x))" |
|
280 |
proof (rule lfp_greatest) |
|
| 63400 | 281 |
fix u |
| 63540 | 282 |
assume u: "g (f u) \<le> u" |
283 |
then have "g (lfp (\<lambda>x. f (g x))) \<le> g (f u)" |
|
| 60173 | 284 |
by (intro assms[THEN monoD] lfp_lowerbound) |
| 63540 | 285 |
with u show "g (lfp (\<lambda>x. f (g x))) \<le> u" |
| 60173 | 286 |
by auto |
287 |
qed |
|
288 |
qed |
|
289 |
||
290 |
lemma lfp_lfp: |
|
291 |
assumes f: "\<And>x y w z. x \<le> y \<Longrightarrow> w \<le> z \<Longrightarrow> f x w \<le> f y z" |
|
292 |
shows "lfp (\<lambda>x. lfp (f x)) = lfp (\<lambda>x. f x x)" |
|
293 |
proof (rule antisym) |
|
294 |
have *: "mono (\<lambda>x. f x x)" |
|
295 |
by (blast intro: monoI f) |
|
296 |
show "lfp (\<lambda>x. lfp (f x)) \<le> lfp (\<lambda>x. f x x)" |
|
297 |
by (intro lfp_lowerbound) (simp add: lfp_unfold[OF *, symmetric]) |
|
298 |
show "lfp (\<lambda>x. lfp (f x)) \<ge> lfp (\<lambda>x. f x x)" (is "?F \<ge> _") |
|
299 |
proof (intro lfp_lowerbound) |
|
300 |
have *: "?F = lfp (f ?F)" |
|
301 |
by (rule lfp_unfold) (blast intro: monoI lfp_mono f) |
|
302 |
also have "\<dots> = f ?F (lfp (f ?F))" |
|
303 |
by (rule lfp_unfold) (blast intro: monoI lfp_mono f) |
|
304 |
finally show "f ?F ?F \<le> ?F" |
|
305 |
by (simp add: *[symmetric]) |
|
306 |
qed |
|
307 |
qed |
|
308 |
||
309 |
lemma gfp_rolling: |
|
310 |
assumes "mono g" "mono f" |
|
311 |
shows "g (gfp (\<lambda>x. f (g x))) = gfp (\<lambda>x. g (f x))" |
|
312 |
proof (rule antisym) |
|
313 |
have *: "mono (\<lambda>x. f (g x))" |
|
314 |
using assms by (auto simp: mono_def) |
|
315 |
show "g (gfp (\<lambda>x. f (g x))) \<le> gfp (\<lambda>x. g (f x))" |
|
316 |
by (rule gfp_upperbound) (simp add: gfp_unfold[OF *, symmetric]) |
|
317 |
show "gfp (\<lambda>x. g (f x)) \<le> g (gfp (\<lambda>x. f (g x)))" |
|
318 |
proof (rule gfp_least) |
|
| 63540 | 319 |
fix u |
320 |
assume u: "u \<le> g (f u)" |
|
321 |
then have "g (f u) \<le> g (gfp (\<lambda>x. f (g x)))" |
|
| 60173 | 322 |
by (intro assms[THEN monoD] gfp_upperbound) |
| 63540 | 323 |
with u show "u \<le> g (gfp (\<lambda>x. f (g x)))" |
| 60173 | 324 |
by auto |
325 |
qed |
|
326 |
qed |
|
327 |
||
328 |
lemma gfp_gfp: |
|
329 |
assumes f: "\<And>x y w z. x \<le> y \<Longrightarrow> w \<le> z \<Longrightarrow> f x w \<le> f y z" |
|
330 |
shows "gfp (\<lambda>x. gfp (f x)) = gfp (\<lambda>x. f x x)" |
|
331 |
proof (rule antisym) |
|
332 |
have *: "mono (\<lambda>x. f x x)" |
|
333 |
by (blast intro: monoI f) |
|
334 |
show "gfp (\<lambda>x. f x x) \<le> gfp (\<lambda>x. gfp (f x))" |
|
335 |
by (intro gfp_upperbound) (simp add: gfp_unfold[OF *, symmetric]) |
|
336 |
show "gfp (\<lambda>x. gfp (f x)) \<le> gfp (\<lambda>x. f x x)" (is "?F \<le> _") |
|
337 |
proof (intro gfp_upperbound) |
|
338 |
have *: "?F = gfp (f ?F)" |
|
339 |
by (rule gfp_unfold) (blast intro: monoI gfp_mono f) |
|
340 |
also have "\<dots> = f ?F (gfp (f ?F))" |
|
341 |
by (rule gfp_unfold) (blast intro: monoI gfp_mono f) |
|
342 |
finally show "?F \<le> f ?F ?F" |
|
343 |
by (simp add: *[symmetric]) |
|
344 |
qed |
|
345 |
qed |
|
| 24915 | 346 |
|
| 63400 | 347 |
|
| 60758 | 348 |
subsection \<open>Inductive predicates and sets\<close> |
| 11688 | 349 |
|
| 60758 | 350 |
text \<open>Package setup.\<close> |
| 10402 | 351 |
|
| 61337 | 352 |
lemmas basic_monos = |
| 22218 | 353 |
subset_refl imp_refl disj_mono conj_mono ex_mono all_mono if_bool_eq_conj |
| 11688 | 354 |
Collect_mono in_mono vimage_mono |
355 |
||
| 48891 | 356 |
ML_file "Tools/inductive.ML" |
|
21018
e6b8d6784792
Added new package for inductive definitions, moved old package
berghofe
parents:
20604
diff
changeset
|
357 |
|
| 61337 | 358 |
lemmas [mono] = |
| 22218 | 359 |
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:
32701
diff
changeset
|
360 |
imp_mono not_mono |
|
21018
e6b8d6784792
Added new package for inductive definitions, moved old package
berghofe
parents:
20604
diff
changeset
|
361 |
Ball_def Bex_def |
|
e6b8d6784792
Added new package for inductive definitions, moved old package
berghofe
parents:
20604
diff
changeset
|
362 |
induct_rulify_fallback |
|
e6b8d6784792
Added new package for inductive definitions, moved old package
berghofe
parents:
20604
diff
changeset
|
363 |
|
| 11688 | 364 |
|
| 60758 | 365 |
subsection \<open>Inductive datatypes and primitive recursion\<close> |
| 11688 | 366 |
|
| 60758 | 367 |
text \<open>Package setup.\<close> |
| 11825 | 368 |
|
|
58112
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents:
56146
diff
changeset
|
369 |
ML_file "Tools/Old_Datatype/old_datatype_aux.ML" |
|
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents:
56146
diff
changeset
|
370 |
ML_file "Tools/Old_Datatype/old_datatype_prop.ML" |
|
58187
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
blanchet
parents:
58112
diff
changeset
|
371 |
ML_file "Tools/Old_Datatype/old_datatype_data.ML" |
|
58112
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents:
56146
diff
changeset
|
372 |
ML_file "Tools/Old_Datatype/old_rep_datatype.ML" |
|
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents:
56146
diff
changeset
|
373 |
ML_file "Tools/Old_Datatype/old_datatype_codegen.ML" |
|
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents:
56146
diff
changeset
|
374 |
ML_file "Tools/Old_Datatype/old_primrec.ML" |
|
12437
6d4e02b6dd43
Moved code generator setup from Recdef to Inductive.
berghofe
parents:
12023
diff
changeset
|
375 |
|
|
55575
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
blanchet
parents:
55534
diff
changeset
|
376 |
ML_file "Tools/BNF/bnf_fp_rec_sugar_util.ML" |
|
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
blanchet
parents:
55534
diff
changeset
|
377 |
ML_file "Tools/BNF/bnf_lfp_rec_sugar.ML" |
|
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
blanchet
parents:
55534
diff
changeset
|
378 |
|
|
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61952
diff
changeset
|
379 |
text \<open>Lambda-abstractions with pattern matching:\<close> |
|
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61952
diff
changeset
|
380 |
syntax (ASCII) |
|
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61952
diff
changeset
|
381 |
"_lam_pats_syntax" :: "cases_syn \<Rightarrow> 'a \<Rightarrow> 'b" ("(%_)" 10)
|
| 23526 | 382 |
syntax |
|
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61952
diff
changeset
|
383 |
"_lam_pats_syntax" :: "cases_syn \<Rightarrow> 'a \<Rightarrow> 'b" ("(\<lambda>_)" 10)
|
| 60758 | 384 |
parse_translation \<open> |
| 52143 | 385 |
let |
386 |
fun fun_tr ctxt [cs] = |
|
387 |
let |
|
388 |
val x = Syntax.free (fst (Name.variant "x" (Term.declare_term_frees cs Name.context))); |
|
389 |
val ft = Case_Translation.case_tr true ctxt [x, cs]; |
|
390 |
in lambda x ft end |
|
391 |
in [(@{syntax_const "_lam_pats_syntax"}, fun_tr)] end
|
|
| 60758 | 392 |
\<close> |
| 23526 | 393 |
|
394 |
end |