author | huffman |
Tue, 26 Oct 2010 14:19:59 -0700 | |
changeset 40216 | 366309dfaf60 |
parent 40089 | 8adc57fb8454 |
child 40327 | 1dfdbd66093a |
permissions | -rw-r--r-- |
15599
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
huffman
parents:
15593
diff
changeset
|
1 |
(* Title: HOLCF/Up.thy |
16070
4a83dd540b88
removed LICENCE note -- everything is subject to Isabelle licence as
wenzelm
parents:
15599
diff
changeset
|
2 |
Author: Franz Regensburger and Brian Huffman |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
3 |
*) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
4 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
5 |
header {* The type of lifted values *} |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
6 |
|
15577 | 7 |
theory Up |
39987
8c2f449af35a
move all bifinite class instances to Bifinite.thy
huffman
parents:
39986
diff
changeset
|
8 |
imports Deflation |
15577 | 9 |
begin |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
10 |
|
36452 | 11 |
default_sort cpo |
15599
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
huffman
parents:
15593
diff
changeset
|
12 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
13 |
subsection {* Definition of new type for lifting *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
14 |
|
16753 | 15 |
datatype 'a u = Ibottom | Iup 'a |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
16 |
|
35427 | 17 |
type_notation (xsymbols) |
18 |
u ("(_\<^sub>\<bottom>)" [1000] 999) |
|
18290 | 19 |
|
34941 | 20 |
primrec Ifup :: "('a \<rightarrow> 'b::pcpo) \<Rightarrow> 'a u \<Rightarrow> 'b" where |
21 |
"Ifup f Ibottom = \<bottom>" |
|
22 |
| "Ifup f (Iup x) = f\<cdot>x" |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
23 |
|
18290 | 24 |
subsection {* Ordering on lifted cpo *} |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
25 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
26 |
instantiation u :: (cpo) below |
25787 | 27 |
begin |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
28 |
|
25787 | 29 |
definition |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
30 |
below_up_def: |
16753 | 31 |
"(op \<sqsubseteq>) \<equiv> (\<lambda>x y. case x of Ibottom \<Rightarrow> True | Iup a \<Rightarrow> |
32 |
(case y of Ibottom \<Rightarrow> False | Iup b \<Rightarrow> a \<sqsubseteq> b))" |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
33 |
|
25787 | 34 |
instance .. |
35 |
end |
|
36 |
||
16753 | 37 |
lemma minimal_up [iff]: "Ibottom \<sqsubseteq> z" |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
38 |
by (simp add: below_up_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
39 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
40 |
lemma not_Iup_below [iff]: "\<not> Iup x \<sqsubseteq> Ibottom" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
41 |
by (simp add: below_up_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
42 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
43 |
lemma Iup_below [iff]: "(Iup x \<sqsubseteq> Iup y) = (x \<sqsubseteq> y)" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
44 |
by (simp add: below_up_def) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
45 |
|
18290 | 46 |
subsection {* Lifted cpo is a partial order *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
47 |
|
15599
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
huffman
parents:
15593
diff
changeset
|
48 |
instance u :: (cpo) po |
25787 | 49 |
proof |
50 |
fix x :: "'a u" |
|
51 |
show "x \<sqsubseteq> x" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
52 |
unfolding below_up_def by (simp split: u.split) |
25787 | 53 |
next |
54 |
fix x y :: "'a u" |
|
55 |
assume "x \<sqsubseteq> y" "y \<sqsubseteq> x" thus "x = y" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
56 |
unfolding below_up_def |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
57 |
by (auto split: u.split_asm intro: below_antisym) |
25787 | 58 |
next |
59 |
fix x y z :: "'a u" |
|
60 |
assume "x \<sqsubseteq> y" "y \<sqsubseteq> z" thus "x \<sqsubseteq> z" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
61 |
unfolding below_up_def |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
62 |
by (auto split: u.split_asm intro: below_trans) |
25787 | 63 |
qed |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
64 |
|
18290 | 65 |
subsection {* Lifted cpo is a cpo *} |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
66 |
|
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
67 |
lemma is_lub_Iup: |
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
68 |
"range S <<| x \<Longrightarrow> range (\<lambda>i. Iup (S i)) <<| Iup x" |
40084 | 69 |
unfolding is_lub_def is_ub_def ball_simps |
70 |
by (auto simp add: below_up_def split: u.split) |
|
15599
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
huffman
parents:
15593
diff
changeset
|
71 |
|
17838 | 72 |
lemma up_chain_lemma: |
40084 | 73 |
assumes Y: "chain Y" obtains "\<forall>i. Y i = Ibottom" |
74 |
| A k where "\<forall>i. Iup (A i) = Y (i + k)" and "chain A" and "range Y <<| Iup (\<Squnion>i. A i)" |
|
75 |
proof (cases "\<exists>k. Y k \<noteq> Ibottom") |
|
76 |
case True |
|
77 |
then obtain k where k: "Y k \<noteq> Ibottom" .. |
|
78 |
def A \<equiv> "\<lambda>i. THE a. Iup a = Y (i + k)" |
|
79 |
have Iup_A: "\<forall>i. Iup (A i) = Y (i + k)" |
|
80 |
proof |
|
81 |
fix i :: nat |
|
82 |
from Y le_add2 have "Y k \<sqsubseteq> Y (i + k)" by (rule chain_mono) |
|
83 |
with k have "Y (i + k) \<noteq> Ibottom" by (cases "Y k", auto) |
|
84 |
thus "Iup (A i) = Y (i + k)" |
|
85 |
by (cases "Y (i + k)", simp_all add: A_def) |
|
86 |
qed |
|
87 |
from Y have chain_A: "chain A" |
|
88 |
unfolding chain_def Iup_below [symmetric] |
|
89 |
by (simp add: Iup_A) |
|
90 |
hence "range A <<| (\<Squnion>i. A i)" |
|
91 |
by (rule cpo_lubI) |
|
92 |
hence "range (\<lambda>i. Iup (A i)) <<| Iup (\<Squnion>i. A i)" |
|
93 |
by (rule is_lub_Iup) |
|
94 |
hence "range (\<lambda>i. Y (i + k)) <<| Iup (\<Squnion>i. A i)" |
|
95 |
by (simp only: Iup_A) |
|
96 |
hence "range (\<lambda>i. Y i) <<| Iup (\<Squnion>i. A i)" |
|
97 |
by (simp only: is_lub_range_shift [OF Y]) |
|
98 |
with Iup_A chain_A show ?thesis .. |
|
99 |
next |
|
100 |
case False |
|
101 |
then have "\<forall>i. Y i = Ibottom" by simp |
|
102 |
then show ?thesis .. |
|
103 |
qed |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
104 |
|
15599
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
huffman
parents:
15593
diff
changeset
|
105 |
instance u :: (cpo) cpo |
40084 | 106 |
proof |
107 |
fix S :: "nat \<Rightarrow> 'a u" |
|
108 |
assume S: "chain S" |
|
109 |
thus "\<exists>x. range (\<lambda>i. S i) <<| x" |
|
110 |
proof (rule up_chain_lemma) |
|
111 |
assume "\<forall>i. S i = Ibottom" |
|
112 |
hence "range (\<lambda>i. S i) <<| Ibottom" |
|
113 |
by (simp add: lub_const) |
|
114 |
thus ?thesis .. |
|
115 |
next |
|
40085 | 116 |
fix A :: "nat \<Rightarrow> 'a" |
117 |
assume "range S <<| Iup (\<Squnion>i. A i)" |
|
40084 | 118 |
thus ?thesis .. |
119 |
qed |
|
120 |
qed |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
121 |
|
18290 | 122 |
subsection {* Lifted cpo is pointed *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
123 |
|
15599
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
huffman
parents:
15593
diff
changeset
|
124 |
instance u :: (cpo) pcpo |
40084 | 125 |
by intro_classes fast |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
126 |
|
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
127 |
text {* for compatibility with old HOLCF-Version *} |
16753 | 128 |
lemma inst_up_pcpo: "\<bottom> = Ibottom" |
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
129 |
by (rule minimal_up [THEN UU_I, symmetric]) |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
130 |
|
35900
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
huffman
parents:
35783
diff
changeset
|
131 |
subsection {* Continuity of \emph{Iup} and \emph{Ifup} *} |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
132 |
|
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
133 |
text {* continuity for @{term Iup} *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
134 |
|
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
135 |
lemma cont_Iup: "cont Iup" |
16215 | 136 |
apply (rule contI) |
15599
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
huffman
parents:
15593
diff
changeset
|
137 |
apply (rule is_lub_Iup) |
26027 | 138 |
apply (erule cpo_lubI) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
139 |
done |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
140 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
141 |
text {* continuity for @{term Ifup} *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
142 |
|
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
143 |
lemma cont_Ifup1: "cont (\<lambda>f. Ifup f x)" |
16753 | 144 |
by (induct x, simp_all) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
145 |
|
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
146 |
lemma monofun_Ifup2: "monofun (\<lambda>x. Ifup f x)" |
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
147 |
apply (rule monofunI) |
16753 | 148 |
apply (case_tac x, simp) |
149 |
apply (case_tac y, simp) |
|
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
150 |
apply (simp add: monofun_cfun_arg) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
151 |
done |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
152 |
|
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
153 |
lemma cont_Ifup2: "cont (\<lambda>x. Ifup f x)" |
40084 | 154 |
proof (rule contI2) |
155 |
fix Y assume Y: "chain Y" and Y': "chain (\<lambda>i. Ifup f (Y i))" |
|
156 |
from Y show "Ifup f (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. Ifup f (Y i))" |
|
157 |
proof (rule up_chain_lemma) |
|
158 |
fix A and k |
|
159 |
assume A: "\<forall>i. Iup (A i) = Y (i + k)" |
|
160 |
assume "chain A" and "range Y <<| Iup (\<Squnion>i. A i)" |
|
161 |
hence "Ifup f (\<Squnion>i. Y i) = (\<Squnion>i. Ifup f (Iup (A i)))" |
|
162 |
by (simp add: thelubI contlub_cfun_arg) |
|
163 |
also have "\<dots> = (\<Squnion>i. Ifup f (Y (i + k)))" |
|
164 |
by (simp add: A) |
|
165 |
also have "\<dots> = (\<Squnion>i. Ifup f (Y i))" |
|
166 |
using Y' by (rule lub_range_shift) |
|
167 |
finally show ?thesis by simp |
|
168 |
qed simp |
|
169 |
qed (rule monofun_Ifup2) |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
170 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
171 |
subsection {* Continuous versions of constants *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
172 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
19105
diff
changeset
|
173 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
19105
diff
changeset
|
174 |
up :: "'a \<rightarrow> 'a u" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
19105
diff
changeset
|
175 |
"up = (\<Lambda> x. Iup x)" |
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
176 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
19105
diff
changeset
|
177 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
19105
diff
changeset
|
178 |
fup :: "('a \<rightarrow> 'b::pcpo) \<rightarrow> 'a u \<rightarrow> 'b" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
19105
diff
changeset
|
179 |
"fup = (\<Lambda> f p. Ifup f p)" |
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
180 |
|
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
181 |
translations |
26046 | 182 |
"case l of XCONST up\<cdot>x \<Rightarrow> t" == "CONST fup\<cdot>(\<Lambda> x. t)\<cdot>l" |
183 |
"\<Lambda>(XCONST up\<cdot>x). t" == "CONST fup\<cdot>(\<Lambda> x. t)" |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
184 |
|
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
185 |
text {* continuous versions of lemmas for @{typ "('a)u"} *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
186 |
|
16753 | 187 |
lemma Exh_Up: "z = \<bottom> \<or> (\<exists>x. z = up\<cdot>x)" |
188 |
apply (induct z) |
|
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
189 |
apply (simp add: inst_up_pcpo) |
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
190 |
apply (simp add: up_def cont_Iup) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
191 |
done |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
192 |
|
16753 | 193 |
lemma up_eq [simp]: "(up\<cdot>x = up\<cdot>y) = (x = y)" |
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
194 |
by (simp add: up_def cont_Iup) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
195 |
|
16753 | 196 |
lemma up_inject: "up\<cdot>x = up\<cdot>y \<Longrightarrow> x = y" |
197 |
by simp |
|
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
198 |
|
17838 | 199 |
lemma up_defined [simp]: "up\<cdot>x \<noteq> \<bottom>" |
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
200 |
by (simp add: up_def cont_Iup inst_up_pcpo) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
201 |
|
25785 | 202 |
lemma not_up_less_UU: "\<not> up\<cdot>x \<sqsubseteq> \<bottom>" |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
203 |
by simp (* FIXME: remove? *) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
204 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
29530
diff
changeset
|
205 |
lemma up_below [simp]: "up\<cdot>x \<sqsubseteq> up\<cdot>y \<longleftrightarrow> x \<sqsubseteq> y" |
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
206 |
by (simp add: up_def cont_Iup) |
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
207 |
|
35783 | 208 |
lemma upE [case_names bottom up, cases type: u]: |
209 |
"\<lbrakk>p = \<bottom> \<Longrightarrow> Q; \<And>x. p = up\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" |
|
25788
30cbe0764599
declare upE as cases rule; add new rule up_induct
huffman
parents:
25787
diff
changeset
|
210 |
apply (cases p) |
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
211 |
apply (simp add: inst_up_pcpo) |
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
212 |
apply (simp add: up_def cont_Iup) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
213 |
done |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
214 |
|
35783 | 215 |
lemma up_induct [case_names bottom up, induct type: u]: |
216 |
"\<lbrakk>P \<bottom>; \<And>x. P (up\<cdot>x)\<rbrakk> \<Longrightarrow> P x" |
|
25788
30cbe0764599
declare upE as cases rule; add new rule up_induct
huffman
parents:
25787
diff
changeset
|
217 |
by (cases x, simp_all) |
30cbe0764599
declare upE as cases rule; add new rule up_induct
huffman
parents:
25787
diff
changeset
|
218 |
|
25827
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25789
diff
changeset
|
219 |
text {* lifting preserves chain-finiteness *} |
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25789
diff
changeset
|
220 |
|
17838 | 221 |
lemma up_chain_cases: |
40084 | 222 |
assumes Y: "chain Y" obtains "\<forall>i. Y i = \<bottom>" |
223 |
| A k where "\<forall>i. up\<cdot>(A i) = Y (i + k)" and "chain A" and "(\<Squnion>i. Y i) = up\<cdot>(\<Squnion>i. A i)" |
|
224 |
apply (rule up_chain_lemma [OF Y]) |
|
225 |
apply (simp_all add: inst_up_pcpo up_def cont_Iup thelubI) |
|
226 |
done |
|
17838 | 227 |
|
25879 | 228 |
lemma compact_up: "compact x \<Longrightarrow> compact (up\<cdot>x)" |
229 |
apply (rule compactI2) |
|
40084 | 230 |
apply (erule up_chain_cases) |
231 |
apply simp |
|
25879 | 232 |
apply (drule (1) compactD2, simp) |
40084 | 233 |
apply (erule exE) |
234 |
apply (drule_tac f="up" and x="x" in monofun_cfun_arg) |
|
235 |
apply (simp, erule exI) |
|
25879 | 236 |
done |
237 |
||
238 |
lemma compact_upD: "compact (up\<cdot>x) \<Longrightarrow> compact x" |
|
239 |
unfolding compact_def |
|
240 |
by (drule adm_subst [OF cont_Rep_CFun2 [where f=up]], simp) |
|
241 |
||
242 |
lemma compact_up_iff [simp]: "compact (up\<cdot>x) = compact x" |
|
243 |
by (safe elim!: compact_up compact_upD) |
|
244 |
||
25827
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
huffman
parents:
25789
diff
changeset
|
245 |
instance u :: (chfin) chfin |
25921 | 246 |
apply intro_classes |
25879 | 247 |
apply (erule compact_imp_max_in_chain) |
25898 | 248 |
apply (rule_tac p="\<Squnion>i. Y i" in upE, simp_all) |
17838 | 249 |
done |
250 |
||
251 |
text {* properties of fup *} |
|
252 |
||
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
253 |
lemma fup1 [simp]: "fup\<cdot>f\<cdot>\<bottom> = \<bottom>" |
29530
9905b660612b
change to simpler, more extensible continuity simproc
huffman
parents:
29138
diff
changeset
|
254 |
by (simp add: fup_def cont_Ifup1 cont_Ifup2 inst_up_pcpo cont2cont_LAM) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
255 |
|
16319
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents:
16215
diff
changeset
|
256 |
lemma fup2 [simp]: "fup\<cdot>f\<cdot>(up\<cdot>x) = f\<cdot>x" |
29530
9905b660612b
change to simpler, more extensible continuity simproc
huffman
parents:
29138
diff
changeset
|
257 |
by (simp add: up_def fup_def cont_Iup cont_Ifup1 cont_Ifup2 cont2cont_LAM) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
258 |
|
16553 | 259 |
lemma fup3 [simp]: "fup\<cdot>up\<cdot>x = x" |
25788
30cbe0764599
declare upE as cases rule; add new rule up_induct
huffman
parents:
25787
diff
changeset
|
260 |
by (cases x, simp_all) |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
261 |
|
33504
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
262 |
subsection {* Map function for lifted cpo *} |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
263 |
|
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
264 |
definition |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
265 |
u_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a u \<rightarrow> 'b u" |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
266 |
where |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
267 |
"u_map = (\<Lambda> f. fup\<cdot>(up oo f))" |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
268 |
|
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
269 |
lemma u_map_strict [simp]: "u_map\<cdot>f\<cdot>\<bottom> = \<bottom>" |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
270 |
unfolding u_map_def by simp |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
271 |
|
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
272 |
lemma u_map_up [simp]: "u_map\<cdot>f\<cdot>(up\<cdot>x) = up\<cdot>(f\<cdot>x)" |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
273 |
unfolding u_map_def by simp |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
274 |
|
33808 | 275 |
lemma u_map_ID: "u_map\<cdot>ID = ID" |
40002
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents:
39987
diff
changeset
|
276 |
unfolding u_map_def by (simp add: cfun_eq_iff eta_cfun) |
33808 | 277 |
|
33587 | 278 |
lemma u_map_map: "u_map\<cdot>f\<cdot>(u_map\<cdot>g\<cdot>p) = u_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>p" |
279 |
by (induct p) simp_all |
|
280 |
||
33504
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
281 |
lemma ep_pair_u_map: "ep_pair e p \<Longrightarrow> ep_pair (u_map\<cdot>e) (u_map\<cdot>p)" |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
282 |
apply default |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
283 |
apply (case_tac x, simp, simp add: ep_pair.e_inverse) |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
284 |
apply (case_tac y, simp, simp add: ep_pair.e_p_below) |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
285 |
done |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
286 |
|
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
287 |
lemma deflation_u_map: "deflation d \<Longrightarrow> deflation (u_map\<cdot>d)" |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
288 |
apply default |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
289 |
apply (case_tac x, simp, simp add: deflation.idem) |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
290 |
apply (case_tac x, simp, simp add: deflation.below) |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
291 |
done |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
292 |
|
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
293 |
lemma finite_deflation_u_map: |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
294 |
assumes "finite_deflation d" shows "finite_deflation (u_map\<cdot>d)" |
39973
c62b4ff97bfc
add lemma finite_deflation_intro
Brian Huffman <brianh@cs.pdx.edu>
parents:
39302
diff
changeset
|
295 |
proof (rule finite_deflation_intro) |
33504
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
296 |
interpret d: finite_deflation d by fact |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
297 |
have "deflation d" by fact |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
298 |
thus "deflation (u_map\<cdot>d)" by (rule deflation_u_map) |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
299 |
have "{x. u_map\<cdot>d\<cdot>x = x} \<subseteq> insert \<bottom> ((\<lambda>x. up\<cdot>x) ` {x. d\<cdot>x = x})" |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
300 |
by (rule subsetI, case_tac x, simp_all) |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
301 |
thus "finite {x. u_map\<cdot>d\<cdot>x = x}" |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
302 |
by (rule finite_subset, simp add: d.finite_fixes) |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
303 |
qed |
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents:
31076
diff
changeset
|
304 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
305 |
end |