| author | blanchet | 
| Thu, 13 Mar 2014 13:18:13 +0100 | |
| changeset 56087 | 2cd8fcb4804d | 
| parent 48891 | c0eafbd55de3 | 
| child 58880 | 0baae4311a9f | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/Cpodef.thy | 
| 16697 | 2 | Author: Brian Huffman | 
| 3 | *) | |
| 4 | ||
| 5 | header {* Subtypes of pcpos *}
 | |
| 6 | ||
| 40772 | 7 | theory Cpodef | 
| 16697 | 8 | imports Adm | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
45606diff
changeset | 9 | keywords "pcpodef" "cpodef" :: thy_goal | 
| 16697 | 10 | begin | 
| 11 | ||
| 12 | subsection {* Proving a subtype is a partial order *}
 | |
| 13 | ||
| 14 | text {*
 | |
| 15 | A subtype of a partial order is itself a partial order, | |
| 16 | if the ordering is defined in the standard way. | |
| 17 | *} | |
| 18 | ||
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 19 | setup {* Sign.add_const_constraint (@{const_name Porder.below}, NONE) *}
 | 
| 28073 | 20 | |
| 16697 | 21 | theorem typedef_po: | 
| 28073 | 22 | fixes Abs :: "'a::po \<Rightarrow> 'b::type" | 
| 16697 | 23 | assumes type: "type_definition Rep Abs A" | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 24 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 25 |   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: 
29138diff
changeset | 26 | apply (intro_classes, unfold below) | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 27 | 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: 
29138diff
changeset | 28 | apply (erule (1) below_trans) | 
| 26420 
57a626f64875
make preorder locale into a superclass of class po
 huffman parents: 
26027diff
changeset | 29 | apply (rule type_definition.Rep_inject [OF type, THEN iffD1]) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 30 | apply (erule (1) below_antisym) | 
| 16697 | 31 | done | 
| 32 | ||
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 33 | setup {* Sign.add_const_constraint (@{const_name Porder.below},
 | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 34 |   SOME @{typ "'a::below \<Rightarrow> 'a::below \<Rightarrow> bool"}) *}
 | 
| 28073 | 35 | |
| 25827 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 36 | subsection {* Proving a subtype is finite *}
 | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 37 | |
| 27296 
eec7a1889ca5
moved Abs_image to Typedef.thy; prove finite_UNIV outside the locale
 huffman parents: 
26420diff
changeset | 38 | lemma typedef_finite_UNIV: | 
| 
eec7a1889ca5
moved Abs_image to Typedef.thy; prove finite_UNIV outside the locale
 huffman parents: 
26420diff
changeset | 39 | fixes Abs :: "'a::type \<Rightarrow> 'b::type" | 
| 
eec7a1889ca5
moved Abs_image to Typedef.thy; prove finite_UNIV outside the locale
 huffman parents: 
26420diff
changeset | 40 | assumes type: "type_definition Rep Abs A" | 
| 
eec7a1889ca5
moved Abs_image to Typedef.thy; prove finite_UNIV outside the locale
 huffman parents: 
26420diff
changeset | 41 | shows "finite A \<Longrightarrow> finite (UNIV :: 'b set)" | 
| 25827 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 42 | proof - | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 43 | assume "finite A" | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 44 | hence "finite (Abs ` A)" by (rule finite_imageI) | 
| 27296 
eec7a1889ca5
moved Abs_image to Typedef.thy; prove finite_UNIV outside the locale
 huffman parents: 
26420diff
changeset | 45 | thus "finite (UNIV :: 'b set)" | 
| 
eec7a1889ca5
moved Abs_image to Typedef.thy; prove finite_UNIV outside the locale
 huffman parents: 
26420diff
changeset | 46 | by (simp only: type_definition.Abs_image [OF type]) | 
| 25827 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 47 | qed | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 48 | |
| 17812 | 49 | subsection {* Proving a subtype is chain-finite *}
 | 
| 50 | ||
| 40035 | 51 | lemma ch2ch_Rep: | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 52 | assumes below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 40035 | 53 | shows "chain S \<Longrightarrow> chain (\<lambda>i. Rep (S i))" | 
| 54 | unfolding chain_def below . | |
| 17812 | 55 | |
| 56 | theorem typedef_chfin: | |
| 57 | fixes Abs :: "'a::chfin \<Rightarrow> 'b::po" | |
| 58 | assumes type: "type_definition Rep Abs A" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 59 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 17812 | 60 |   shows "OFCLASS('b, chfin_class)"
 | 
| 25921 | 61 | apply intro_classes | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 62 | apply (drule ch2ch_Rep [OF below]) | 
| 25921 | 63 | apply (drule chfin) | 
| 17812 | 64 | apply (unfold max_in_chain_def) | 
| 65 | apply (simp add: type_definition.Rep_inject [OF type]) | |
| 66 | done | |
| 67 | ||
| 16697 | 68 | subsection {* Proving a subtype is complete *}
 | 
| 69 | ||
| 70 | text {*
 | |
| 71 | A subtype of a cpo is itself a cpo if the ordering is | |
| 72 | defined in the standard way, and the defining subset | |
| 73 | is closed with respect to limits of chains. A set is | |
| 74 | closed if and only if membership in the set is an | |
| 75 | admissible predicate. | |
| 76 | *} | |
| 77 | ||
| 40035 | 78 | lemma typedef_is_lubI: | 
| 79 | assumes below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | |
| 80 | shows "range (\<lambda>i. Rep (S i)) <<| Rep x \<Longrightarrow> range S <<| x" | |
| 81 | unfolding is_lub_def is_ub_def below by simp | |
| 82 | ||
| 16918 | 83 | lemma Abs_inverse_lub_Rep: | 
| 16697 | 84 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::po" | 
| 85 | assumes type: "type_definition Rep Abs A" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 86 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 87 | and adm: "adm (\<lambda>x. x \<in> A)" | 
| 16918 | 88 | shows "chain S \<Longrightarrow> Rep (Abs (\<Squnion>i. Rep (S i))) = (\<Squnion>i. Rep (S i))" | 
| 89 | apply (rule type_definition.Abs_inverse [OF type]) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 90 | apply (erule admD [OF adm ch2ch_Rep [OF below]]) | 
| 16697 | 91 | apply (rule type_definition.Rep [OF type]) | 
| 92 | done | |
| 93 | ||
| 40770 
6023808b38d4
rename cpodef theorems: lub_foo -> is_lub_foo, thelub_foo -> lub_foo
 huffman parents: 
40325diff
changeset | 94 | theorem typedef_is_lub: | 
| 16697 | 95 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::po" | 
| 96 | assumes type: "type_definition Rep Abs A" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 97 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 98 | and adm: "adm (\<lambda>x. x \<in> A)" | 
| 16918 | 99 | shows "chain S \<Longrightarrow> range S <<| Abs (\<Squnion>i. Rep (S i))" | 
| 40035 | 100 | proof - | 
| 101 | assume S: "chain S" | |
| 102 | hence "chain (\<lambda>i. Rep (S i))" by (rule ch2ch_Rep [OF below]) | |
| 103 | hence "range (\<lambda>i. Rep (S i)) <<| (\<Squnion>i. Rep (S i))" by (rule cpo_lubI) | |
| 104 | hence "range (\<lambda>i. Rep (S i)) <<| Rep (Abs (\<Squnion>i. Rep (S i)))" | |
| 105 | by (simp only: Abs_inverse_lub_Rep [OF type below adm S]) | |
| 106 | thus "range S <<| Abs (\<Squnion>i. Rep (S i))" | |
| 107 | by (rule typedef_is_lubI [OF below]) | |
| 108 | qed | |
| 16697 | 109 | |
| 45606 | 110 | lemmas typedef_lub = typedef_is_lub [THEN lub_eqI] | 
| 16918 | 111 | |
| 16697 | 112 | theorem typedef_cpo: | 
| 113 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::po" | |
| 114 | assumes type: "type_definition Rep Abs A" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 115 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 116 | and adm: "adm (\<lambda>x. x \<in> A)" | 
| 117 |   shows "OFCLASS('b, cpo_class)"
 | |
| 16918 | 118 | proof | 
| 119 | fix S::"nat \<Rightarrow> 'b" assume "chain S" | |
| 120 | hence "range S <<| Abs (\<Squnion>i. Rep (S i))" | |
| 40770 
6023808b38d4
rename cpodef theorems: lub_foo -> is_lub_foo, thelub_foo -> lub_foo
 huffman parents: 
40325diff
changeset | 121 | by (rule typedef_is_lub [OF type below adm]) | 
| 16918 | 122 | thus "\<exists>x. range S <<| x" .. | 
| 123 | qed | |
| 16697 | 124 | |
| 35900 
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
 huffman parents: 
31738diff
changeset | 125 | subsubsection {* Continuity of \emph{Rep} and \emph{Abs} *}
 | 
| 16697 | 126 | |
| 127 | text {* For any sub-cpo, the @{term Rep} function is continuous. *}
 | |
| 128 | ||
| 129 | theorem typedef_cont_Rep: | |
| 130 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo" | |
| 131 | assumes type: "type_definition Rep Abs A" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 132 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 133 | and adm: "adm (\<lambda>x. x \<in> A)" | 
| 40834 
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
 huffman parents: 
40774diff
changeset | 134 | shows "cont (\<lambda>x. f x) \<Longrightarrow> cont (\<lambda>x. Rep (f x))" | 
| 
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
 huffman parents: 
40774diff
changeset | 135 | apply (erule cont_apply [OF _ _ cont_const]) | 
| 16697 | 136 | apply (rule contI) | 
| 40770 
6023808b38d4
rename cpodef theorems: lub_foo -> is_lub_foo, thelub_foo -> lub_foo
 huffman parents: 
40325diff
changeset | 137 | apply (simp only: typedef_lub [OF type below adm]) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 138 | apply (simp only: Abs_inverse_lub_Rep [OF type below adm]) | 
| 26027 | 139 | apply (rule cpo_lubI) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 140 | apply (erule ch2ch_Rep [OF below]) | 
| 16697 | 141 | done | 
| 142 | ||
| 143 | text {*
 | |
| 144 |   For a sub-cpo, we can make the @{term Abs} function continuous
 | |
| 145 | only if we restrict its domain to the defining subset by | |
| 146 | composing it with another continuous function. | |
| 147 | *} | |
| 148 | ||
| 149 | theorem typedef_cont_Abs: | |
| 150 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo" | |
| 151 | fixes f :: "'c::cpo \<Rightarrow> 'a::cpo" | |
| 152 | assumes type: "type_definition Rep Abs A" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 153 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16918 | 154 | and adm: "adm (\<lambda>x. x \<in> A)" (* not used *) | 
| 16697 | 155 | and f_in_A: "\<And>x. f x \<in> A" | 
| 40325 | 156 | shows "cont f \<Longrightarrow> cont (\<lambda>x. Abs (f x))" | 
| 157 | unfolding cont_def is_lub_def is_ub_def ball_simps below | |
| 158 | by (simp add: type_definition.Abs_inverse [OF type f_in_A]) | |
| 16697 | 159 | |
| 17833 | 160 | subsection {* Proving subtype elements are compact *}
 | 
| 161 | ||
| 162 | theorem typedef_compact: | |
| 163 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo" | |
| 164 | assumes type: "type_definition Rep Abs A" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 165 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 17833 | 166 | and adm: "adm (\<lambda>x. x \<in> A)" | 
| 167 | shows "compact (Rep k) \<Longrightarrow> compact k" | |
| 168 | proof (unfold compact_def) | |
| 169 | have cont_Rep: "cont Rep" | |
| 40834 
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
 huffman parents: 
40774diff
changeset | 170 | by (rule typedef_cont_Rep [OF type below adm cont_id]) | 
| 41182 | 171 | assume "adm (\<lambda>x. Rep k \<notsqsubseteq> x)" | 
| 172 | with cont_Rep have "adm (\<lambda>x. Rep k \<notsqsubseteq> Rep x)" by (rule adm_subst) | |
| 173 | thus "adm (\<lambda>x. k \<notsqsubseteq> x)" by (unfold below) | |
| 17833 | 174 | qed | 
| 175 | ||
| 16697 | 176 | subsection {* Proving a subtype is pointed *}
 | 
| 177 | ||
| 178 | text {*
 | |
| 179 | A subtype of a cpo has a least element if and only if | |
| 180 | the defining subset has a least element. | |
| 181 | *} | |
| 182 | ||
| 16918 | 183 | theorem typedef_pcpo_generic: | 
| 16697 | 184 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo" | 
| 185 | assumes type: "type_definition Rep Abs A" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 186 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 187 | and z_in_A: "z \<in> A" | 
| 188 | and z_least: "\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x" | |
| 189 |   shows "OFCLASS('b, pcpo_class)"
 | |
| 190 | apply (intro_classes) | |
| 191 | apply (rule_tac x="Abs z" in exI, rule allI) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 192 | apply (unfold below) | 
| 16697 | 193 | apply (subst type_definition.Abs_inverse [OF type z_in_A]) | 
| 194 | apply (rule z_least [OF type_definition.Rep [OF type]]) | |
| 195 | done | |
| 196 | ||
| 197 | text {*
 | |
| 198 | As a special case, a subtype of a pcpo has a least element | |
| 199 |   if the defining subset contains @{term \<bottom>}.
 | |
| 200 | *} | |
| 201 | ||
| 16918 | 202 | theorem typedef_pcpo: | 
| 16697 | 203 | fixes Abs :: "'a::pcpo \<Rightarrow> 'b::cpo" | 
| 204 | assumes type: "type_definition Rep Abs A" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 205 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 206 | and bottom_in_A: "\<bottom> \<in> A" | 
| 16697 | 207 |   shows "OFCLASS('b, pcpo_class)"
 | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 208 | by (rule typedef_pcpo_generic [OF type below bottom_in_A], rule minimal) | 
| 16697 | 209 | |
| 35900 
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
 huffman parents: 
31738diff
changeset | 210 | subsubsection {* Strictness of \emph{Rep} and \emph{Abs} *}
 | 
| 16697 | 211 | |
| 212 | text {*
 | |
| 213 |   For a sub-pcpo where @{term \<bottom>} is a member of the defining
 | |
| 214 |   subset, @{term Rep} and @{term Abs} are both strict.
 | |
| 215 | *} | |
| 216 | ||
| 217 | theorem typedef_Abs_strict: | |
| 218 | assumes type: "type_definition Rep Abs A" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 219 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 220 | and bottom_in_A: "\<bottom> \<in> A" | 
| 16697 | 221 | shows "Abs \<bottom> = \<bottom>" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 222 | apply (rule bottomI, unfold below) | 
| 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 223 | apply (simp add: type_definition.Abs_inverse [OF type bottom_in_A]) | 
| 16697 | 224 | done | 
| 225 | ||
| 226 | theorem typedef_Rep_strict: | |
| 227 | assumes type: "type_definition Rep Abs A" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 228 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 229 | and bottom_in_A: "\<bottom> \<in> A" | 
| 16697 | 230 | shows "Rep \<bottom> = \<bottom>" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 231 | apply (rule typedef_Abs_strict [OF type below bottom_in_A, THEN subst]) | 
| 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 232 | apply (rule type_definition.Abs_inverse [OF type bottom_in_A]) | 
| 16697 | 233 | done | 
| 234 | ||
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40089diff
changeset | 235 | theorem typedef_Abs_bottom_iff: | 
| 25926 | 236 | assumes type: "type_definition Rep Abs A" | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 237 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 238 | and bottom_in_A: "\<bottom> \<in> A" | 
| 25926 | 239 | shows "x \<in> A \<Longrightarrow> (Abs x = \<bottom>) = (x = \<bottom>)" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 240 | apply (rule typedef_Abs_strict [OF type below bottom_in_A, THEN subst]) | 
| 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 241 | apply (simp add: type_definition.Abs_inject [OF type] bottom_in_A) | 
| 25926 | 242 | done | 
| 243 | ||
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40089diff
changeset | 244 | theorem typedef_Rep_bottom_iff: | 
| 25926 | 245 | assumes type: "type_definition Rep Abs A" | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 246 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 247 | and bottom_in_A: "\<bottom> \<in> A" | 
| 25926 | 248 | shows "(Rep x = \<bottom>) = (x = \<bottom>)" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 249 | apply (rule typedef_Rep_strict [OF type below bottom_in_A, THEN subst]) | 
| 25926 | 250 | apply (simp add: type_definition.Rep_inject [OF type]) | 
| 251 | done | |
| 252 | ||
| 19519 | 253 | subsection {* Proving a subtype is flat *}
 | 
| 254 | ||
| 255 | theorem typedef_flat: | |
| 256 | fixes Abs :: "'a::flat \<Rightarrow> 'b::pcpo" | |
| 257 | assumes type: "type_definition Rep Abs A" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 258 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 259 | and bottom_in_A: "\<bottom> \<in> A" | 
| 19519 | 260 |   shows "OFCLASS('b, flat_class)"
 | 
| 261 | apply (intro_classes) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 262 | apply (unfold below) | 
| 19519 | 263 | apply (simp add: type_definition.Rep_inject [OF type, symmetric]) | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41182diff
changeset | 264 | apply (simp add: typedef_Rep_strict [OF type below bottom_in_A]) | 
| 19519 | 265 | apply (simp add: ax_flat) | 
| 266 | done | |
| 267 | ||
| 16697 | 268 | subsection {* HOLCF type definition package *}
 | 
| 269 | ||
| 48891 | 270 | ML_file "Tools/cpodef.ML" | 
| 16697 | 271 | |
| 272 | end |