author | huffman |
Wed, 17 Feb 2010 09:08:58 -0800 | |
changeset 35169 | 31cbcb019003 |
parent 33586 | 0e745228d605 |
child 35901 | 12f09bf2c77f |
permissions | -rw-r--r-- |
27409 | 1 |
(* Title: HOLCF/Algebraic.thy |
2 |
Author: Brian Huffman |
|
3 |
*) |
|
4 |
||
5 |
header {* Algebraic deflations *} |
|
6 |
||
7 |
theory Algebraic |
|
8 |
imports Completion Fix Eventual |
|
9 |
begin |
|
10 |
||
11 |
subsection {* Constructing finite deflations by iteration *} |
|
12 |
||
13 |
lemma finite_deflation_imp_deflation: |
|
14 |
"finite_deflation d \<Longrightarrow> deflation d" |
|
15 |
unfolding finite_deflation_def by simp |
|
16 |
||
17 |
lemma le_Suc_induct: |
|
18 |
assumes le: "i \<le> j" |
|
19 |
assumes step: "\<And>i. P i (Suc i)" |
|
20 |
assumes refl: "\<And>i. P i i" |
|
21 |
assumes trans: "\<And>i j k. \<lbrakk>P i j; P j k\<rbrakk> \<Longrightarrow> P i k" |
|
22 |
shows "P i j" |
|
23 |
proof (cases "i = j") |
|
24 |
assume "i = j" |
|
25 |
thus "P i j" by (simp add: refl) |
|
26 |
next |
|
27 |
assume "i \<noteq> j" |
|
28 |
with le have "i < j" by simp |
|
29 |
thus "P i j" using step trans by (rule less_Suc_induct) |
|
30 |
qed |
|
31 |
||
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
32 |
definition |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
33 |
eventual_iterate :: "('a \<rightarrow> 'a::cpo) \<Rightarrow> ('a \<rightarrow> 'a)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
34 |
where |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
35 |
"eventual_iterate f = eventual (\<lambda>n. iterate n\<cdot>f)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
36 |
|
27409 | 37 |
text {* A pre-deflation is like a deflation, but not idempotent. *} |
38 |
||
39 |
locale pre_deflation = |
|
40 |
fixes f :: "'a \<rightarrow> 'a::cpo" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
41 |
assumes below: "\<And>x. f\<cdot>x \<sqsubseteq> x" |
27409 | 42 |
assumes finite_range: "finite (range (\<lambda>x. f\<cdot>x))" |
43 |
begin |
|
44 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
45 |
lemma iterate_below: "iterate i\<cdot>f\<cdot>x \<sqsubseteq> x" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
46 |
by (induct i, simp_all add: below_trans [OF below]) |
27409 | 47 |
|
48 |
lemma iterate_fixed: "f\<cdot>x = x \<Longrightarrow> iterate i\<cdot>f\<cdot>x = x" |
|
49 |
by (induct i, simp_all) |
|
50 |
||
51 |
lemma antichain_iterate_app: "i \<le> j \<Longrightarrow> iterate j\<cdot>f\<cdot>x \<sqsubseteq> iterate i\<cdot>f\<cdot>x" |
|
52 |
apply (erule le_Suc_induct) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
53 |
apply (simp add: below) |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
54 |
apply (rule below_refl) |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
55 |
apply (erule (1) below_trans) |
27409 | 56 |
done |
57 |
||
58 |
lemma finite_range_iterate_app: "finite (range (\<lambda>i. iterate i\<cdot>f\<cdot>x))" |
|
59 |
proof (rule finite_subset) |
|
60 |
show "range (\<lambda>i. iterate i\<cdot>f\<cdot>x) \<subseteq> insert x (range (\<lambda>x. f\<cdot>x))" |
|
61 |
by (clarify, case_tac i, simp_all) |
|
62 |
show "finite (insert x (range (\<lambda>x. f\<cdot>x)))" |
|
63 |
by (simp add: finite_range) |
|
64 |
qed |
|
65 |
||
66 |
lemma eventually_constant_iterate_app: |
|
67 |
"eventually_constant (\<lambda>i. iterate i\<cdot>f\<cdot>x)" |
|
68 |
unfolding eventually_constant_def MOST_nat_le |
|
69 |
proof - |
|
70 |
let ?Y = "\<lambda>i. iterate i\<cdot>f\<cdot>x" |
|
71 |
have "\<exists>j. \<forall>k. ?Y j \<sqsubseteq> ?Y k" |
|
72 |
apply (rule finite_range_has_max) |
|
73 |
apply (erule antichain_iterate_app) |
|
74 |
apply (rule finite_range_iterate_app) |
|
75 |
done |
|
76 |
then obtain j where j: "\<And>k. ?Y j \<sqsubseteq> ?Y k" by fast |
|
77 |
show "\<exists>z m. \<forall>n\<ge>m. ?Y n = z" |
|
78 |
proof (intro exI allI impI) |
|
79 |
fix k |
|
80 |
assume "j \<le> k" |
|
81 |
hence "?Y k \<sqsubseteq> ?Y j" by (rule antichain_iterate_app) |
|
82 |
also have "?Y j \<sqsubseteq> ?Y k" by (rule j) |
|
83 |
finally show "?Y k = ?Y j" . |
|
84 |
qed |
|
85 |
qed |
|
86 |
||
87 |
lemma eventually_constant_iterate: |
|
88 |
"eventually_constant (\<lambda>n. iterate n\<cdot>f)" |
|
89 |
proof - |
|
90 |
have "\<forall>y\<in>range (\<lambda>x. f\<cdot>x). eventually_constant (\<lambda>i. iterate i\<cdot>f\<cdot>y)" |
|
91 |
by (simp add: eventually_constant_iterate_app) |
|
92 |
hence "\<forall>y\<in>range (\<lambda>x. f\<cdot>x). MOST i. MOST j. iterate j\<cdot>f\<cdot>y = iterate i\<cdot>f\<cdot>y" |
|
93 |
unfolding eventually_constant_MOST_MOST . |
|
94 |
hence "MOST i. MOST j. \<forall>y\<in>range (\<lambda>x. f\<cdot>x). iterate j\<cdot>f\<cdot>y = iterate i\<cdot>f\<cdot>y" |
|
95 |
by (simp only: MOST_finite_Ball_distrib [OF finite_range]) |
|
96 |
hence "MOST i. MOST j. \<forall>x. iterate j\<cdot>f\<cdot>(f\<cdot>x) = iterate i\<cdot>f\<cdot>(f\<cdot>x)" |
|
97 |
by simp |
|
98 |
hence "MOST i. MOST j. \<forall>x. iterate (Suc j)\<cdot>f\<cdot>x = iterate (Suc i)\<cdot>f\<cdot>x" |
|
99 |
by (simp only: iterate_Suc2) |
|
100 |
hence "MOST i. MOST j. iterate (Suc j)\<cdot>f = iterate (Suc i)\<cdot>f" |
|
101 |
by (simp only: expand_cfun_eq) |
|
102 |
hence "eventually_constant (\<lambda>i. iterate (Suc i)\<cdot>f)" |
|
103 |
unfolding eventually_constant_MOST_MOST . |
|
104 |
thus "eventually_constant (\<lambda>i. iterate i\<cdot>f)" |
|
105 |
by (rule eventually_constant_SucD) |
|
106 |
qed |
|
107 |
||
108 |
abbreviation |
|
109 |
d :: "'a \<rightarrow> 'a" |
|
110 |
where |
|
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
111 |
"d \<equiv> eventual_iterate f" |
27409 | 112 |
|
113 |
lemma MOST_d: "MOST n. P (iterate n\<cdot>f) \<Longrightarrow> P d" |
|
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
114 |
unfolding eventual_iterate_def |
27409 | 115 |
using eventually_constant_iterate by (rule MOST_eventual) |
116 |
||
117 |
lemma f_d: "f\<cdot>(d\<cdot>x) = d\<cdot>x" |
|
118 |
apply (rule MOST_d) |
|
119 |
apply (subst iterate_Suc [symmetric]) |
|
120 |
apply (rule eventually_constant_MOST_Suc_eq) |
|
121 |
apply (rule eventually_constant_iterate_app) |
|
122 |
done |
|
123 |
||
124 |
lemma d_fixed_iff: "d\<cdot>x = x \<longleftrightarrow> f\<cdot>x = x" |
|
125 |
proof |
|
126 |
assume "d\<cdot>x = x" |
|
127 |
with f_d [where x=x] |
|
128 |
show "f\<cdot>x = x" by simp |
|
129 |
next |
|
130 |
assume f: "f\<cdot>x = x" |
|
131 |
have "\<forall>n. iterate n\<cdot>f\<cdot>x = x" |
|
132 |
by (rule allI, rule nat.induct, simp, simp add: f) |
|
133 |
hence "MOST n. iterate n\<cdot>f\<cdot>x = x" |
|
134 |
by (rule ALL_MOST) |
|
135 |
thus "d\<cdot>x = x" |
|
136 |
by (rule MOST_d) |
|
137 |
qed |
|
138 |
||
139 |
lemma finite_deflation_d: "finite_deflation d" |
|
140 |
proof |
|
141 |
fix x :: 'a |
|
142 |
have "d \<in> range (\<lambda>n. iterate n\<cdot>f)" |
|
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
143 |
unfolding eventual_iterate_def |
27409 | 144 |
using eventually_constant_iterate |
145 |
by (rule eventual_mem_range) |
|
146 |
then obtain n where n: "d = iterate n\<cdot>f" .. |
|
147 |
have "iterate n\<cdot>f\<cdot>(d\<cdot>x) = d\<cdot>x" |
|
148 |
using f_d by (rule iterate_fixed) |
|
149 |
thus "d\<cdot>(d\<cdot>x) = d\<cdot>x" |
|
150 |
by (simp add: n) |
|
151 |
next |
|
152 |
fix x :: 'a |
|
153 |
show "d\<cdot>x \<sqsubseteq> x" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
154 |
by (rule MOST_d, simp add: iterate_below) |
27409 | 155 |
next |
156 |
from finite_range |
|
157 |
have "finite {x. f\<cdot>x = x}" |
|
158 |
by (rule finite_range_imp_finite_fixes) |
|
159 |
thus "finite {x. d\<cdot>x = x}" |
|
160 |
by (simp add: d_fixed_iff) |
|
161 |
qed |
|
162 |
||
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
163 |
lemma deflation_d: "deflation d" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
164 |
using finite_deflation_d |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
165 |
by (rule finite_deflation_imp_deflation) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
166 |
|
27409 | 167 |
end |
168 |
||
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
169 |
lemma finite_deflation_eventual_iterate: |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
170 |
"pre_deflation d \<Longrightarrow> finite_deflation (eventual_iterate d)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
171 |
by (rule pre_deflation.finite_deflation_d) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
172 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
173 |
lemma pre_deflation_oo: |
28611 | 174 |
assumes "finite_deflation d" |
27409 | 175 |
assumes f: "\<And>x. f\<cdot>x \<sqsubseteq> x" |
176 |
shows "pre_deflation (d oo f)" |
|
177 |
proof |
|
29237 | 178 |
interpret d: finite_deflation d by fact |
27409 | 179 |
fix x |
180 |
show "\<And>x. (d oo f)\<cdot>x \<sqsubseteq> x" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
181 |
by (simp, rule below_trans [OF d.below f]) |
27409 | 182 |
show "finite (range (\<lambda>x. (d oo f)\<cdot>x))" |
183 |
by (rule finite_subset [OF _ d.finite_range], auto) |
|
184 |
qed |
|
185 |
||
186 |
lemma eventual_iterate_oo_fixed_iff: |
|
28611 | 187 |
assumes "finite_deflation d" |
27409 | 188 |
assumes f: "\<And>x. f\<cdot>x \<sqsubseteq> x" |
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
189 |
shows "eventual_iterate (d oo f)\<cdot>x = x \<longleftrightarrow> d\<cdot>x = x \<and> f\<cdot>x = x" |
27409 | 190 |
proof - |
29237 | 191 |
interpret d: finite_deflation d by fact |
27409 | 192 |
let ?e = "d oo f" |
29237 | 193 |
interpret e: pre_deflation "d oo f" |
27409 | 194 |
using `finite_deflation d` f |
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
195 |
by (rule pre_deflation_oo) |
27409 | 196 |
let ?g = "eventual (\<lambda>n. iterate n\<cdot>?e)" |
197 |
show ?thesis |
|
198 |
apply (subst e.d_fixed_iff) |
|
199 |
apply simp |
|
200 |
apply safe |
|
201 |
apply (erule subst) |
|
202 |
apply (rule d.idem) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
203 |
apply (rule below_antisym) |
27409 | 204 |
apply (rule f) |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
205 |
apply (erule subst, rule d.below) |
27409 | 206 |
apply simp |
207 |
done |
|
208 |
qed |
|
209 |
||
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
210 |
lemma eventual_mono: |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
211 |
assumes A: "eventually_constant A" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
212 |
assumes B: "eventually_constant B" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
213 |
assumes below: "\<And>n. A n \<sqsubseteq> B n" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
214 |
shows "eventual A \<sqsubseteq> eventual B" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
215 |
proof - |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
216 |
from A have "MOST n. A n = eventual A" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
217 |
by (rule MOST_eq_eventual) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
218 |
then have "MOST n. eventual A \<sqsubseteq> B n" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
219 |
by (rule MOST_mono) (erule subst, rule below) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
220 |
with B show "eventual A \<sqsubseteq> eventual B" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
221 |
by (rule MOST_eventual) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
222 |
qed |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
223 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
224 |
lemma eventual_iterate_mono: |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
225 |
assumes f: "pre_deflation f" and g: "pre_deflation g" and "f \<sqsubseteq> g" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
226 |
shows "eventual_iterate f \<sqsubseteq> eventual_iterate g" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
227 |
unfolding eventual_iterate_def |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
228 |
apply (rule eventual_mono) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
229 |
apply (rule pre_deflation.eventually_constant_iterate [OF f]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
230 |
apply (rule pre_deflation.eventually_constant_iterate [OF g]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
231 |
apply (rule monofun_cfun_arg [OF `f \<sqsubseteq> g`]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
232 |
done |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
233 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
234 |
lemma cont2cont_eventual_iterate_oo: |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
235 |
assumes d: "finite_deflation d" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
236 |
assumes cont: "cont f" and below: "\<And>x y. f x\<cdot>y \<sqsubseteq> y" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
237 |
shows "cont (\<lambda>x. eventual_iterate (d oo f x))" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
238 |
(is "cont ?e") |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
239 |
proof (rule contI2) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
240 |
show "monofun ?e" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
241 |
apply (rule monofunI) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
242 |
apply (rule eventual_iterate_mono) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
243 |
apply (rule pre_deflation_oo [OF d below]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
244 |
apply (rule pre_deflation_oo [OF d below]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
245 |
apply (rule monofun_cfun_arg) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
246 |
apply (erule cont2monofunE [OF cont]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
247 |
done |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
248 |
next |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
249 |
fix Y :: "nat \<Rightarrow> 'b" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
250 |
assume Y: "chain Y" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
251 |
with cont have fY: "chain (\<lambda>i. f (Y i))" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
252 |
by (rule ch2ch_cont) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
253 |
assume eY: "chain (\<lambda>i. ?e (Y i))" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
254 |
have lub_below: "\<And>x. f (\<Squnion>i. Y i)\<cdot>x \<sqsubseteq> x" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
255 |
by (rule admD [OF _ Y], simp add: cont, rule below) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
256 |
have "deflation (?e (\<Squnion>i. Y i))" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
257 |
apply (rule pre_deflation.deflation_d) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
258 |
apply (rule pre_deflation_oo [OF d lub_below]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
259 |
done |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
260 |
then show "?e (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. ?e (Y i))" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
261 |
proof (rule deflation.belowI) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
262 |
fix x :: 'a |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
263 |
assume "?e (\<Squnion>i. Y i)\<cdot>x = x" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
264 |
hence "d\<cdot>x = x" and "f (\<Squnion>i. Y i)\<cdot>x = x" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
265 |
by (simp_all add: eventual_iterate_oo_fixed_iff [OF d lub_below]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
266 |
hence "(\<Squnion>i. f (Y i)\<cdot>x) = x" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
267 |
apply (simp only: cont2contlubE [OF cont Y]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
268 |
apply (simp only: contlub_cfun_fun [OF fY]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
269 |
done |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
270 |
have "compact (d\<cdot>x)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
271 |
using d by (rule finite_deflation.compact) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
272 |
then have "compact x" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
273 |
using `d\<cdot>x = x` by simp |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
274 |
then have "compact (\<Squnion>i. f (Y i)\<cdot>x)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
275 |
using `(\<Squnion>i. f (Y i)\<cdot>x) = x` by simp |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
276 |
then have "\<exists>n. max_in_chain n (\<lambda>i. f (Y i)\<cdot>x)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
277 |
by - (rule compact_imp_max_in_chain, simp add: fY, assumption) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
278 |
then obtain n where n: "max_in_chain n (\<lambda>i. f (Y i)\<cdot>x)" .. |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
279 |
then have "f (Y n)\<cdot>x = x" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
280 |
using `(\<Squnion>i. f (Y i)\<cdot>x) = x` fY by (simp add: maxinch_is_thelub) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
281 |
with `d\<cdot>x = x` have "?e (Y n)\<cdot>x = x" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
282 |
by (simp add: eventual_iterate_oo_fixed_iff [OF d below]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
283 |
moreover have "?e (Y n)\<cdot>x \<sqsubseteq> (\<Squnion>i. ?e (Y i)\<cdot>x)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
284 |
by (rule is_ub_thelub, simp add: eY) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
285 |
ultimately have "x \<sqsubseteq> (\<Squnion>i. ?e (Y i))\<cdot>x" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
286 |
by (simp add: contlub_cfun_fun eY) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
287 |
also have "(\<Squnion>i. ?e (Y i))\<cdot>x \<sqsubseteq> x" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
288 |
apply (rule deflation.below) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
289 |
apply (rule admD [OF adm_deflation eY]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
290 |
apply (rule pre_deflation.deflation_d) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
291 |
apply (rule pre_deflation_oo [OF d below]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
292 |
done |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
293 |
finally show "(\<Squnion>i. ?e (Y i))\<cdot>x = x" .. |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
294 |
qed |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
295 |
qed |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
296 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
297 |
|
27409 | 298 |
subsection {* Type constructor for finite deflations *} |
299 |
||
300 |
defaultsort profinite |
|
301 |
||
302 |
typedef (open) 'a fin_defl = "{d::'a \<rightarrow> 'a. finite_deflation d}" |
|
303 |
by (fast intro: finite_deflation_approx) |
|
304 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
305 |
instantiation fin_defl :: (profinite) below |
27409 | 306 |
begin |
307 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
308 |
definition below_fin_defl_def: |
27409 | 309 |
"op \<sqsubseteq> \<equiv> \<lambda>x y. Rep_fin_defl x \<sqsubseteq> Rep_fin_defl y" |
310 |
||
311 |
instance .. |
|
312 |
end |
|
313 |
||
314 |
instance fin_defl :: (profinite) po |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
315 |
by (rule typedef_po [OF type_definition_fin_defl below_fin_defl_def]) |
27409 | 316 |
|
317 |
lemma finite_deflation_Rep_fin_defl: "finite_deflation (Rep_fin_defl d)" |
|
318 |
using Rep_fin_defl by simp |
|
319 |
||
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
320 |
lemma deflation_Rep_fin_defl: "deflation (Rep_fin_defl d)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
321 |
using finite_deflation_Rep_fin_defl |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
322 |
by (rule finite_deflation_imp_deflation) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
323 |
|
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29252
diff
changeset
|
324 |
interpretation Rep_fin_defl: finite_deflation "Rep_fin_defl d" |
27409 | 325 |
by (rule finite_deflation_Rep_fin_defl) |
326 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
327 |
lemma fin_defl_belowI: |
27409 | 328 |
"(\<And>x. Rep_fin_defl a\<cdot>x = x \<Longrightarrow> Rep_fin_defl b\<cdot>x = x) \<Longrightarrow> a \<sqsubseteq> b" |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
329 |
unfolding below_fin_defl_def |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
330 |
by (rule Rep_fin_defl.belowI) |
27409 | 331 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
332 |
lemma fin_defl_belowD: |
27409 | 333 |
"\<lbrakk>a \<sqsubseteq> b; Rep_fin_defl a\<cdot>x = x\<rbrakk> \<Longrightarrow> Rep_fin_defl b\<cdot>x = x" |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
334 |
unfolding below_fin_defl_def |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
335 |
by (rule Rep_fin_defl.belowD) |
27409 | 336 |
|
337 |
lemma fin_defl_eqI: |
|
338 |
"(\<And>x. Rep_fin_defl a\<cdot>x = x \<longleftrightarrow> Rep_fin_defl b\<cdot>x = x) \<Longrightarrow> a = b" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
339 |
apply (rule below_antisym) |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
340 |
apply (rule fin_defl_belowI, simp) |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
341 |
apply (rule fin_defl_belowI, simp) |
27409 | 342 |
done |
343 |
||
344 |
lemma Abs_fin_defl_mono: |
|
345 |
"\<lbrakk>finite_deflation a; finite_deflation b; a \<sqsubseteq> b\<rbrakk> |
|
346 |
\<Longrightarrow> Abs_fin_defl a \<sqsubseteq> Abs_fin_defl b" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
347 |
unfolding below_fin_defl_def |
27409 | 348 |
by (simp add: Abs_fin_defl_inverse) |
349 |
||
350 |
||
351 |
subsection {* Take function for finite deflations *} |
|
352 |
||
353 |
definition |
|
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
354 |
defl_approx :: "nat \<Rightarrow> ('a \<rightarrow> 'a) \<Rightarrow> ('a \<rightarrow> 'a)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
355 |
where |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
356 |
"defl_approx i d = eventual_iterate (approx i oo d)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
357 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
358 |
lemma finite_deflation_defl_approx: |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
359 |
"deflation d \<Longrightarrow> finite_deflation (defl_approx i d)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
360 |
unfolding defl_approx_def |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
361 |
apply (rule pre_deflation.finite_deflation_d) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
362 |
apply (rule pre_deflation_oo) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
363 |
apply (rule finite_deflation_approx) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
364 |
apply (erule deflation.below) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
365 |
done |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
366 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
367 |
lemma deflation_defl_approx: |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
368 |
"deflation d \<Longrightarrow> deflation (defl_approx i d)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
369 |
apply (rule finite_deflation_imp_deflation) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
370 |
apply (erule finite_deflation_defl_approx) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
371 |
done |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
372 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
373 |
lemma defl_approx_fixed_iff: |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
374 |
"deflation d \<Longrightarrow> defl_approx i d\<cdot>x = x \<longleftrightarrow> approx i\<cdot>x = x \<and> d\<cdot>x = x" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
375 |
unfolding defl_approx_def |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
376 |
apply (rule eventual_iterate_oo_fixed_iff) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
377 |
apply (rule finite_deflation_approx) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
378 |
apply (erule deflation.below) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
379 |
done |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
380 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
381 |
lemma defl_approx_below: |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
382 |
"\<lbrakk>a \<sqsubseteq> b; deflation a; deflation b\<rbrakk> \<Longrightarrow> defl_approx i a \<sqsubseteq> defl_approx i b" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
383 |
apply (rule deflation.belowI) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
384 |
apply (erule deflation_defl_approx) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
385 |
apply (simp add: defl_approx_fixed_iff) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
386 |
apply (erule (1) deflation.belowD) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
387 |
apply (erule conjunct2) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
388 |
done |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
389 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
390 |
lemma cont2cont_defl_approx: |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
391 |
assumes cont: "cont f" and below: "\<And>x y. f x\<cdot>y \<sqsubseteq> y" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
392 |
shows "cont (\<lambda>x. defl_approx i (f x))" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
393 |
unfolding defl_approx_def |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
394 |
using finite_deflation_approx assms |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
395 |
by (rule cont2cont_eventual_iterate_oo) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
396 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
397 |
definition |
27409 | 398 |
fd_take :: "nat \<Rightarrow> 'a fin_defl \<Rightarrow> 'a fin_defl" |
399 |
where |
|
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
400 |
"fd_take i d = Abs_fin_defl (defl_approx i (Rep_fin_defl d))" |
27409 | 401 |
|
402 |
lemma Rep_fin_defl_fd_take: |
|
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
403 |
"Rep_fin_defl (fd_take i d) = defl_approx i (Rep_fin_defl d)" |
27409 | 404 |
unfolding fd_take_def |
405 |
apply (rule Abs_fin_defl_inverse [unfolded mem_Collect_eq]) |
|
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
406 |
apply (rule finite_deflation_defl_approx) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
407 |
apply (rule deflation_Rep_fin_defl) |
27409 | 408 |
done |
409 |
||
410 |
lemma fd_take_fixed_iff: |
|
411 |
"Rep_fin_defl (fd_take i d)\<cdot>x = x \<longleftrightarrow> |
|
412 |
approx i\<cdot>x = x \<and> Rep_fin_defl d\<cdot>x = x" |
|
413 |
unfolding Rep_fin_defl_fd_take |
|
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
414 |
apply (rule defl_approx_fixed_iff) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
415 |
apply (rule deflation_Rep_fin_defl) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
416 |
done |
27409 | 417 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
418 |
lemma fd_take_below: "fd_take n d \<sqsubseteq> d" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
419 |
apply (rule fin_defl_belowI) |
27409 | 420 |
apply (simp add: fd_take_fixed_iff) |
421 |
done |
|
422 |
||
423 |
lemma fd_take_idem: "fd_take n (fd_take n d) = fd_take n d" |
|
424 |
apply (rule fin_defl_eqI) |
|
425 |
apply (simp add: fd_take_fixed_iff) |
|
426 |
done |
|
427 |
||
428 |
lemma fd_take_mono: "a \<sqsubseteq> b \<Longrightarrow> fd_take n a \<sqsubseteq> fd_take n b" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
429 |
apply (rule fin_defl_belowI) |
27409 | 430 |
apply (simp add: fd_take_fixed_iff) |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
431 |
apply (simp add: fin_defl_belowD) |
27409 | 432 |
done |
433 |
||
434 |
lemma approx_fixed_le_lemma: "\<lbrakk>i \<le> j; approx i\<cdot>x = x\<rbrakk> \<Longrightarrow> approx j\<cdot>x = x" |
|
435 |
by (erule subst, simp add: min_def) |
|
436 |
||
437 |
lemma fd_take_chain: "m \<le> n \<Longrightarrow> fd_take m a \<sqsubseteq> fd_take n a" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
438 |
apply (rule fin_defl_belowI) |
27409 | 439 |
apply (simp add: fd_take_fixed_iff) |
440 |
apply (simp add: approx_fixed_le_lemma) |
|
441 |
done |
|
442 |
||
443 |
lemma finite_range_fd_take: "finite (range (fd_take n))" |
|
444 |
apply (rule finite_imageD [where f="\<lambda>a. {x. Rep_fin_defl a\<cdot>x = x}"]) |
|
445 |
apply (rule finite_subset [where B="Pow {x. approx n\<cdot>x = x}"]) |
|
446 |
apply (clarify, simp add: fd_take_fixed_iff) |
|
447 |
apply (simp add: finite_fixes_approx) |
|
448 |
apply (rule inj_onI, clarify) |
|
449 |
apply (simp add: expand_set_eq fin_defl_eqI) |
|
450 |
done |
|
451 |
||
452 |
lemma fd_take_covers: "\<exists>n. fd_take n a = a" |
|
453 |
apply (rule_tac x= |
|
454 |
"Max ((\<lambda>x. LEAST n. approx n\<cdot>x = x) ` {x. Rep_fin_defl a\<cdot>x = x})" in exI) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
455 |
apply (rule below_antisym) |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
456 |
apply (rule fd_take_below) |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
457 |
apply (rule fin_defl_belowI) |
27409 | 458 |
apply (simp add: fd_take_fixed_iff) |
459 |
apply (rule approx_fixed_le_lemma) |
|
460 |
apply (rule Max_ge) |
|
461 |
apply (rule finite_imageI) |
|
462 |
apply (rule Rep_fin_defl.finite_fixes) |
|
463 |
apply (rule imageI) |
|
464 |
apply (erule CollectI) |
|
465 |
apply (rule LeastI_ex) |
|
466 |
apply (rule profinite_compact_eq_approx) |
|
467 |
apply (erule subst) |
|
468 |
apply (rule Rep_fin_defl.compact) |
|
469 |
done |
|
470 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
471 |
interpretation fin_defl: basis_take below fd_take |
27409 | 472 |
apply default |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
473 |
apply (rule fd_take_below) |
27409 | 474 |
apply (rule fd_take_idem) |
475 |
apply (erule fd_take_mono) |
|
476 |
apply (rule fd_take_chain, simp) |
|
477 |
apply (rule finite_range_fd_take) |
|
478 |
apply (rule fd_take_covers) |
|
479 |
done |
|
480 |
||
33586
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
481 |
|
27409 | 482 |
subsection {* Defining algebraic deflations by ideal completion *} |
483 |
||
484 |
typedef (open) 'a alg_defl = |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
485 |
"{S::'a fin_defl set. below.ideal S}" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
486 |
by (fast intro: below.ideal_principal) |
27409 | 487 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
488 |
instantiation alg_defl :: (profinite) below |
27409 | 489 |
begin |
490 |
||
491 |
definition |
|
492 |
"x \<sqsubseteq> y \<longleftrightarrow> Rep_alg_defl x \<subseteq> Rep_alg_defl y" |
|
493 |
||
494 |
instance .. |
|
495 |
end |
|
496 |
||
497 |
instance alg_defl :: (profinite) po |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
498 |
by (rule below.typedef_ideal_po |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
499 |
[OF type_definition_alg_defl below_alg_defl_def]) |
27409 | 500 |
|
501 |
instance alg_defl :: (profinite) cpo |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
502 |
by (rule below.typedef_ideal_cpo |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
503 |
[OF type_definition_alg_defl below_alg_defl_def]) |
27409 | 504 |
|
505 |
lemma Rep_alg_defl_lub: |
|
506 |
"chain Y \<Longrightarrow> Rep_alg_defl (\<Squnion>i. Y i) = (\<Union>i. Rep_alg_defl (Y i))" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
507 |
by (rule below.typedef_ideal_rep_contlub |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
508 |
[OF type_definition_alg_defl below_alg_defl_def]) |
27409 | 509 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
510 |
lemma ideal_Rep_alg_defl: "below.ideal (Rep_alg_defl xs)" |
27409 | 511 |
by (rule Rep_alg_defl [unfolded mem_Collect_eq]) |
512 |
||
513 |
definition |
|
514 |
alg_defl_principal :: "'a fin_defl \<Rightarrow> 'a alg_defl" where |
|
515 |
"alg_defl_principal t = Abs_alg_defl {u. u \<sqsubseteq> t}" |
|
516 |
||
517 |
lemma Rep_alg_defl_principal: |
|
518 |
"Rep_alg_defl (alg_defl_principal t) = {u. u \<sqsubseteq> t}" |
|
519 |
unfolding alg_defl_principal_def |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
520 |
by (simp add: Abs_alg_defl_inverse below.ideal_principal) |
27409 | 521 |
|
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29252
diff
changeset
|
522 |
interpretation alg_defl: |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
523 |
ideal_completion below fd_take alg_defl_principal Rep_alg_defl |
27409 | 524 |
apply default |
525 |
apply (rule ideal_Rep_alg_defl) |
|
526 |
apply (erule Rep_alg_defl_lub) |
|
527 |
apply (rule Rep_alg_defl_principal) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
528 |
apply (simp only: below_alg_defl_def) |
27409 | 529 |
done |
530 |
||
531 |
text {* Algebraic deflations are pointed *} |
|
532 |
||
533 |
lemma finite_deflation_UU: "finite_deflation \<bottom>" |
|
534 |
by default simp_all |
|
535 |
||
536 |
lemma alg_defl_minimal: |
|
537 |
"alg_defl_principal (Abs_fin_defl \<bottom>) \<sqsubseteq> x" |
|
538 |
apply (induct x rule: alg_defl.principal_induct, simp) |
|
539 |
apply (rule alg_defl.principal_mono) |
|
540 |
apply (induct_tac a) |
|
541 |
apply (rule Abs_fin_defl_mono) |
|
542 |
apply (rule finite_deflation_UU) |
|
543 |
apply simp |
|
544 |
apply (rule minimal) |
|
545 |
done |
|
546 |
||
547 |
instance alg_defl :: (bifinite) pcpo |
|
548 |
by intro_classes (fast intro: alg_defl_minimal) |
|
549 |
||
550 |
lemma inst_alg_defl_pcpo: "\<bottom> = alg_defl_principal (Abs_fin_defl \<bottom>)" |
|
551 |
by (rule alg_defl_minimal [THEN UU_I, symmetric]) |
|
552 |
||
553 |
text {* Algebraic deflations are profinite *} |
|
554 |
||
555 |
instantiation alg_defl :: (profinite) profinite |
|
556 |
begin |
|
557 |
||
558 |
definition |
|
559 |
approx_alg_defl_def: "approx = alg_defl.completion_approx" |
|
560 |
||
561 |
instance |
|
562 |
apply (intro_classes, unfold approx_alg_defl_def) |
|
563 |
apply (rule alg_defl.chain_completion_approx) |
|
564 |
apply (rule alg_defl.lub_completion_approx) |
|
565 |
apply (rule alg_defl.completion_approx_idem) |
|
566 |
apply (rule alg_defl.finite_fixes_completion_approx) |
|
567 |
done |
|
568 |
||
569 |
end |
|
570 |
||
571 |
instance alg_defl :: (bifinite) bifinite .. |
|
572 |
||
573 |
lemma approx_alg_defl_principal [simp]: |
|
574 |
"approx n\<cdot>(alg_defl_principal t) = alg_defl_principal (fd_take n t)" |
|
575 |
unfolding approx_alg_defl_def |
|
576 |
by (rule alg_defl.completion_approx_principal) |
|
577 |
||
578 |
lemma approx_eq_alg_defl_principal: |
|
579 |
"\<exists>t\<in>Rep_alg_defl xs. approx n\<cdot>xs = alg_defl_principal (fd_take n t)" |
|
580 |
unfolding approx_alg_defl_def |
|
581 |
by (rule alg_defl.completion_approx_eq_principal) |
|
582 |
||
583 |
||
584 |
subsection {* Applying algebraic deflations *} |
|
585 |
||
586 |
definition |
|
587 |
cast :: "'a alg_defl \<rightarrow> 'a \<rightarrow> 'a" |
|
588 |
where |
|
589 |
"cast = alg_defl.basis_fun Rep_fin_defl" |
|
590 |
||
591 |
lemma cast_alg_defl_principal: |
|
592 |
"cast\<cdot>(alg_defl_principal a) = Rep_fin_defl a" |
|
593 |
unfolding cast_def |
|
594 |
apply (rule alg_defl.basis_fun_principal) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
595 |
apply (simp only: below_fin_defl_def) |
27409 | 596 |
done |
597 |
||
598 |
lemma deflation_cast: "deflation (cast\<cdot>d)" |
|
599 |
apply (induct d rule: alg_defl.principal_induct) |
|
600 |
apply (rule adm_subst [OF _ adm_deflation], simp) |
|
601 |
apply (simp add: cast_alg_defl_principal) |
|
602 |
apply (rule finite_deflation_imp_deflation) |
|
603 |
apply (rule finite_deflation_Rep_fin_defl) |
|
604 |
done |
|
605 |
||
606 |
lemma finite_deflation_cast: |
|
607 |
"compact d \<Longrightarrow> finite_deflation (cast\<cdot>d)" |
|
608 |
apply (drule alg_defl.compact_imp_principal, clarify) |
|
609 |
apply (simp add: cast_alg_defl_principal) |
|
610 |
apply (rule finite_deflation_Rep_fin_defl) |
|
611 |
done |
|
612 |
||
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29252
diff
changeset
|
613 |
interpretation cast: deflation "cast\<cdot>d" |
27409 | 614 |
by (rule deflation_cast) |
615 |
||
33586
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
616 |
declare cast.idem [simp] |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
617 |
|
31164
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
618 |
lemma cast_approx: "cast\<cdot>(approx n\<cdot>A) = defl_approx n (cast\<cdot>A)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
619 |
apply (rule alg_defl.principal_induct) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
620 |
apply (rule adm_eq) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
621 |
apply simp |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
622 |
apply (simp add: cont2cont_defl_approx cast.below) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
623 |
apply (simp only: approx_alg_defl_principal) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
624 |
apply (simp only: cast_alg_defl_principal) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
625 |
apply (simp only: Rep_fin_defl_fd_take) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
626 |
done |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
627 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
628 |
lemma cast_approx_fixed_iff: |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
629 |
"cast\<cdot>(approx i\<cdot>A)\<cdot>x = x \<longleftrightarrow> approx i\<cdot>x = x \<and> cast\<cdot>A\<cdot>x = x" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
630 |
apply (simp only: cast_approx) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
631 |
apply (rule defl_approx_fixed_iff) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
632 |
apply (rule deflation_cast) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
633 |
done |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
634 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
635 |
lemma defl_approx_cast: "defl_approx i (cast\<cdot>A) = cast\<cdot>(approx i\<cdot>A)" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
636 |
by (rule cast_approx [symmetric]) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
637 |
|
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
638 |
lemma cast_below_imp_below: "cast\<cdot>A \<sqsubseteq> cast\<cdot>B \<Longrightarrow> A \<sqsubseteq> B" |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
639 |
apply (rule profinite_below_ext) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
640 |
apply (drule_tac i=i in defl_approx_below) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
641 |
apply (rule deflation_cast) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
642 |
apply (rule deflation_cast) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
643 |
apply (simp only: defl_approx_cast) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
644 |
apply (cut_tac x="approx i\<cdot>A" in alg_defl.compact_imp_principal) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
645 |
apply (rule compact_approx) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
646 |
apply (cut_tac x="approx i\<cdot>B" in alg_defl.compact_imp_principal) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
647 |
apply (rule compact_approx) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
648 |
apply clarsimp |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
649 |
apply (simp add: cast_alg_defl_principal) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
650 |
apply (simp add: below_fin_defl_def) |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
651 |
done |
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents:
31076
diff
changeset
|
652 |
|
33586
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
653 |
lemma cast_eq_imp_eq: "cast\<cdot>A = cast\<cdot>B \<Longrightarrow> A = B" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
654 |
by (simp add: below_antisym cast_below_imp_below) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
655 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
656 |
lemma cast_strict1 [simp]: "cast\<cdot>\<bottom> = \<bottom>" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
657 |
apply (subst inst_alg_defl_pcpo) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
658 |
apply (subst cast_alg_defl_principal) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
659 |
apply (rule Abs_fin_defl_inverse) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
660 |
apply (simp add: finite_deflation_UU) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
661 |
done |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
662 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
663 |
lemma cast_strict2 [simp]: "cast\<cdot>A\<cdot>\<bottom> = \<bottom>" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
664 |
by (rule cast.below [THEN UU_I]) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
665 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
666 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
667 |
subsection {* Deflation membership relation *} |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
668 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
669 |
definition |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
670 |
in_deflation :: "'a::profinite \<Rightarrow> 'a alg_defl \<Rightarrow> bool" (infixl ":::" 50) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
671 |
where |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
672 |
"x ::: A \<longleftrightarrow> cast\<cdot>A\<cdot>x = x" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
673 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
674 |
lemma adm_in_deflation: "adm (\<lambda>x. x ::: A)" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
675 |
unfolding in_deflation_def by simp |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
676 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
677 |
lemma in_deflationI: "cast\<cdot>A\<cdot>x = x \<Longrightarrow> x ::: A" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
678 |
unfolding in_deflation_def . |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
679 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
680 |
lemma cast_fixed: "x ::: A \<Longrightarrow> cast\<cdot>A\<cdot>x = x" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
681 |
unfolding in_deflation_def . |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
682 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
683 |
lemma cast_in_deflation [simp]: "cast\<cdot>A\<cdot>x ::: A" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
684 |
unfolding in_deflation_def by (rule cast.idem) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
685 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
686 |
lemma bottom_in_deflation [simp]: "\<bottom> ::: A" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
687 |
unfolding in_deflation_def by (rule cast_strict2) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
688 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
689 |
lemma subdeflationD: "A \<sqsubseteq> B \<Longrightarrow> x ::: A \<Longrightarrow> x ::: B" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
690 |
unfolding in_deflation_def |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
691 |
apply (rule deflation.belowD) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
692 |
apply (rule deflation_cast) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
693 |
apply (erule monofun_cfun_arg) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
694 |
apply assumption |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
695 |
done |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
696 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
697 |
lemma rev_subdeflationD: "x ::: A \<Longrightarrow> A \<sqsubseteq> B \<Longrightarrow> x ::: B" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
698 |
by (rule subdeflationD) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
699 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
700 |
lemma subdeflationI: "(\<And>x. x ::: A \<Longrightarrow> x ::: B) \<Longrightarrow> A \<sqsubseteq> B" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
701 |
apply (rule cast_below_imp_below) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
702 |
apply (rule cast.belowI) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
703 |
apply (simp add: in_deflation_def) |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
704 |
done |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
705 |
|
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
706 |
text "Identity deflation:" |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
707 |
|
27409 | 708 |
lemma "cast\<cdot>(\<Squnion>i. alg_defl_principal (Abs_fin_defl (approx i)))\<cdot>x = x" |
709 |
apply (subst contlub_cfun_arg) |
|
710 |
apply (rule chainI) |
|
711 |
apply (rule alg_defl.principal_mono) |
|
712 |
apply (rule Abs_fin_defl_mono) |
|
713 |
apply (rule finite_deflation_approx) |
|
714 |
apply (rule finite_deflation_approx) |
|
715 |
apply (rule chainE) |
|
716 |
apply (rule chain_approx) |
|
717 |
apply (simp add: cast_alg_defl_principal Abs_fin_defl_inverse finite_deflation_approx) |
|
718 |
done |
|
719 |
||
33586
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
720 |
subsection {* Bifinite domains and algebraic deflations *} |
0e745228d605
add in_deflation relation, more lemmas about cast
huffman
parents:
31164
diff
changeset
|
721 |
|
27409 | 722 |
text {* This lemma says that if we have an ep-pair from |
723 |
a bifinite domain into a universal domain, then e oo p |
|
724 |
is an algebraic deflation. *} |
|
725 |
||
726 |
lemma |
|
28611 | 727 |
assumes "ep_pair e p" |
27409 | 728 |
constrains e :: "'a::profinite \<rightarrow> 'b::profinite" |
729 |
shows "\<exists>d. cast\<cdot>d = e oo p" |
|
730 |
proof |
|
29237 | 731 |
interpret ep_pair e p by fact |
27409 | 732 |
let ?a = "\<lambda>i. e oo approx i oo p" |
733 |
have a: "\<And>i. finite_deflation (?a i)" |
|
734 |
apply (rule finite_deflation_e_d_p) |
|
735 |
apply (rule finite_deflation_approx) |
|
736 |
done |
|
737 |
let ?d = "\<Squnion>i. alg_defl_principal (Abs_fin_defl (?a i))" |
|
738 |
show "cast\<cdot>?d = e oo p" |
|
739 |
apply (subst contlub_cfun_arg) |
|
740 |
apply (rule chainI) |
|
741 |
apply (rule alg_defl.principal_mono) |
|
742 |
apply (rule Abs_fin_defl_mono [OF a a]) |
|
743 |
apply (rule chainE, simp) |
|
744 |
apply (subst cast_alg_defl_principal) |
|
745 |
apply (simp add: Abs_fin_defl_inverse a) |
|
746 |
apply (simp add: expand_cfun_eq lub_distribs) |
|
747 |
done |
|
748 |
qed |
|
749 |
||
750 |
text {* This lemma says that if we have an ep-pair |
|
751 |
from a cpo into a bifinite domain, and e oo p is |
|
752 |
an algebraic deflation, then the cpo is bifinite. *} |
|
753 |
||
754 |
lemma |
|
28611 | 755 |
assumes "ep_pair e p" |
27409 | 756 |
constrains e :: "'a::cpo \<rightarrow> 'b::profinite" |
757 |
assumes d: "\<And>x. cast\<cdot>d\<cdot>x = e\<cdot>(p\<cdot>x)" |
|
758 |
obtains a :: "nat \<Rightarrow> 'a \<rightarrow> 'a" where |
|
759 |
"\<And>i. finite_deflation (a i)" |
|
760 |
"(\<Squnion>i. a i) = ID" |
|
761 |
proof |
|
29237 | 762 |
interpret ep_pair e p by fact |
27409 | 763 |
let ?a = "\<lambda>i. p oo cast\<cdot>(approx i\<cdot>d) oo e" |
764 |
show "\<And>i. finite_deflation (?a i)" |
|
765 |
apply (rule finite_deflation_p_d_e) |
|
766 |
apply (rule finite_deflation_cast) |
|
767 |
apply (rule compact_approx) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
768 |
apply (rule below_eq_trans [OF _ d]) |
27409 | 769 |
apply (rule monofun_cfun_fun) |
770 |
apply (rule monofun_cfun_arg) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
771 |
apply (rule approx_below) |
27409 | 772 |
done |
773 |
show "(\<Squnion>i. ?a i) = ID" |
|
774 |
apply (rule ext_cfun, simp) |
|
775 |
apply (simp add: lub_distribs) |
|
776 |
apply (simp add: d) |
|
777 |
done |
|
778 |
qed |
|
779 |
||
780 |
end |