author | wenzelm |
Thu, 12 Dec 2024 15:45:29 +0100 | |
changeset 81583 | b6df83045178 |
parent 81579 | cf4bebd770b5 |
child 81584 | a065d8bcfd3d |
permissions | -rw-r--r-- |
42151 | 1 |
(* Title: HOL/HOLCF/Universal.thy |
27411 | 2 |
Author: Brian Huffman |
3 |
*) |
|
4 |
||
62175 | 5 |
section \<open>A universal bifinite domain\<close> |
35794 | 6 |
|
27411 | 7 |
theory Universal |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
65552
diff
changeset
|
8 |
imports Bifinite Completion "HOL-Library.Nat_Bijection" |
27411 | 9 |
begin |
10 |
||
81579 | 11 |
unbundle no binomial_syntax |
65552
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
wenzelm
parents:
64267
diff
changeset
|
12 |
|
81577 | 13 |
|
62175 | 14 |
subsection \<open>Basis for universal domain\<close> |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
15 |
|
62175 | 16 |
subsubsection \<open>Basis datatype\<close> |
27411 | 17 |
|
41295 | 18 |
type_synonym ubasis = nat |
27411 | 19 |
|
20 |
definition |
|
21 |
node :: "nat \<Rightarrow> ubasis \<Rightarrow> ubasis set \<Rightarrow> ubasis" |
|
22 |
where |
|
35701 | 23 |
"node i a S = Suc (prod_encode (i, prod_encode (a, set_encode S)))" |
27411 | 24 |
|
30505 | 25 |
lemma node_not_0 [simp]: "node i a S \<noteq> 0" |
27411 | 26 |
unfolding node_def by simp |
27 |
||
30505 | 28 |
lemma node_gt_0 [simp]: "0 < node i a S" |
27411 | 29 |
unfolding node_def by simp |
30 |
||
31 |
lemma node_inject [simp]: |
|
30505 | 32 |
"\<lbrakk>finite S; finite T\<rbrakk> |
33 |
\<Longrightarrow> node i a S = node j b T \<longleftrightarrow> i = j \<and> a = b \<and> S = T" |
|
35701 | 34 |
unfolding node_def by (simp add: prod_encode_eq set_encode_eq) |
27411 | 35 |
|
30505 | 36 |
lemma node_gt0: "i < node i a S" |
27411 | 37 |
unfolding node_def less_Suc_eq_le |
35701 | 38 |
by (rule le_prod_encode_1) |
27411 | 39 |
|
30505 | 40 |
lemma node_gt1: "a < node i a S" |
27411 | 41 |
unfolding node_def less_Suc_eq_le |
35701 | 42 |
by (rule order_trans [OF le_prod_encode_1 le_prod_encode_2]) |
27411 | 43 |
|
44 |
lemma nat_less_power2: "n < 2^n" |
|
73869 | 45 |
by (fact less_exp) |
27411 | 46 |
|
30505 | 47 |
lemma node_gt2: "\<lbrakk>finite S; b \<in> S\<rbrakk> \<Longrightarrow> b < node i a S" |
35701 | 48 |
unfolding node_def less_Suc_eq_le set_encode_def |
49 |
apply (rule order_trans [OF _ le_prod_encode_2]) |
|
50 |
apply (rule order_trans [OF _ le_prod_encode_2]) |
|
67399 | 51 |
apply (rule order_trans [where y="sum ((^) 2) {b}"]) |
27411 | 52 |
apply (simp add: nat_less_power2 [THEN order_less_imp_le]) |
64267 | 53 |
apply (erule sum_mono2, simp, simp) |
27411 | 54 |
done |
55 |
||
35701 | 56 |
lemma eq_prod_encode_pairI: |
57 |
"\<lbrakk>fst (prod_decode x) = a; snd (prod_decode x) = b\<rbrakk> \<Longrightarrow> x = prod_encode (a, b)" |
|
80175
200107cdd3ac
Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents:
73869
diff
changeset
|
58 |
by auto |
27411 | 59 |
|
60 |
lemma node_cases: |
|
61 |
assumes 1: "x = 0 \<Longrightarrow> P" |
|
30505 | 62 |
assumes 2: "\<And>i a S. \<lbrakk>finite S; x = node i a S\<rbrakk> \<Longrightarrow> P" |
27411 | 63 |
shows "P" |
64 |
apply (cases x) |
|
65 |
apply (erule 1) |
|
66 |
apply (rule 2) |
|
35701 | 67 |
apply (rule finite_set_decode) |
27411 | 68 |
apply (simp add: node_def) |
35701 | 69 |
apply (rule eq_prod_encode_pairI [OF refl]) |
70 |
apply (rule eq_prod_encode_pairI [OF refl refl]) |
|
27411 | 71 |
done |
72 |
||
73 |
lemma node_induct: |
|
74 |
assumes 1: "P 0" |
|
30505 | 75 |
assumes 2: "\<And>i a S. \<lbrakk>P a; finite S; \<forall>b\<in>S. P b\<rbrakk> \<Longrightarrow> P (node i a S)" |
27411 | 76 |
shows "P x" |
77 |
apply (induct x rule: nat_less_induct) |
|
78 |
apply (case_tac n rule: node_cases) |
|
79 |
apply (simp add: 1) |
|
80 |
apply (simp add: 2 node_gt1 node_gt2) |
|
81 |
done |
|
82 |
||
81577 | 83 |
|
62175 | 84 |
subsubsection \<open>Basis ordering\<close> |
27411 | 85 |
|
86 |
inductive |
|
87 |
ubasis_le :: "nat \<Rightarrow> nat \<Rightarrow> bool" |
|
88 |
where |
|
30505 | 89 |
ubasis_le_refl: "ubasis_le a a" |
27411 | 90 |
| ubasis_le_trans: |
30505 | 91 |
"\<lbrakk>ubasis_le a b; ubasis_le b c\<rbrakk> \<Longrightarrow> ubasis_le a c" |
27411 | 92 |
| ubasis_le_lower: |
30505 | 93 |
"finite S \<Longrightarrow> ubasis_le a (node i a S)" |
27411 | 94 |
| ubasis_le_upper: |
30505 | 95 |
"\<lbrakk>finite S; b \<in> S; ubasis_le a b\<rbrakk> \<Longrightarrow> ubasis_le (node i a S) b" |
27411 | 96 |
|
97 |
lemma ubasis_le_minimal: "ubasis_le 0 x" |
|
98 |
apply (induct x rule: node_induct) |
|
99 |
apply (rule ubasis_le_refl) |
|
100 |
apply (erule ubasis_le_trans) |
|
101 |
apply (erule ubasis_le_lower) |
|
102 |
done |
|
103 |
||
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
104 |
interpretation udom: preorder ubasis_le |
61169 | 105 |
apply standard |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
106 |
apply (rule ubasis_le_refl) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
107 |
apply (erule (1) ubasis_le_trans) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
108 |
done |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
109 |
|
81577 | 110 |
|
62175 | 111 |
subsubsection \<open>Generic take function\<close> |
27411 | 112 |
|
113 |
function |
|
114 |
ubasis_until :: "(ubasis \<Rightarrow> bool) \<Rightarrow> ubasis \<Rightarrow> ubasis" |
|
115 |
where |
|
116 |
"ubasis_until P 0 = 0" |
|
30505 | 117 |
| "finite S \<Longrightarrow> ubasis_until P (node i a S) = |
118 |
(if P (node i a S) then node i a S else ubasis_until P a)" |
|
56073
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
54863
diff
changeset
|
119 |
apply clarify |
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
54863
diff
changeset
|
120 |
apply (rule_tac x=b in node_cases) |
68780 | 121 |
apply simp_all |
27411 | 122 |
done |
123 |
||
124 |
termination ubasis_until |
|
125 |
apply (relation "measure snd") |
|
126 |
apply (rule wf_measure) |
|
127 |
apply (simp add: node_gt1) |
|
128 |
done |
|
129 |
||
130 |
lemma ubasis_until: "P 0 \<Longrightarrow> P (ubasis_until P x)" |
|
131 |
by (induct x rule: node_induct) simp_all |
|
132 |
||
133 |
lemma ubasis_until': "0 < ubasis_until P x \<Longrightarrow> P (ubasis_until P x)" |
|
134 |
by (induct x rule: node_induct) auto |
|
135 |
||
136 |
lemma ubasis_until_same: "P x \<Longrightarrow> ubasis_until P x = x" |
|
137 |
by (induct x rule: node_induct) simp_all |
|
138 |
||
139 |
lemma ubasis_until_idem: |
|
140 |
"P 0 \<Longrightarrow> ubasis_until P (ubasis_until P x) = ubasis_until P x" |
|
141 |
by (rule ubasis_until_same [OF ubasis_until]) |
|
142 |
||
143 |
lemma ubasis_until_0: |
|
144 |
"\<forall>x. x \<noteq> 0 \<longrightarrow> \<not> P x \<Longrightarrow> ubasis_until P x = 0" |
|
145 |
by (induct x rule: node_induct) simp_all |
|
146 |
||
147 |
lemma ubasis_until_less: "ubasis_le (ubasis_until P x) x" |
|
148 |
apply (induct x rule: node_induct) |
|
149 |
apply (simp add: ubasis_le_refl) |
|
80175
200107cdd3ac
Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents:
73869
diff
changeset
|
150 |
by (metis ubasis_le.simps ubasis_until.simps(2)) |
27411 | 151 |
|
152 |
lemma ubasis_until_chain: |
|
153 |
assumes PQ: "\<And>x. P x \<Longrightarrow> Q x" |
|
154 |
shows "ubasis_le (ubasis_until P x) (ubasis_until Q x)" |
|
155 |
apply (induct x rule: node_induct) |
|
156 |
apply (simp add: ubasis_le_refl) |
|
80175
200107cdd3ac
Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents:
73869
diff
changeset
|
157 |
by (metis assms ubasis_until.simps(2) ubasis_until_less) |
27411 | 158 |
|
159 |
lemma ubasis_until_mono: |
|
30505 | 160 |
assumes "\<And>i a S b. \<lbrakk>finite S; P (node i a S); b \<in> S; ubasis_le a b\<rbrakk> \<Longrightarrow> P b" |
161 |
shows "ubasis_le a b \<Longrightarrow> ubasis_le (ubasis_until P a) (ubasis_until P b)" |
|
30561 | 162 |
proof (induct set: ubasis_le) |
163 |
case (ubasis_le_refl a) show ?case by (rule ubasis_le.ubasis_le_refl) |
|
164 |
next |
|
165 |
case (ubasis_le_trans a b c) thus ?case by - (rule ubasis_le.ubasis_le_trans) |
|
166 |
next |
|
167 |
case (ubasis_le_lower S a i) thus ?case |
|
80175
200107cdd3ac
Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents:
73869
diff
changeset
|
168 |
by (metis ubasis_le.simps ubasis_until.simps(2) ubasis_until_less) |
30561 | 169 |
next |
170 |
case (ubasis_le_upper S b a i) thus ?case |
|
80175
200107cdd3ac
Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents:
73869
diff
changeset
|
171 |
by (metis assms ubasis_le.simps ubasis_until.simps(2) ubasis_until_same) |
30561 | 172 |
qed |
27411 | 173 |
|
174 |
lemma finite_range_ubasis_until: |
|
175 |
"finite {x. P x} \<Longrightarrow> finite (range (ubasis_until P))" |
|
176 |
apply (rule finite_subset [where B="insert 0 {x. P x}"]) |
|
177 |
apply (clarsimp simp add: ubasis_until') |
|
178 |
apply simp |
|
179 |
done |
|
180 |
||
181 |
||
62175 | 182 |
subsection \<open>Defining the universal domain by ideal completion\<close> |
27411 | 183 |
|
49834 | 184 |
typedef udom = "{S. udom.ideal S}" |
40888
28cd51cff70c
reformulate lemma preorder.ex_ideal, and use it for typedefs
huffman
parents:
40774
diff
changeset
|
185 |
by (rule udom.ex_ideal) |
27411 | 186 |
|
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
|
187 |
instantiation udom :: below |
27411 | 188 |
begin |
189 |
||
190 |
definition |
|
191 |
"x \<sqsubseteq> y \<longleftrightarrow> Rep_udom x \<subseteq> Rep_udom y" |
|
192 |
||
193 |
instance .. |
|
194 |
end |
|
195 |
||
196 |
instance udom :: po |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
197 |
using type_definition_udom below_udom_def |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
198 |
by (rule udom.typedef_ideal_po) |
27411 | 199 |
|
200 |
instance udom :: cpo |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
201 |
using type_definition_udom below_udom_def |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
202 |
by (rule udom.typedef_ideal_cpo) |
27411 | 203 |
|
204 |
definition |
|
205 |
udom_principal :: "nat \<Rightarrow> udom" where |
|
206 |
"udom_principal t = Abs_udom {u. ubasis_le u t}" |
|
207 |
||
39984 | 208 |
lemma ubasis_countable: "\<exists>f::ubasis \<Rightarrow> nat. inj f" |
209 |
by (rule exI, rule inj_on_id) |
|
27411 | 210 |
|
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
30561
diff
changeset
|
211 |
interpretation udom: |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
212 |
ideal_completion ubasis_le udom_principal Rep_udom |
39984 | 213 |
using type_definition_udom below_udom_def |
214 |
using udom_principal_def ubasis_countable |
|
215 |
by (rule udom.typedef_ideal_completion) |
|
27411 | 216 |
|
62175 | 217 |
text \<open>Universal domain is pointed\<close> |
27411 | 218 |
|
219 |
lemma udom_minimal: "udom_principal 0 \<sqsubseteq> x" |
|
220 |
apply (induct x rule: udom.principal_induct) |
|
221 |
apply (simp, simp add: ubasis_le_minimal) |
|
222 |
done |
|
223 |
||
224 |
instance udom :: pcpo |
|
225 |
by intro_classes (fast intro: udom_minimal) |
|
226 |
||
227 |
lemma inst_udom_pcpo: "\<bottom> = udom_principal 0" |
|
41430
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
huffman
parents:
41413
diff
changeset
|
228 |
by (rule udom_minimal [THEN bottomI, symmetric]) |
27411 | 229 |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
230 |
|
62175 | 231 |
subsection \<open>Compact bases of domains\<close> |
27411 | 232 |
|
49834 | 233 |
typedef 'a compact_basis = "{x::'a::pcpo. compact x}" |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
234 |
by auto |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
235 |
|
41370
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
236 |
lemma Rep_compact_basis' [simp]: "compact (Rep_compact_basis a)" |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
237 |
by (rule Rep_compact_basis [unfolded mem_Collect_eq]) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
238 |
|
41370
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
239 |
lemma Abs_compact_basis_inverse' [simp]: |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
240 |
"compact x \<Longrightarrow> Rep_compact_basis (Abs_compact_basis x) = x" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
241 |
by (rule Abs_compact_basis_inverse [unfolded mem_Collect_eq]) |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
242 |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
243 |
instantiation compact_basis :: (pcpo) below |
27411 | 244 |
begin |
245 |
||
246 |
definition |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
247 |
compact_le_def: |
67399 | 248 |
"(\<sqsubseteq>) \<equiv> (\<lambda>x y. Rep_compact_basis x \<sqsubseteq> Rep_compact_basis y)" |
27411 | 249 |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
250 |
instance .. |
27411 | 251 |
end |
252 |
||
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
253 |
instance compact_basis :: (pcpo) po |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
254 |
using type_definition_compact_basis compact_le_def |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
255 |
by (rule typedef_po) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
256 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
257 |
definition |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81579
diff
changeset
|
258 |
approximants :: "'a::pcpo \<Rightarrow> 'a compact_basis set" where |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
259 |
"approximants = (\<lambda>x. {a. Rep_compact_basis a \<sqsubseteq> x})" |
27411 | 260 |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
261 |
definition |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
262 |
compact_bot :: "'a::pcpo compact_basis" where |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
263 |
"compact_bot = Abs_compact_basis \<bottom>" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
264 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
265 |
lemma Rep_compact_bot [simp]: "Rep_compact_basis compact_bot = \<bottom>" |
41370
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
266 |
unfolding compact_bot_def by simp |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
267 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
268 |
lemma compact_bot_minimal [simp]: "compact_bot \<sqsubseteq> a" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
269 |
unfolding compact_le_def Rep_compact_bot by simp |
27411 | 270 |
|
271 |
||
62175 | 272 |
subsection \<open>Universality of \emph{udom}\<close> |
27411 | 273 |
|
62175 | 274 |
text \<open>We use a locale to parameterize the construction over a chain |
275 |
of approx functions on the type to be embedded.\<close> |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
276 |
|
46868 | 277 |
locale bifinite_approx_chain = |
278 |
approx_chain approx for approx :: "nat \<Rightarrow> 'a::bifinite \<rightarrow> 'a" |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
279 |
begin |
27411 | 280 |
|
81577 | 281 |
|
62175 | 282 |
subsubsection \<open>Choosing a maximal element from a finite set\<close> |
27411 | 283 |
|
284 |
lemma finite_has_maximal: |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
285 |
fixes A :: "'a compact_basis set" |
27411 | 286 |
shows "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> \<exists>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y" |
287 |
proof (induct rule: finite_ne_induct) |
|
288 |
case (singleton x) |
|
289 |
show ?case by simp |
|
290 |
next |
|
291 |
case (insert a A) |
|
62175 | 292 |
from \<open>\<exists>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y\<close> |
27411 | 293 |
obtain x where x: "x \<in> A" |
294 |
and x_eq: "\<And>y. \<lbrakk>y \<in> A; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> x = y" by fast |
|
295 |
show ?case |
|
296 |
proof (intro bexI ballI impI) |
|
297 |
fix y |
|
298 |
assume "y \<in> insert a A" and "(if x \<sqsubseteq> a then a else x) \<sqsubseteq> y" |
|
299 |
thus "(if x \<sqsubseteq> a then a else x) = y" |
|
300 |
apply auto |
|
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
|
301 |
apply (frule (1) below_trans) |
27411 | 302 |
apply (frule (1) x_eq) |
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
|
303 |
apply (rule below_antisym, assumption) |
27411 | 304 |
apply simp |
305 |
apply (erule (1) x_eq) |
|
306 |
done |
|
307 |
next |
|
308 |
show "(if x \<sqsubseteq> a then a else x) \<in> insert a A" |
|
309 |
by (simp add: x) |
|
310 |
qed |
|
311 |
qed |
|
312 |
||
313 |
definition |
|
314 |
choose :: "'a compact_basis set \<Rightarrow> 'a compact_basis" |
|
315 |
where |
|
316 |
"choose A = (SOME x. x \<in> {x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y})" |
|
317 |
||
318 |
lemma choose_lemma: |
|
319 |
"\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> choose A \<in> {x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y}" |
|
320 |
unfolding choose_def |
|
321 |
apply (rule someI_ex) |
|
322 |
apply (frule (1) finite_has_maximal, fast) |
|
323 |
done |
|
324 |
||
325 |
lemma maximal_choose: |
|
326 |
"\<lbrakk>finite A; y \<in> A; choose A \<sqsubseteq> y\<rbrakk> \<Longrightarrow> choose A = y" |
|
327 |
apply (cases "A = {}", simp) |
|
328 |
apply (frule (1) choose_lemma, simp) |
|
329 |
done |
|
330 |
||
331 |
lemma choose_in: "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> choose A \<in> A" |
|
332 |
by (frule (1) choose_lemma, simp) |
|
333 |
||
334 |
function |
|
335 |
choose_pos :: "'a compact_basis set \<Rightarrow> 'a compact_basis \<Rightarrow> nat" |
|
336 |
where |
|
337 |
"choose_pos A x = |
|
338 |
(if finite A \<and> x \<in> A \<and> x \<noteq> choose A |
|
339 |
then Suc (choose_pos (A - {choose A}) x) else 0)" |
|
340 |
by auto |
|
341 |
||
342 |
termination choose_pos |
|
343 |
apply (relation "measure (card \<circ> fst)", simp) |
|
344 |
apply clarsimp |
|
345 |
apply (rule card_Diff1_less) |
|
346 |
apply assumption |
|
347 |
apply (erule choose_in) |
|
348 |
apply clarsimp |
|
349 |
done |
|
350 |
||
351 |
declare choose_pos.simps [simp del] |
|
352 |
||
353 |
lemma choose_pos_choose: "finite A \<Longrightarrow> choose_pos A (choose A) = 0" |
|
354 |
by (simp add: choose_pos.simps) |
|
355 |
||
356 |
lemma inj_on_choose_pos [OF refl]: |
|
357 |
"\<lbrakk>card A = n; finite A\<rbrakk> \<Longrightarrow> inj_on (choose_pos A) A" |
|
358 |
apply (induct n arbitrary: A) |
|
359 |
apply simp |
|
360 |
apply (case_tac "A = {}", simp) |
|
361 |
apply (frule (1) choose_in) |
|
362 |
apply (rule inj_onI) |
|
363 |
apply (drule_tac x="A - {choose A}" in meta_spec, simp) |
|
364 |
apply (simp add: choose_pos.simps) |
|
62390 | 365 |
apply (simp split: if_split_asm) |
27411 | 366 |
apply (erule (1) inj_onD, simp, simp) |
367 |
done |
|
368 |
||
369 |
lemma choose_pos_bounded [OF refl]: |
|
370 |
"\<lbrakk>card A = n; finite A; x \<in> A\<rbrakk> \<Longrightarrow> choose_pos A x < n" |
|
371 |
apply (induct n arbitrary: A) |
|
372 |
apply simp |
|
373 |
apply (case_tac "A = {}", simp) |
|
374 |
apply (frule (1) choose_in) |
|
375 |
apply (subst choose_pos.simps) |
|
376 |
apply simp |
|
377 |
done |
|
378 |
||
379 |
lemma choose_pos_lessD: |
|
41182 | 380 |
"\<lbrakk>choose_pos A x < choose_pos A y; finite A; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> x \<notsqsubseteq> y" |
27411 | 381 |
apply (induct A x arbitrary: y rule: choose_pos.induct) |
382 |
apply simp |
|
383 |
apply (case_tac "x = choose A") |
|
384 |
apply simp |
|
385 |
apply (rule notI) |
|
386 |
apply (frule (2) maximal_choose) |
|
387 |
apply simp |
|
388 |
apply (case_tac "y = choose A") |
|
389 |
apply (simp add: choose_pos_choose) |
|
390 |
apply (drule_tac x=y in meta_spec) |
|
391 |
apply simp |
|
392 |
apply (erule meta_mp) |
|
393 |
apply (simp add: choose_pos.simps) |
|
394 |
done |
|
395 |
||
81577 | 396 |
|
62175 | 397 |
subsubsection \<open>Compact basis take function\<close> |
27411 | 398 |
|
399 |
primrec |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
400 |
cb_take :: "nat \<Rightarrow> 'a compact_basis \<Rightarrow> 'a compact_basis" where |
27411 | 401 |
"cb_take 0 = (\<lambda>x. compact_bot)" |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
402 |
| "cb_take (Suc n) = (\<lambda>a. Abs_compact_basis (approx n\<cdot>(Rep_compact_basis a)))" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
403 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
404 |
declare cb_take.simps [simp del] |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
405 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
406 |
lemma cb_take_zero [simp]: "cb_take 0 a = compact_bot" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
407 |
by (simp only: cb_take.simps) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
408 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
409 |
lemma Rep_cb_take: |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
410 |
"Rep_compact_basis (cb_take (Suc n) a) = approx n\<cdot>(Rep_compact_basis a)" |
41370
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
411 |
by (simp add: cb_take.simps(2)) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
412 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
413 |
lemmas approx_Rep_compact_basis = Rep_cb_take [symmetric] |
27411 | 414 |
|
415 |
lemma cb_take_covers: "\<exists>n. cb_take n x = x" |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
416 |
apply (subgoal_tac "\<exists>n. cb_take (Suc n) x = x", fast) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
417 |
apply (simp add: Rep_compact_basis_inject [symmetric]) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
418 |
apply (simp add: Rep_cb_take) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
419 |
apply (rule compact_eq_approx) |
41370
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
420 |
apply (rule Rep_compact_basis') |
27411 | 421 |
done |
422 |
||
423 |
lemma cb_take_less: "cb_take n x \<sqsubseteq> x" |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
424 |
unfolding compact_le_def |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
425 |
by (cases n, simp, simp add: Rep_cb_take approx_below) |
27411 | 426 |
|
427 |
lemma cb_take_idem: "cb_take n (cb_take n x) = cb_take n x" |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
428 |
unfolding Rep_compact_basis_inject [symmetric] |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
429 |
by (cases n, simp, simp add: Rep_cb_take approx_idem) |
27411 | 430 |
|
431 |
lemma cb_take_mono: "x \<sqsubseteq> y \<Longrightarrow> cb_take n x \<sqsubseteq> cb_take n y" |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
432 |
unfolding compact_le_def |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
433 |
by (cases n, simp, simp add: Rep_cb_take monofun_cfun_arg) |
27411 | 434 |
|
435 |
lemma cb_take_chain_le: "m \<le> n \<Longrightarrow> cb_take m x \<sqsubseteq> cb_take n x" |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
436 |
unfolding compact_le_def |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
437 |
apply (cases m, simp, cases n, simp) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
438 |
apply (simp add: Rep_cb_take, rule chain_mono, simp, simp) |
27411 | 439 |
done |
440 |
||
441 |
lemma finite_range_cb_take: "finite (range (cb_take n))" |
|
442 |
apply (cases n) |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
443 |
apply (subgoal_tac "range (cb_take 0) = {compact_bot}", simp, force) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
444 |
apply (rule finite_imageD [where f="Rep_compact_basis"]) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
445 |
apply (rule finite_subset [where B="range (\<lambda>x. approx (n - 1)\<cdot>x)"]) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
446 |
apply (clarsimp simp add: Rep_cb_take) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
447 |
apply (rule finite_range_approx) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
448 |
apply (rule inj_onI, simp add: Rep_compact_basis_inject) |
27411 | 449 |
done |
450 |
||
81577 | 451 |
|
62175 | 452 |
subsubsection \<open>Rank of basis elements\<close> |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
453 |
|
27411 | 454 |
definition |
455 |
rank :: "'a compact_basis \<Rightarrow> nat" |
|
456 |
where |
|
457 |
"rank x = (LEAST n. cb_take n x = x)" |
|
458 |
||
459 |
lemma compact_approx_rank: "cb_take (rank x) x = x" |
|
460 |
unfolding rank_def |
|
461 |
apply (rule LeastI_ex) |
|
462 |
apply (rule cb_take_covers) |
|
463 |
done |
|
464 |
||
465 |
lemma rank_leD: "rank x \<le> n \<Longrightarrow> cb_take n 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
|
466 |
apply (rule below_antisym [OF cb_take_less]) |
27411 | 467 |
apply (subst compact_approx_rank [symmetric]) |
468 |
apply (erule cb_take_chain_le) |
|
469 |
done |
|
470 |
||
471 |
lemma rank_leI: "cb_take n x = x \<Longrightarrow> rank x \<le> n" |
|
472 |
unfolding rank_def by (rule Least_le) |
|
473 |
||
474 |
lemma rank_le_iff: "rank x \<le> n \<longleftrightarrow> cb_take n x = x" |
|
475 |
by (rule iffI [OF rank_leD rank_leI]) |
|
476 |
||
30505 | 477 |
lemma rank_compact_bot [simp]: "rank compact_bot = 0" |
478 |
using rank_leI [of 0 compact_bot] by simp |
|
479 |
||
480 |
lemma rank_eq_0_iff [simp]: "rank x = 0 \<longleftrightarrow> x = compact_bot" |
|
481 |
using rank_le_iff [of x 0] by auto |
|
482 |
||
27411 | 483 |
definition |
484 |
rank_le :: "'a compact_basis \<Rightarrow> 'a compact_basis set" |
|
485 |
where |
|
486 |
"rank_le x = {y. rank y \<le> rank x}" |
|
487 |
||
488 |
definition |
|
489 |
rank_lt :: "'a compact_basis \<Rightarrow> 'a compact_basis set" |
|
490 |
where |
|
491 |
"rank_lt x = {y. rank y < rank x}" |
|
492 |
||
493 |
definition |
|
494 |
rank_eq :: "'a compact_basis \<Rightarrow> 'a compact_basis set" |
|
495 |
where |
|
496 |
"rank_eq x = {y. rank y = rank x}" |
|
497 |
||
498 |
lemma rank_eq_cong: "rank x = rank y \<Longrightarrow> rank_eq x = rank_eq y" |
|
499 |
unfolding rank_eq_def by simp |
|
500 |
||
501 |
lemma rank_lt_cong: "rank x = rank y \<Longrightarrow> rank_lt x = rank_lt y" |
|
502 |
unfolding rank_lt_def by simp |
|
503 |
||
504 |
lemma rank_eq_subset: "rank_eq x \<subseteq> rank_le x" |
|
505 |
unfolding rank_eq_def rank_le_def by auto |
|
506 |
||
507 |
lemma rank_lt_subset: "rank_lt x \<subseteq> rank_le x" |
|
508 |
unfolding rank_lt_def rank_le_def by auto |
|
509 |
||
510 |
lemma finite_rank_le: "finite (rank_le x)" |
|
511 |
unfolding rank_le_def |
|
512 |
apply (rule finite_subset [where B="range (cb_take (rank x))"]) |
|
513 |
apply clarify |
|
514 |
apply (rule range_eqI) |
|
515 |
apply (erule rank_leD [symmetric]) |
|
516 |
apply (rule finite_range_cb_take) |
|
517 |
done |
|
518 |
||
519 |
lemma finite_rank_eq: "finite (rank_eq x)" |
|
520 |
by (rule finite_subset [OF rank_eq_subset finite_rank_le]) |
|
521 |
||
522 |
lemma finite_rank_lt: "finite (rank_lt x)" |
|
523 |
by (rule finite_subset [OF rank_lt_subset finite_rank_le]) |
|
524 |
||
525 |
lemma rank_lt_Int_rank_eq: "rank_lt x \<inter> rank_eq x = {}" |
|
526 |
unfolding rank_lt_def rank_eq_def rank_le_def by auto |
|
527 |
||
528 |
lemma rank_lt_Un_rank_eq: "rank_lt x \<union> rank_eq x = rank_le x" |
|
529 |
unfolding rank_lt_def rank_eq_def rank_le_def by auto |
|
530 |
||
81577 | 531 |
|
62175 | 532 |
subsubsection \<open>Sequencing basis elements\<close> |
27411 | 533 |
|
534 |
definition |
|
30505 | 535 |
place :: "'a compact_basis \<Rightarrow> nat" |
27411 | 536 |
where |
30505 | 537 |
"place x = card (rank_lt x) + choose_pos (rank_eq x) x" |
27411 | 538 |
|
30505 | 539 |
lemma place_bounded: "place x < card (rank_le x)" |
540 |
unfolding place_def |
|
27411 | 541 |
apply (rule ord_less_eq_trans) |
542 |
apply (rule add_strict_left_mono) |
|
543 |
apply (rule choose_pos_bounded) |
|
544 |
apply (rule finite_rank_eq) |
|
545 |
apply (simp add: rank_eq_def) |
|
546 |
apply (subst card_Un_disjoint [symmetric]) |
|
547 |
apply (rule finite_rank_lt) |
|
548 |
apply (rule finite_rank_eq) |
|
549 |
apply (rule rank_lt_Int_rank_eq) |
|
550 |
apply (simp add: rank_lt_Un_rank_eq) |
|
551 |
done |
|
552 |
||
30505 | 553 |
lemma place_ge: "card (rank_lt x) \<le> place x" |
554 |
unfolding place_def by simp |
|
27411 | 555 |
|
30505 | 556 |
lemma place_rank_mono: |
27411 | 557 |
fixes x y :: "'a compact_basis" |
30505 | 558 |
shows "rank x < rank y \<Longrightarrow> place x < place y" |
559 |
apply (rule less_le_trans [OF place_bounded]) |
|
560 |
apply (rule order_trans [OF _ place_ge]) |
|
27411 | 561 |
apply (rule card_mono) |
562 |
apply (rule finite_rank_lt) |
|
563 |
apply (simp add: rank_le_def rank_lt_def subset_eq) |
|
564 |
done |
|
565 |
||
30505 | 566 |
lemma place_eqD: "place x = place y \<Longrightarrow> x = y" |
27411 | 567 |
apply (rule linorder_cases [where x="rank x" and y="rank y"]) |
30505 | 568 |
apply (drule place_rank_mono, simp) |
569 |
apply (simp add: place_def) |
|
27411 | 570 |
apply (rule inj_on_choose_pos [where A="rank_eq x", THEN inj_onD]) |
571 |
apply (rule finite_rank_eq) |
|
572 |
apply (simp cong: rank_lt_cong rank_eq_cong) |
|
573 |
apply (simp add: rank_eq_def) |
|
574 |
apply (simp add: rank_eq_def) |
|
30505 | 575 |
apply (drule place_rank_mono, simp) |
27411 | 576 |
done |
577 |
||
30505 | 578 |
lemma inj_place: "inj place" |
579 |
by (rule inj_onI, erule place_eqD) |
|
27411 | 580 |
|
81577 | 581 |
|
62175 | 582 |
subsubsection \<open>Embedding and projection on basis elements\<close> |
27411 | 583 |
|
30505 | 584 |
definition |
585 |
sub :: "'a compact_basis \<Rightarrow> 'a compact_basis" |
|
586 |
where |
|
587 |
"sub x = (case rank x of 0 \<Rightarrow> compact_bot | Suc k \<Rightarrow> cb_take k x)" |
|
588 |
||
589 |
lemma rank_sub_less: "x \<noteq> compact_bot \<Longrightarrow> rank (sub x) < rank x" |
|
590 |
unfolding sub_def |
|
591 |
apply (cases "rank x", simp) |
|
592 |
apply (simp add: less_Suc_eq_le) |
|
593 |
apply (rule rank_leI) |
|
594 |
apply (rule cb_take_idem) |
|
595 |
done |
|
596 |
||
597 |
lemma place_sub_less: "x \<noteq> compact_bot \<Longrightarrow> place (sub x) < place x" |
|
598 |
apply (rule place_rank_mono) |
|
599 |
apply (erule rank_sub_less) |
|
600 |
done |
|
601 |
||
602 |
lemma sub_below: "sub x \<sqsubseteq> x" |
|
603 |
unfolding sub_def by (cases "rank x", simp_all add: cb_take_less) |
|
604 |
||
605 |
lemma rank_less_imp_below_sub: "\<lbrakk>x \<sqsubseteq> y; rank x < rank y\<rbrakk> \<Longrightarrow> x \<sqsubseteq> sub y" |
|
606 |
unfolding sub_def |
|
607 |
apply (cases "rank y", simp) |
|
608 |
apply (simp add: less_Suc_eq_le) |
|
609 |
apply (subgoal_tac "cb_take nat x \<sqsubseteq> cb_take nat y") |
|
610 |
apply (simp add: rank_leD) |
|
611 |
apply (erule cb_take_mono) |
|
612 |
done |
|
613 |
||
69661 | 614 |
function basis_emb :: "'a compact_basis \<Rightarrow> ubasis" |
615 |
where "basis_emb x = (if x = compact_bot then 0 else |
|
30505 | 616 |
node (place x) (basis_emb (sub x)) |
617 |
(basis_emb ` {y. place y < place x \<and> x \<sqsubseteq> y}))" |
|
69661 | 618 |
by simp_all |
27411 | 619 |
|
620 |
termination basis_emb |
|
69661 | 621 |
by (relation "measure place") (simp_all add: place_sub_less) |
27411 | 622 |
|
623 |
declare basis_emb.simps [simp del] |
|
624 |
||
69661 | 625 |
lemma basis_emb_compact_bot [simp]: |
626 |
"basis_emb compact_bot = 0" |
|
627 |
using basis_emb.simps [of compact_bot] by simp |
|
628 |
||
629 |
lemma basis_emb_rec: |
|
630 |
"basis_emb x = node (place x) (basis_emb (sub x)) (basis_emb ` {y. place y < place x \<and> x \<sqsubseteq> y})" |
|
631 |
if "x \<noteq> compact_bot" |
|
632 |
using that basis_emb.simps [of x] by simp |
|
633 |
||
634 |
lemma basis_emb_eq_0_iff [simp]: |
|
635 |
"basis_emb x = 0 \<longleftrightarrow> x = compact_bot" |
|
636 |
by (cases "x = compact_bot") (simp_all add: basis_emb_rec) |
|
27411 | 637 |
|
30505 | 638 |
lemma fin1: "finite {y. place y < place x \<and> x \<sqsubseteq> y}" |
27411 | 639 |
apply (subst Collect_conj_eq) |
640 |
apply (rule finite_Int) |
|
641 |
apply (rule disjI1) |
|
30505 | 642 |
apply (subgoal_tac "finite (place -` {n. n < place x})", simp) |
643 |
apply (rule finite_vimageI [OF _ inj_place]) |
|
27411 | 644 |
apply (simp add: lessThan_def [symmetric]) |
645 |
done |
|
646 |
||
30505 | 647 |
lemma fin2: "finite (basis_emb ` {y. place y < place x \<and> x \<sqsubseteq> y})" |
27411 | 648 |
by (rule finite_imageI [OF fin1]) |
649 |
||
30505 | 650 |
lemma rank_place_mono: |
651 |
"\<lbrakk>place x < place y; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> rank x < rank y" |
|
652 |
apply (rule linorder_cases, assumption) |
|
653 |
apply (simp add: place_def cong: rank_lt_cong rank_eq_cong) |
|
654 |
apply (drule choose_pos_lessD) |
|
655 |
apply (rule finite_rank_eq) |
|
656 |
apply (simp add: rank_eq_def) |
|
657 |
apply (simp add: rank_eq_def) |
|
658 |
apply simp |
|
659 |
apply (drule place_rank_mono, simp) |
|
660 |
done |
|
661 |
||
662 |
lemma basis_emb_mono: |
|
663 |
"x \<sqsubseteq> y \<Longrightarrow> ubasis_le (basis_emb x) (basis_emb y)" |
|
34915 | 664 |
proof (induct "max (place x) (place y)" arbitrary: x y rule: less_induct) |
665 |
case less |
|
30505 | 666 |
show ?case proof (rule linorder_cases) |
667 |
assume "place x < place y" |
|
668 |
then have "rank x < rank y" |
|
62175 | 669 |
using \<open>x \<sqsubseteq> y\<close> by (rule rank_place_mono) |
670 |
with \<open>place x < place y\<close> show ?case |
|
30505 | 671 |
apply (case_tac "y = compact_bot", simp) |
672 |
apply (simp add: basis_emb.simps [of y]) |
|
673 |
apply (rule ubasis_le_trans [OF _ ubasis_le_lower [OF fin2]]) |
|
34915 | 674 |
apply (rule less) |
30505 | 675 |
apply (simp add: less_max_iff_disj) |
676 |
apply (erule place_sub_less) |
|
62175 | 677 |
apply (erule rank_less_imp_below_sub [OF \<open>x \<sqsubseteq> y\<close>]) |
27411 | 678 |
done |
30505 | 679 |
next |
680 |
assume "place x = place y" |
|
681 |
hence "x = y" by (rule place_eqD) |
|
682 |
thus ?case by (simp add: ubasis_le_refl) |
|
683 |
next |
|
684 |
assume "place x > place y" |
|
62175 | 685 |
with \<open>x \<sqsubseteq> y\<close> show ?case |
30505 | 686 |
apply (case_tac "x = compact_bot", simp add: ubasis_le_minimal) |
687 |
apply (simp add: basis_emb.simps [of x]) |
|
688 |
apply (rule ubasis_le_upper [OF fin2], simp) |
|
34915 | 689 |
apply (rule less) |
30505 | 690 |
apply (simp add: less_max_iff_disj) |
691 |
apply (erule place_sub_less) |
|
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
|
692 |
apply (erule rev_below_trans) |
30505 | 693 |
apply (rule sub_below) |
694 |
done |
|
27411 | 695 |
qed |
696 |
qed |
|
697 |
||
698 |
lemma inj_basis_emb: "inj basis_emb" |
|
69661 | 699 |
proof (rule injI) |
700 |
fix x y |
|
701 |
assume "basis_emb x = basis_emb y" |
|
702 |
then show "x = y" |
|
703 |
by (cases "x = compact_bot \<or> y = compact_bot") (auto simp add: basis_emb_rec fin2 place_eqD) |
|
704 |
qed |
|
27411 | 705 |
|
706 |
definition |
|
30505 | 707 |
basis_prj :: "ubasis \<Rightarrow> 'a compact_basis" |
27411 | 708 |
where |
709 |
"basis_prj x = inv basis_emb |
|
30505 | 710 |
(ubasis_until (\<lambda>x. x \<in> range (basis_emb :: 'a compact_basis \<Rightarrow> ubasis)) x)" |
27411 | 711 |
|
712 |
lemma basis_prj_basis_emb: "\<And>x. basis_prj (basis_emb x) = x" |
|
713 |
unfolding basis_prj_def |
|
714 |
apply (subst ubasis_until_same) |
|
715 |
apply (rule rangeI) |
|
716 |
apply (rule inv_f_f) |
|
717 |
apply (rule inj_basis_emb) |
|
718 |
done |
|
719 |
||
720 |
lemma basis_prj_node: |
|
30505 | 721 |
"\<lbrakk>finite S; node i a S \<notin> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)\<rbrakk> |
722 |
\<Longrightarrow> basis_prj (node i a S) = (basis_prj a :: 'a compact_basis)" |
|
27411 | 723 |
unfolding basis_prj_def by simp |
724 |
||
725 |
lemma basis_prj_0: "basis_prj 0 = compact_bot" |
|
726 |
apply (subst basis_emb_compact_bot [symmetric]) |
|
727 |
apply (rule basis_prj_basis_emb) |
|
728 |
done |
|
729 |
||
30505 | 730 |
lemma node_eq_basis_emb_iff: |
731 |
"finite S \<Longrightarrow> node i a S = basis_emb x \<longleftrightarrow> |
|
732 |
x \<noteq> compact_bot \<and> i = place x \<and> a = basis_emb (sub x) \<and> |
|
733 |
S = basis_emb ` {y. place y < place x \<and> x \<sqsubseteq> y}" |
|
734 |
apply (cases "x = compact_bot", simp) |
|
735 |
apply (simp add: basis_emb.simps [of x]) |
|
736 |
apply (simp add: fin2) |
|
27411 | 737 |
done |
738 |
||
30505 | 739 |
lemma basis_prj_mono: "ubasis_le a b \<Longrightarrow> basis_prj a \<sqsubseteq> basis_prj b" |
740 |
proof (induct a b rule: ubasis_le.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
|
741 |
case (ubasis_le_refl a) show ?case by (rule below_refl) |
30505 | 742 |
next |
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
|
743 |
case (ubasis_le_trans a b c) thus ?case by - (rule below_trans) |
30505 | 744 |
next |
745 |
case (ubasis_le_lower S a i) thus ?case |
|
30561 | 746 |
apply (cases "node i a S \<in> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)") |
30505 | 747 |
apply (erule rangeE, rename_tac x) |
748 |
apply (simp add: basis_prj_basis_emb) |
|
749 |
apply (simp add: node_eq_basis_emb_iff) |
|
750 |
apply (simp add: basis_prj_basis_emb) |
|
751 |
apply (rule sub_below) |
|
752 |
apply (simp add: basis_prj_node) |
|
753 |
done |
|
754 |
next |
|
755 |
case (ubasis_le_upper S b a i) thus ?case |
|
30561 | 756 |
apply (cases "node i a S \<in> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)") |
30505 | 757 |
apply (erule rangeE, rename_tac x) |
758 |
apply (simp add: basis_prj_basis_emb) |
|
759 |
apply (clarsimp simp add: node_eq_basis_emb_iff) |
|
760 |
apply (simp add: basis_prj_basis_emb) |
|
761 |
apply (simp add: basis_prj_node) |
|
762 |
done |
|
763 |
qed |
|
764 |
||
27411 | 765 |
lemma basis_emb_prj_less: "ubasis_le (basis_emb (basis_prj x)) x" |
766 |
unfolding basis_prj_def |
|
33071
362f59fe5092
renamed f_inv_onto_f to f_inv_into_f (cf. 764547b68538);
wenzelm
parents:
32997
diff
changeset
|
767 |
apply (subst f_inv_into_f [where f=basis_emb]) |
27411 | 768 |
apply (rule ubasis_until) |
769 |
apply (rule range_eqI [where x=compact_bot]) |
|
770 |
apply simp |
|
771 |
apply (rule ubasis_until_less) |
|
772 |
done |
|
773 |
||
41286
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
774 |
lemma ideal_completion: |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
775 |
"ideal_completion below Rep_compact_basis (approximants :: 'a \<Rightarrow> _)" |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
776 |
proof |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
777 |
fix w :: "'a" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
778 |
show "below.ideal (approximants w)" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
779 |
proof (rule below.idealI) |
41370
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
780 |
have "Abs_compact_basis (approx 0\<cdot>w) \<in> approximants w" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
781 |
by (simp add: approximants_def approx_below) |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
782 |
thus "\<exists>x. x \<in> approximants w" .. |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
783 |
next |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
784 |
fix x y :: "'a compact_basis" |
41370
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
785 |
assume x: "x \<in> approximants w" and y: "y \<in> approximants w" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
786 |
obtain i where i: "approx i\<cdot>(Rep_compact_basis x) = Rep_compact_basis x" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
787 |
using compact_eq_approx Rep_compact_basis' by fast |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
788 |
obtain j where j: "approx j\<cdot>(Rep_compact_basis y) = Rep_compact_basis y" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
789 |
using compact_eq_approx Rep_compact_basis' by fast |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
790 |
let ?z = "Abs_compact_basis (approx (max i j)\<cdot>w)" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
791 |
have "?z \<in> approximants w" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
792 |
by (simp add: approximants_def approx_below) |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
793 |
moreover from x y have "x \<sqsubseteq> ?z \<and> y \<sqsubseteq> ?z" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
794 |
by (simp add: approximants_def compact_le_def) |
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
49834
diff
changeset
|
795 |
(metis i j monofun_cfun chain_mono chain_approx max.cobounded1 max.cobounded2) |
41370
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
796 |
ultimately show "\<exists>z \<in> approximants w. x \<sqsubseteq> z \<and> y \<sqsubseteq> z" .. |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
797 |
next |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
798 |
fix x y :: "'a compact_basis" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
799 |
assume "x \<sqsubseteq> y" "y \<in> approximants w" thus "x \<in> approximants w" |
41370
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
800 |
unfolding approximants_def compact_le_def |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
801 |
by (auto elim: below_trans) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
802 |
qed |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
803 |
next |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
804 |
fix Y :: "nat \<Rightarrow> 'a" |
41370
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
805 |
assume "chain Y" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
806 |
thus "approximants (\<Squnion>i. Y i) = (\<Union>i. approximants (Y i))" |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
807 |
unfolding approximants_def |
41370
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
808 |
by (auto simp add: compact_below_lub_iff) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
809 |
next |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
810 |
fix a :: "'a compact_basis" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
811 |
show "approximants (Rep_compact_basis a) = {b. b \<sqsubseteq> a}" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
812 |
unfolding approximants_def compact_le_def .. |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
813 |
next |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
814 |
fix x y :: "'a" |
41370
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
815 |
assume "approximants x \<subseteq> approximants y" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
816 |
hence "\<forall>z. compact z \<longrightarrow> z \<sqsubseteq> x \<longrightarrow> z \<sqsubseteq> y" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
817 |
by (simp add: approximants_def subset_eq) |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
818 |
(metis Abs_compact_basis_inverse') |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
819 |
hence "(\<Squnion>i. approx i\<cdot>x) \<sqsubseteq> y" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
820 |
by (simp add: lub_below approx_below) |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
821 |
thus "x \<sqsubseteq> y" |
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
huffman
parents:
41295
diff
changeset
|
822 |
by (simp add: lub_distribs) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
823 |
next |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
824 |
show "\<exists>f::'a compact_basis \<Rightarrow> nat. inj f" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
825 |
by (rule exI, rule inj_place) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
826 |
qed |
27411 | 827 |
|
41286
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
828 |
end |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
829 |
|
61605 | 830 |
interpretation compact_basis: |
41286
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
831 |
ideal_completion below Rep_compact_basis |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
832 |
"approximants :: 'a::bifinite \<Rightarrow> 'a compact_basis set" |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
833 |
proof - |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
834 |
obtain a :: "nat \<Rightarrow> 'a \<rightarrow> 'a" where "approx_chain a" |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
835 |
using bifinite .. |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
836 |
hence "bifinite_approx_chain a" |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
837 |
unfolding bifinite_approx_chain_def . |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
838 |
thus "ideal_completion below Rep_compact_basis (approximants :: 'a \<Rightarrow> _)" |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
839 |
by (rule bifinite_approx_chain.ideal_completion) |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
840 |
qed |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
841 |
|
81577 | 842 |
|
62175 | 843 |
subsubsection \<open>EP-pair from any bifinite domain into \emph{udom}\<close> |
27411 | 844 |
|
41286
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
845 |
context bifinite_approx_chain begin |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
846 |
|
27411 | 847 |
definition |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
848 |
udom_emb :: "'a \<rightarrow> udom" |
27411 | 849 |
where |
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41370
diff
changeset
|
850 |
"udom_emb = compact_basis.extension (\<lambda>x. udom_principal (basis_emb x))" |
27411 | 851 |
|
852 |
definition |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
853 |
udom_prj :: "udom \<rightarrow> 'a" |
27411 | 854 |
where |
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41370
diff
changeset
|
855 |
"udom_prj = udom.extension (\<lambda>x. Rep_compact_basis (basis_prj x))" |
27411 | 856 |
|
857 |
lemma udom_emb_principal: |
|
858 |
"udom_emb\<cdot>(Rep_compact_basis x) = udom_principal (basis_emb x)" |
|
859 |
unfolding udom_emb_def |
|
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41370
diff
changeset
|
860 |
apply (rule compact_basis.extension_principal) |
27411 | 861 |
apply (rule udom.principal_mono) |
862 |
apply (erule basis_emb_mono) |
|
863 |
done |
|
864 |
||
865 |
lemma udom_prj_principal: |
|
866 |
"udom_prj\<cdot>(udom_principal x) = Rep_compact_basis (basis_prj x)" |
|
867 |
unfolding udom_prj_def |
|
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41370
diff
changeset
|
868 |
apply (rule udom.extension_principal) |
27411 | 869 |
apply (rule compact_basis.principal_mono) |
870 |
apply (erule basis_prj_mono) |
|
871 |
done |
|
872 |
||
873 |
lemma ep_pair_udom: "ep_pair udom_emb udom_prj" |
|
61169 | 874 |
apply standard |
27411 | 875 |
apply (rule compact_basis.principal_induct, simp) |
876 |
apply (simp add: udom_emb_principal udom_prj_principal) |
|
877 |
apply (simp add: basis_prj_basis_emb) |
|
878 |
apply (rule udom.principal_induct, simp) |
|
879 |
apply (simp add: udom_emb_principal udom_prj_principal) |
|
880 |
apply (rule basis_emb_prj_less) |
|
881 |
done |
|
882 |
||
883 |
end |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
884 |
|
41286
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
885 |
abbreviation "udom_emb \<equiv> bifinite_approx_chain.udom_emb" |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
886 |
abbreviation "udom_prj \<equiv> bifinite_approx_chain.udom_prj" |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
887 |
|
41286
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
888 |
lemmas ep_pair_udom = |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
889 |
bifinite_approx_chain.ep_pair_udom [unfolded bifinite_approx_chain_def] |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
890 |
|
81577 | 891 |
|
62175 | 892 |
subsection \<open>Chain of approx functions for type \emph{udom}\<close> |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
893 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
894 |
definition |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
895 |
udom_approx :: "nat \<Rightarrow> udom \<rightarrow> udom" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
896 |
where |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
897 |
"udom_approx i = |
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41370
diff
changeset
|
898 |
udom.extension (\<lambda>x. udom_principal (ubasis_until (\<lambda>y. y \<le> i) x))" |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
899 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
900 |
lemma udom_approx_mono: |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
901 |
"ubasis_le a b \<Longrightarrow> |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
902 |
udom_principal (ubasis_until (\<lambda>y. y \<le> i) a) \<sqsubseteq> |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
903 |
udom_principal (ubasis_until (\<lambda>y. y \<le> i) b)" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
904 |
apply (rule udom.principal_mono) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
905 |
apply (rule ubasis_until_mono) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
906 |
apply (frule (2) order_less_le_trans [OF node_gt2]) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
907 |
apply (erule order_less_imp_le) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
908 |
apply assumption |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
909 |
done |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
910 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
911 |
lemma adm_mem_finite: "\<lbrakk>cont f; finite S\<rbrakk> \<Longrightarrow> adm (\<lambda>x. f x \<in> S)" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
912 |
by (erule adm_subst, induct set: finite, simp_all) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
913 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
914 |
lemma udom_approx_principal: |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
915 |
"udom_approx i\<cdot>(udom_principal x) = |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
916 |
udom_principal (ubasis_until (\<lambda>y. y \<le> i) x)" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
917 |
unfolding udom_approx_def |
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41370
diff
changeset
|
918 |
apply (rule udom.extension_principal) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
919 |
apply (erule udom_approx_mono) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
920 |
done |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
921 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
922 |
lemma finite_deflation_udom_approx: "finite_deflation (udom_approx i)" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
923 |
proof |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
924 |
fix x show "udom_approx i\<cdot>(udom_approx i\<cdot>x) = udom_approx i\<cdot>x" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
925 |
by (induct x rule: udom.principal_induct, simp) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
926 |
(simp add: udom_approx_principal ubasis_until_idem) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
927 |
next |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
928 |
fix x show "udom_approx i\<cdot>x \<sqsubseteq> x" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
929 |
by (induct x rule: udom.principal_induct, simp) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
930 |
(simp add: udom_approx_principal ubasis_until_less) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
931 |
next |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
932 |
have *: "finite (range (\<lambda>x. udom_principal (ubasis_until (\<lambda>y. y \<le> i) x)))" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
933 |
apply (subst range_composition [where f=udom_principal]) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
934 |
apply (simp add: finite_range_ubasis_until) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
935 |
done |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
936 |
show "finite {x. udom_approx i\<cdot>x = x}" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
937 |
apply (rule finite_range_imp_finite_fixes) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
938 |
apply (rule rev_finite_subset [OF *]) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
939 |
apply (clarsimp, rename_tac x) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
940 |
apply (induct_tac x rule: udom.principal_induct) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
941 |
apply (simp add: adm_mem_finite *) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
942 |
apply (simp add: udom_approx_principal) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
943 |
done |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
944 |
qed |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
945 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
946 |
interpretation udom_approx: finite_deflation "udom_approx i" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
947 |
by (rule finite_deflation_udom_approx) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
948 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
949 |
lemma chain_udom_approx [simp]: "chain (\<lambda>i. udom_approx i)" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
950 |
unfolding udom_approx_def |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
951 |
apply (rule chainI) |
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41370
diff
changeset
|
952 |
apply (rule udom.extension_mono) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
953 |
apply (erule udom_approx_mono) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
954 |
apply (erule udom_approx_mono) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
955 |
apply (rule udom.principal_mono) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
956 |
apply (rule ubasis_until_chain, simp) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
957 |
done |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
958 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
959 |
lemma lub_udom_approx [simp]: "(\<Squnion>i. udom_approx i) = ID" |
40002
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents:
39984
diff
changeset
|
960 |
apply (rule cfun_eqI, simp add: contlub_cfun_fun) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
961 |
apply (rule below_antisym) |
40500
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents:
40002
diff
changeset
|
962 |
apply (rule lub_below) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
963 |
apply (simp) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
964 |
apply (rule udom_approx.below) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
965 |
apply (rule_tac x=x in udom.principal_induct) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
966 |
apply (simp add: lub_distribs) |
40500
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents:
40002
diff
changeset
|
967 |
apply (rule_tac i=a in below_lub) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
968 |
apply simp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
969 |
apply (simp add: udom_approx_principal) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
970 |
apply (simp add: ubasis_until_same ubasis_le_refl) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
971 |
done |
65552
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
wenzelm
parents:
64267
diff
changeset
|
972 |
|
41286
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
973 |
lemma udom_approx [simp]: "approx_chain udom_approx" |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
974 |
proof |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
975 |
show "chain (\<lambda>i. udom_approx i)" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
976 |
by (rule chain_udom_approx) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
977 |
show "(\<Squnion>i. udom_approx i) = ID" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
978 |
by (rule lub_udom_approx) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
979 |
qed |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
980 |
|
41286
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
981 |
instance udom :: bifinite |
61169 | 982 |
by standard (fast intro: udom_approx) |
41286
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41182
diff
changeset
|
983 |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
984 |
hide_const (open) node |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
985 |
|
81579 | 986 |
unbundle binomial_syntax |
65552
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
wenzelm
parents:
64267
diff
changeset
|
987 |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
36452
diff
changeset
|
988 |
end |