| author | wenzelm | 
| Tue, 04 Apr 2017 21:57:43 +0200 | |
| changeset 65380 | ae93953746fc | 
| parent 65379 | 76a96e32bd23 | 
| child 68383 | 93a42bd62ede | 
| permissions | -rw-r--r-- | 
| 42151 | 1  | 
(* Title: HOL/HOLCF/Completion.thy  | 
| 27404 | 2  | 
Author: Brian Huffman  | 
3  | 
*)  | 
|
4  | 
||
| 62175 | 5  | 
section \<open>Defining algebraic domains by ideal completion\<close>  | 
| 27404 | 6  | 
|
7  | 
theory Completion  | 
|
| 
65380
 
ae93953746fc
eliminated Plain_HOLCF.thy (see also 8e92772bc0e8): it was modeled after HOL/Plain.thy which was discontinued later;
 
wenzelm 
parents: 
65379 
diff
changeset
 | 
8  | 
imports Cfun  | 
| 27404 | 9  | 
begin  | 
10  | 
||
| 62175 | 11  | 
subsection \<open>Ideals over a preorder\<close>  | 
| 27404 | 12  | 
|
13  | 
locale preorder =  | 
|
14  | 
fixes r :: "'a::type \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<preceq>" 50)  | 
|
15  | 
assumes r_refl: "x \<preceq> x"  | 
|
16  | 
assumes r_trans: "\<lbrakk>x \<preceq> y; y \<preceq> z\<rbrakk> \<Longrightarrow> x \<preceq> z"  | 
|
17  | 
begin  | 
|
18  | 
||
19  | 
definition  | 
|
20  | 
ideal :: "'a set \<Rightarrow> bool" where  | 
|
21  | 
"ideal A = ((\<exists>x. x \<in> A) \<and> (\<forall>x\<in>A. \<forall>y\<in>A. \<exists>z\<in>A. x \<preceq> z \<and> y \<preceq> z) \<and>  | 
|
22  | 
(\<forall>x y. x \<preceq> y \<longrightarrow> y \<in> A \<longrightarrow> x \<in> A))"  | 
|
23  | 
||
24  | 
lemma idealI:  | 
|
25  | 
assumes "\<exists>x. x \<in> A"  | 
|
26  | 
assumes "\<And>x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> \<exists>z\<in>A. x \<preceq> z \<and> y \<preceq> z"  | 
|
27  | 
assumes "\<And>x y. \<lbrakk>x \<preceq> y; y \<in> A\<rbrakk> \<Longrightarrow> x \<in> A"  | 
|
28  | 
shows "ideal A"  | 
|
| 41529 | 29  | 
unfolding ideal_def using assms by fast  | 
| 27404 | 30  | 
|
31  | 
lemma idealD1:  | 
|
32  | 
"ideal A \<Longrightarrow> \<exists>x. x \<in> A"  | 
|
33  | 
unfolding ideal_def by fast  | 
|
34  | 
||
35  | 
lemma idealD2:  | 
|
36  | 
"\<lbrakk>ideal A; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> \<exists>z\<in>A. x \<preceq> z \<and> y \<preceq> z"  | 
|
37  | 
unfolding ideal_def by fast  | 
|
38  | 
||
39  | 
lemma idealD3:  | 
|
40  | 
"\<lbrakk>ideal A; x \<preceq> y; y \<in> A\<rbrakk> \<Longrightarrow> x \<in> A"  | 
|
41  | 
unfolding ideal_def by fast  | 
|
42  | 
||
43  | 
lemma ideal_principal: "ideal {x. x \<preceq> z}"
 | 
|
44  | 
apply (rule idealI)  | 
|
45  | 
apply (rule_tac x=z in exI)  | 
|
46  | 
apply (fast intro: r_refl)  | 
|
47  | 
apply (rule_tac x=z in bexI, fast)  | 
|
48  | 
apply (fast intro: r_refl)  | 
|
49  | 
apply (fast intro: r_trans)  | 
|
50  | 
done  | 
|
51  | 
||
| 
40888
 
28cd51cff70c
reformulate lemma preorder.ex_ideal, and use it for typedefs
 
huffman 
parents: 
40774 
diff
changeset
 | 
52  | 
lemma ex_ideal: "\<exists>A. A \<in> {A. ideal A}"
 | 
| 
 
28cd51cff70c
reformulate lemma preorder.ex_ideal, and use it for typedefs
 
huffman 
parents: 
40774 
diff
changeset
 | 
53  | 
by (fast intro: ideal_principal)  | 
| 27404 | 54  | 
|
| 62175 | 55  | 
text \<open>The set of ideals is a cpo\<close>  | 
| 27404 | 56  | 
|
57  | 
lemma ideal_UN:  | 
|
58  | 
fixes A :: "nat \<Rightarrow> 'a set"  | 
|
59  | 
assumes ideal_A: "\<And>i. ideal (A i)"  | 
|
60  | 
assumes chain_A: "\<And>i j. i \<le> j \<Longrightarrow> A i \<subseteq> A j"  | 
|
61  | 
shows "ideal (\<Union>i. A i)"  | 
|
62  | 
apply (rule idealI)  | 
|
63  | 
apply (cut_tac idealD1 [OF ideal_A], fast)  | 
|
64  | 
apply (clarify, rename_tac i j)  | 
|
| 
54863
 
82acc20ded73
prefer more canonical names for lemmas on min/max
 
haftmann 
parents: 
42151 
diff
changeset
 | 
65  | 
apply (drule subsetD [OF chain_A [OF max.cobounded1]])  | 
| 
 
82acc20ded73
prefer more canonical names for lemmas on min/max
 
haftmann 
parents: 
42151 
diff
changeset
 | 
66  | 
apply (drule subsetD [OF chain_A [OF max.cobounded2]])  | 
| 27404 | 67  | 
apply (drule (1) idealD2 [OF ideal_A])  | 
68  | 
apply blast  | 
|
69  | 
apply clarify  | 
|
70  | 
apply (drule (1) idealD3 [OF ideal_A])  | 
|
71  | 
apply fast  | 
|
72  | 
done  | 
|
73  | 
||
74  | 
lemma typedef_ideal_po:  | 
|
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
30729 
diff
changeset
 | 
75  | 
fixes Abs :: "'a set \<Rightarrow> 'b::below"  | 
| 27404 | 76  | 
  assumes type: "type_definition Rep Abs {S. ideal S}"
 | 
| 
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
 | 
77  | 
assumes below: "\<And>x y. x \<sqsubseteq> y \<longleftrightarrow> Rep x \<subseteq> Rep y"  | 
| 27404 | 78  | 
  shows "OFCLASS('b, po_class)"
 | 
| 
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
 | 
79  | 
apply (intro_classes, unfold below)  | 
| 27404 | 80  | 
apply (rule subset_refl)  | 
81  | 
apply (erule (1) subset_trans)  | 
|
82  | 
apply (rule type_definition.Rep_inject [OF type, THEN iffD1])  | 
|
83  | 
apply (erule (1) subset_antisym)  | 
|
84  | 
done  | 
|
85  | 
||
86  | 
lemma  | 
|
87  | 
fixes Abs :: "'a set \<Rightarrow> 'b::po"  | 
|
88  | 
  assumes type: "type_definition Rep Abs {S. ideal S}"
 | 
|
| 
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
 | 
89  | 
assumes below: "\<And>x y. x \<sqsubseteq> y \<longleftrightarrow> Rep x \<subseteq> Rep y"  | 
| 27404 | 90  | 
assumes S: "chain S"  | 
91  | 
shows typedef_ideal_lub: "range S <<| Abs (\<Union>i. Rep (S i))"  | 
|
| 40769 | 92  | 
and typedef_ideal_rep_lub: "Rep (\<Squnion>i. S i) = (\<Union>i. Rep (S i))"  | 
| 27404 | 93  | 
proof -  | 
94  | 
have 1: "ideal (\<Union>i. Rep (S i))"  | 
|
95  | 
apply (rule ideal_UN)  | 
|
96  | 
apply (rule type_definition.Rep [OF type, unfolded mem_Collect_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
 | 
97  | 
apply (subst below [symmetric])  | 
| 27404 | 98  | 
apply (erule chain_mono [OF S])  | 
99  | 
done  | 
|
100  | 
hence 2: "Rep (Abs (\<Union>i. Rep (S i))) = (\<Union>i. Rep (S i))"  | 
|
101  | 
by (simp add: type_definition.Abs_inverse [OF type])  | 
|
102  | 
show 3: "range S <<| Abs (\<Union>i. Rep (S i))"  | 
|
103  | 
apply (rule is_lubI)  | 
|
104  | 
apply (rule is_ubI)  | 
|
| 
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
 | 
105  | 
apply (simp add: below 2, fast)  | 
| 
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
30729 
diff
changeset
 | 
106  | 
apply (simp add: below 2 is_ub_def, fast)  | 
| 27404 | 107  | 
done  | 
108  | 
hence 4: "(\<Squnion>i. S i) = Abs (\<Union>i. Rep (S i))"  | 
|
| 40771 | 109  | 
by (rule lub_eqI)  | 
| 27404 | 110  | 
show 5: "Rep (\<Squnion>i. S i) = (\<Union>i. Rep (S i))"  | 
111  | 
by (simp add: 4 2)  | 
|
112  | 
qed  | 
|
113  | 
||
114  | 
lemma typedef_ideal_cpo:  | 
|
115  | 
fixes Abs :: "'a set \<Rightarrow> 'b::po"  | 
|
116  | 
  assumes type: "type_definition Rep Abs {S. ideal S}"
 | 
|
| 
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
 | 
117  | 
assumes below: "\<And>x y. x \<sqsubseteq> y \<longleftrightarrow> Rep x \<subseteq> Rep y"  | 
| 27404 | 118  | 
  shows "OFCLASS('b, cpo_class)"
 | 
| 61169 | 119  | 
by standard (rule exI, erule typedef_ideal_lub [OF type below])  | 
| 27404 | 120  | 
|
121  | 
end  | 
|
122  | 
||
| 
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
 | 
123  | 
interpretation below: preorder "below :: 'a::po \<Rightarrow> 'a \<Rightarrow> bool"  | 
| 27404 | 124  | 
apply unfold_locales  | 
| 
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
 | 
125  | 
apply (rule below_refl)  | 
| 
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
30729 
diff
changeset
 | 
126  | 
apply (erule (1) below_trans)  | 
| 27404 | 127  | 
done  | 
128  | 
||
| 62175 | 129  | 
subsection \<open>Lemmas about least upper bounds\<close>  | 
| 27404 | 130  | 
|
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
131  | 
lemma is_ub_thelub_ex: "\<lbrakk>\<exists>u. S <<| u; x \<in> S\<rbrakk> \<Longrightarrow> x \<sqsubseteq> lub S"  | 
| 40771 | 132  | 
apply (erule exE, drule is_lub_lub)  | 
| 27404 | 133  | 
apply (drule is_lubD1)  | 
134  | 
apply (erule (1) is_ubD)  | 
|
135  | 
done  | 
|
136  | 
||
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
137  | 
lemma is_lub_thelub_ex: "\<lbrakk>\<exists>u. S <<| u; S <| x\<rbrakk> \<Longrightarrow> lub S \<sqsubseteq> x"  | 
| 40771 | 138  | 
by (erule exE, drule is_lub_lub, erule is_lubD2)  | 
| 27404 | 139  | 
|
| 65379 | 140  | 
|
| 62175 | 141  | 
subsection \<open>Locale for ideal completion\<close>  | 
| 28133 | 142  | 
|
| 65379 | 143  | 
hide_const (open) Filter.principal  | 
144  | 
||
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
145  | 
locale ideal_completion = preorder +  | 
| 27404 | 146  | 
fixes principal :: "'a::type \<Rightarrow> 'b::cpo"  | 
147  | 
fixes rep :: "'b::cpo \<Rightarrow> 'a::type set"  | 
|
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
148  | 
assumes ideal_rep: "\<And>x. ideal (rep x)"  | 
| 40769 | 149  | 
assumes rep_lub: "\<And>Y. chain Y \<Longrightarrow> rep (\<Squnion>i. Y i) = (\<Union>i. rep (Y i))"  | 
| 27404 | 150  | 
  assumes rep_principal: "\<And>a. rep (principal a) = {b. b \<preceq> a}"
 | 
| 41033 | 151  | 
assumes belowI: "\<And>x y. rep x \<subseteq> rep y \<Longrightarrow> x \<sqsubseteq> y"  | 
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
152  | 
assumes countable: "\<exists>f::'a \<Rightarrow> nat. inj f"  | 
| 27404 | 153  | 
begin  | 
154  | 
||
| 28133 | 155  | 
lemma rep_mono: "x \<sqsubseteq> y \<Longrightarrow> rep x \<subseteq> rep y"  | 
156  | 
apply (frule bin_chain)  | 
|
| 40769 | 157  | 
apply (drule rep_lub)  | 
| 40771 | 158  | 
apply (simp only: lub_eqI [OF is_lub_bin_chain])  | 
| 28133 | 159  | 
apply (rule subsetI, rule UN_I [where a=0], simp_all)  | 
160  | 
done  | 
|
161  | 
||
| 
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
 | 
162  | 
lemma below_def: "x \<sqsubseteq> y \<longleftrightarrow> rep x \<subseteq> rep y"  | 
| 41033 | 163  | 
by (rule iffI [OF rep_mono belowI])  | 
| 28133 | 164  | 
|
| 
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
 | 
165  | 
lemma principal_below_iff_mem_rep: "principal a \<sqsubseteq> x \<longleftrightarrow> a \<in> rep x"  | 
| 41033 | 166  | 
unfolding below_def rep_principal  | 
167  | 
by (auto intro: r_refl elim: idealD3 [OF ideal_rep])  | 
|
| 28133 | 168  | 
|
| 
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
 | 
169  | 
lemma principal_below_iff [simp]: "principal a \<sqsubseteq> principal b \<longleftrightarrow> a \<preceq> b"  | 
| 
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
30729 
diff
changeset
 | 
170  | 
by (simp add: principal_below_iff_mem_rep rep_principal)  | 
| 28133 | 171  | 
|
172  | 
lemma principal_eq_iff: "principal a = principal b \<longleftrightarrow> a \<preceq> b \<and> b \<preceq> a"  | 
|
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
30729 
diff
changeset
 | 
173  | 
unfolding po_eq_conv [where 'a='b] principal_below_iff ..  | 
| 28133 | 174  | 
|
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
175  | 
lemma eq_iff: "x = y \<longleftrightarrow> rep x = rep y"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
176  | 
unfolding po_eq_conv below_def by auto  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
177  | 
|
| 28133 | 178  | 
lemma principal_mono: "a \<preceq> b \<Longrightarrow> principal a \<sqsubseteq> principal b"  | 
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
30729 
diff
changeset
 | 
179  | 
by (simp only: principal_below_iff)  | 
| 28133 | 180  | 
|
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
181  | 
lemma ch2ch_principal [simp]:  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
182  | 
"\<forall>i. Y i \<preceq> Y (Suc i) \<Longrightarrow> chain (\<lambda>i. principal (Y i))"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
183  | 
by (simp add: chainI principal_mono)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
184  | 
|
| 62175 | 185  | 
subsubsection \<open>Principal ideals approximate all elements\<close>  | 
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
186  | 
|
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
187  | 
lemma compact_principal [simp]: "compact (principal a)"  | 
| 40769 | 188  | 
by (rule compactI2, simp add: principal_below_iff_mem_rep rep_lub)  | 
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
189  | 
|
| 62175 | 190  | 
text \<open>Construct a chain whose lub is the same as a given ideal\<close>  | 
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
191  | 
|
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
192  | 
lemma obtain_principal_chain:  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
193  | 
obtains Y where "\<forall>i. Y i \<preceq> Y (Suc i)" and "x = (\<Squnion>i. principal (Y i))"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
194  | 
proof -  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
195  | 
obtain count :: "'a \<Rightarrow> nat" where inj: "inj count"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
196  | 
using countable ..  | 
| 63040 | 197  | 
define enum where "enum i = (THE a. count a = i)" for i  | 
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
198  | 
have enum_count [simp]: "\<And>x. enum (count x) = x"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
199  | 
unfolding enum_def by (simp add: inj_eq [OF inj])  | 
| 63040 | 200  | 
define a where "a = (LEAST i. enum i \<in> rep x)"  | 
201  | 
define b where "b i = (LEAST j. enum j \<in> rep x \<and> \<not> enum j \<preceq> enum i)" for i  | 
|
202  | 
define c where "c i j = (LEAST k. enum k \<in> rep x \<and> enum i \<preceq> enum k \<and> enum j \<preceq> enum k)" for i j  | 
|
203  | 
define P where "P i \<longleftrightarrow> (\<exists>j. enum j \<in> rep x \<and> \<not> enum j \<preceq> enum i)" for i  | 
|
204  | 
define X where "X = rec_nat a (\<lambda>n i. if P i then c i (b i) else i)"  | 
|
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
205  | 
have X_0: "X 0 = a" unfolding X_def by simp  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
206  | 
have X_Suc: "\<And>n. X (Suc n) = (if P (X n) then c (X n) (b (X n)) else X n)"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
207  | 
unfolding X_def by simp  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
208  | 
have a_mem: "enum a \<in> rep x"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
209  | 
unfolding a_def  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
210  | 
apply (rule LeastI_ex)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
211  | 
apply (cut_tac ideal_rep [of x])  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
212  | 
apply (drule idealD1)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
213  | 
apply (clarify, rename_tac a)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
214  | 
apply (rule_tac x="count a" in exI, simp)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
215  | 
done  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
216  | 
have b: "\<And>i. P i \<Longrightarrow> enum i \<in> rep x  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
217  | 
\<Longrightarrow> enum (b i) \<in> rep x \<and> \<not> enum (b i) \<preceq> enum i"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
218  | 
unfolding P_def b_def by (erule LeastI2_ex, simp)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
219  | 
have c: "\<And>i j. enum i \<in> rep x \<Longrightarrow> enum j \<in> rep x  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
220  | 
\<Longrightarrow> enum (c i j) \<in> rep x \<and> enum i \<preceq> enum (c i j) \<and> enum j \<preceq> enum (c i j)"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
221  | 
unfolding c_def  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
222  | 
apply (drule (1) idealD2 [OF ideal_rep], clarify)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
223  | 
apply (rule_tac a="count z" in LeastI2, simp, simp)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
224  | 
done  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
225  | 
have X_mem: "\<And>n. enum (X n) \<in> rep x"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
226  | 
apply (induct_tac n)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
227  | 
apply (simp add: X_0 a_mem)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
228  | 
apply (clarsimp simp add: X_Suc, rename_tac n)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
229  | 
apply (simp add: b c)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
230  | 
done  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
231  | 
have X_chain: "\<And>n. enum (X n) \<preceq> enum (X (Suc n))"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
232  | 
apply (clarsimp simp add: X_Suc r_refl)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
233  | 
apply (simp add: b c X_mem)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
234  | 
done  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
235  | 
have less_b: "\<And>n i. n < b i \<Longrightarrow> enum n \<in> rep x \<Longrightarrow> enum n \<preceq> enum i"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
236  | 
unfolding b_def by (drule not_less_Least, simp)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
237  | 
have X_covers: "\<And>n. \<forall>k\<le>n. enum k \<in> rep x \<longrightarrow> enum k \<preceq> enum (X n)"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
238  | 
apply (induct_tac n)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
239  | 
apply (clarsimp simp add: X_0 a_def)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
240  | 
apply (drule_tac k=0 in Least_le, simp add: r_refl)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
241  | 
apply (clarsimp, rename_tac n k)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
242  | 
apply (erule le_SucE)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
243  | 
apply (rule r_trans [OF _ X_chain], simp)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
244  | 
apply (case_tac "P (X n)", simp add: X_Suc)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
245  | 
apply (rule_tac x="b (X n)" and y="Suc n" in linorder_cases)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
246  | 
apply (simp only: less_Suc_eq_le)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
247  | 
apply (drule spec, drule (1) mp, simp add: b X_mem)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
248  | 
apply (simp add: c X_mem)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
249  | 
apply (drule (1) less_b)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
250  | 
apply (erule r_trans)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
251  | 
apply (simp add: b c X_mem)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
252  | 
apply (simp add: X_Suc)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
253  | 
apply (simp add: P_def)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
254  | 
done  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
255  | 
have 1: "\<forall>i. enum (X i) \<preceq> enum (X (Suc i))"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
256  | 
by (simp add: X_chain)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
257  | 
have 2: "x = (\<Squnion>n. principal (enum (X n)))"  | 
| 40769 | 258  | 
apply (simp add: eq_iff rep_lub 1 rep_principal)  | 
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
259  | 
apply (auto, rename_tac a)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
260  | 
apply (subgoal_tac "\<exists>i. a = enum i", erule exE)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
261  | 
apply (rule_tac x=i in exI, simp add: X_covers)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
262  | 
apply (rule_tac x="count a" in exI, simp)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
263  | 
apply (erule idealD3 [OF ideal_rep])  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
264  | 
apply (rule X_mem)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
265  | 
done  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
266  | 
from 1 2 show ?thesis ..  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
267  | 
qed  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
268  | 
|
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
269  | 
lemma principal_induct:  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
270  | 
assumes adm: "adm P"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
271  | 
assumes P: "\<And>a. P (principal a)"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
272  | 
shows "P x"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
273  | 
apply (rule obtain_principal_chain [of x])  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
274  | 
apply (simp add: admD [OF adm] P)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
275  | 
done  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
276  | 
|
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
277  | 
lemma compact_imp_principal: "compact x \<Longrightarrow> \<exists>a. x = principal a"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
278  | 
apply (rule obtain_principal_chain [of x])  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
279  | 
apply (drule adm_compact_neq [OF _ cont_id])  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
280  | 
apply (subgoal_tac "chain (\<lambda>i. principal (Y i))")  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
281  | 
apply (drule (2) admD2, fast, simp)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
282  | 
done  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
283  | 
|
| 62175 | 284  | 
subsection \<open>Defining functions in terms of basis elements\<close>  | 
| 28133 | 285  | 
|
286  | 
definition  | 
|
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
287  | 
  extension :: "('a::type \<Rightarrow> 'c::cpo) \<Rightarrow> 'b \<rightarrow> 'c" where
 | 
| 
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
288  | 
"extension = (\<lambda>f. (\<Lambda> x. lub (f ` rep x)))"  | 
| 28133 | 289  | 
|
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
290  | 
lemma extension_lemma:  | 
| 27404 | 291  | 
fixes f :: "'a::type \<Rightarrow> 'c::cpo"  | 
292  | 
assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"  | 
|
293  | 
shows "\<exists>u. f ` rep x <<| u"  | 
|
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
294  | 
proof -  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
295  | 
obtain Y where Y: "\<forall>i. Y i \<preceq> Y (Suc i)"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
296  | 
and x: "x = (\<Squnion>i. principal (Y i))"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
297  | 
by (rule obtain_principal_chain [of x])  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
298  | 
have chain: "chain (\<lambda>i. f (Y i))"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
299  | 
by (rule chainI, simp add: f_mono Y)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
300  | 
  have rep_x: "rep x = (\<Union>n. {a. a \<preceq> Y n})"
 | 
| 40769 | 301  | 
by (simp add: x rep_lub Y rep_principal)  | 
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
302  | 
have "f ` rep x <<| (\<Squnion>n. f (Y n))"  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
303  | 
apply (rule is_lubI)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
304  | 
apply (rule ub_imageI, rename_tac a)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
305  | 
apply (clarsimp simp add: rep_x)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
306  | 
apply (drule f_mono)  | 
| 
40500
 
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
 
huffman 
parents: 
40002 
diff
changeset
 | 
307  | 
apply (erule below_lub [OF chain])  | 
| 
 
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
 
huffman 
parents: 
40002 
diff
changeset
 | 
308  | 
apply (rule lub_below [OF chain])  | 
| 
 
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
 
huffman 
parents: 
40002 
diff
changeset
 | 
309  | 
apply (drule_tac x="Y n" in ub_imageD)  | 
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
310  | 
apply (simp add: rep_x, fast intro: r_refl)  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
311  | 
apply assumption  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
312  | 
done  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
313  | 
thus ?thesis ..  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
314  | 
qed  | 
| 27404 | 315  | 
|
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
316  | 
lemma extension_beta:  | 
| 27404 | 317  | 
fixes f :: "'a::type \<Rightarrow> 'c::cpo"  | 
318  | 
assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"  | 
|
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
319  | 
shows "extension f\<cdot>x = lub (f ` rep x)"  | 
| 
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
320  | 
unfolding extension_def  | 
| 27404 | 321  | 
proof (rule beta_cfun)  | 
322  | 
have lub: "\<And>x. \<exists>u. f ` rep x <<| u"  | 
|
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
323  | 
using f_mono by (rule extension_lemma)  | 
| 27404 | 324  | 
show cont: "cont (\<lambda>x. lub (f ` rep x))"  | 
325  | 
apply (rule contI2)  | 
|
326  | 
apply (rule monofunI)  | 
|
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
327  | 
apply (rule is_lub_thelub_ex [OF lub ub_imageI])  | 
| 
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
328  | 
apply (rule is_ub_thelub_ex [OF lub imageI])  | 
| 27404 | 329  | 
apply (erule (1) subsetD [OF rep_mono])  | 
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
330  | 
apply (rule is_lub_thelub_ex [OF lub ub_imageI])  | 
| 40769 | 331  | 
apply (simp add: rep_lub, clarify)  | 
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
30729 
diff
changeset
 | 
332  | 
apply (erule rev_below_trans [OF is_ub_thelub])  | 
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
333  | 
apply (erule is_ub_thelub_ex [OF lub imageI])  | 
| 27404 | 334  | 
done  | 
335  | 
qed  | 
|
336  | 
||
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
337  | 
lemma extension_principal:  | 
| 27404 | 338  | 
fixes f :: "'a::type \<Rightarrow> 'c::cpo"  | 
339  | 
assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"  | 
|
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
340  | 
shows "extension f\<cdot>(principal a) = f a"  | 
| 
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
341  | 
apply (subst extension_beta, erule f_mono)  | 
| 27404 | 342  | 
apply (subst rep_principal)  | 
| 41033 | 343  | 
apply (rule lub_eqI)  | 
344  | 
apply (rule is_lub_maximal)  | 
|
345  | 
apply (rule ub_imageI)  | 
|
346  | 
apply (simp add: f_mono)  | 
|
347  | 
apply (rule imageI)  | 
|
348  | 
apply (simp add: r_refl)  | 
|
| 27404 | 349  | 
done  | 
350  | 
||
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
351  | 
lemma extension_mono:  | 
| 27404 | 352  | 
assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"  | 
353  | 
assumes g_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> g a \<sqsubseteq> g b"  | 
|
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
30729 
diff
changeset
 | 
354  | 
assumes below: "\<And>a. f a \<sqsubseteq> g a"  | 
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
355  | 
shows "extension f \<sqsubseteq> extension g"  | 
| 
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
 | 
356  | 
apply (rule cfun_belowI)  | 
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
357  | 
apply (simp only: extension_beta f_mono g_mono)  | 
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
358  | 
apply (rule is_lub_thelub_ex)  | 
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
359  | 
apply (rule extension_lemma, erule f_mono)  | 
| 27404 | 360  | 
apply (rule ub_imageI, rename_tac a)  | 
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
30729 
diff
changeset
 | 
361  | 
apply (rule below_trans [OF below])  | 
| 
39974
 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 
huffman 
parents: 
39967 
diff
changeset
 | 
362  | 
apply (rule is_ub_thelub_ex)  | 
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
363  | 
apply (rule extension_lemma, erule g_mono)  | 
| 27404 | 364  | 
apply (erule imageI)  | 
365  | 
done  | 
|
366  | 
||
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
367  | 
lemma cont_extension:  | 
| 41182 | 368  | 
assumes f_mono: "\<And>a b x. a \<preceq> b \<Longrightarrow> f x a \<sqsubseteq> f x b"  | 
369  | 
assumes f_cont: "\<And>a. cont (\<lambda>x. f x a)"  | 
|
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
370  | 
shows "cont (\<lambda>x. extension (\<lambda>a. f x a))"  | 
| 41182 | 371  | 
apply (rule contI2)  | 
372  | 
apply (rule monofunI)  | 
|
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
373  | 
apply (rule extension_mono, erule f_mono, erule f_mono)  | 
| 41182 | 374  | 
apply (erule cont2monofunE [OF f_cont])  | 
375  | 
apply (rule cfun_belowI)  | 
|
376  | 
apply (rule principal_induct, simp)  | 
|
377  | 
apply (simp only: contlub_cfun_fun)  | 
|
| 
41394
 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 
huffman 
parents: 
41182 
diff
changeset
 | 
378  | 
apply (simp only: extension_principal f_mono)  | 
| 41182 | 379  | 
apply (simp add: cont2contlubE [OF f_cont])  | 
380  | 
done  | 
|
381  | 
||
| 27404 | 382  | 
end  | 
383  | 
||
| 39984 | 384  | 
lemma (in preorder) typedef_ideal_completion:  | 
385  | 
fixes Abs :: "'a set \<Rightarrow> 'b::cpo"  | 
|
386  | 
  assumes type: "type_definition Rep Abs {S. ideal S}"
 | 
|
387  | 
assumes below: "\<And>x y. x \<sqsubseteq> y \<longleftrightarrow> Rep x \<subseteq> Rep y"  | 
|
388  | 
  assumes principal: "\<And>a. principal a = Abs {b. b \<preceq> a}"
 | 
|
389  | 
assumes countable: "\<exists>f::'a \<Rightarrow> nat. inj f"  | 
|
390  | 
shows "ideal_completion r principal Rep"  | 
|
391  | 
proof  | 
|
392  | 
  interpret type_definition Rep Abs "{S. ideal S}" by fact
 | 
|
393  | 
fix a b :: 'a and x y :: 'b and Y :: "nat \<Rightarrow> 'b"  | 
|
394  | 
show "ideal (Rep x)"  | 
|
395  | 
using Rep [of x] by simp  | 
|
396  | 
show "chain Y \<Longrightarrow> Rep (\<Squnion>i. Y i) = (\<Union>i. Rep (Y i))"  | 
|
| 40769 | 397  | 
using type below by (rule typedef_ideal_rep_lub)  | 
| 39984 | 398  | 
  show "Rep (principal a) = {b. b \<preceq> a}"
 | 
399  | 
by (simp add: principal Abs_inverse ideal_principal)  | 
|
400  | 
show "Rep x \<subseteq> Rep y \<Longrightarrow> x \<sqsubseteq> y"  | 
|
401  | 
by (simp only: below)  | 
|
402  | 
show "\<exists>f::'a \<Rightarrow> nat. inj f"  | 
|
403  | 
by (rule countable)  | 
|
404  | 
qed  | 
|
405  | 
||
| 27404 | 406  | 
end  |