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