| author | berghofe | 
| Sun, 10 Jan 2010 18:09:11 +0100 | |
| changeset 34910 | b23bd3ee4813 | 
| parent 31738 | 7b9b9ba532ca | 
| child 35900 | aa5dfb03eb1e | 
| permissions | -rw-r--r-- | 
| 16697 | 1 | (* Title: HOLCF/Pcpodef.thy | 
| 2 | Author: Brian Huffman | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Subtypes of pcpos *}
 | |
| 6 | ||
| 7 | theory Pcpodef | |
| 8 | imports Adm | |
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31076diff
changeset | 9 | uses ("Tools/pcpodef.ML")
 | 
| 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 | |
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 49 | theorem typedef_finite_po: | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 50 | fixes Abs :: "'a::finite_po \<Rightarrow> 'b::po" | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 51 | assumes type: "type_definition Rep Abs A" | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 52 |   shows "OFCLASS('b, finite_po_class)"
 | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 53 | apply (intro_classes) | 
| 27296 
eec7a1889ca5
moved Abs_image to Typedef.thy; prove finite_UNIV outside the locale
 huffman parents: 
26420diff
changeset | 54 | apply (rule typedef_finite_UNIV [OF type]) | 
| 25827 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 55 | apply (rule finite) | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 56 | done | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
23152diff
changeset | 57 | |
| 17812 | 58 | subsection {* Proving a subtype is chain-finite *}
 | 
| 59 | ||
| 60 | lemma monofun_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 | 61 | assumes below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 17812 | 62 | shows "monofun 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 | 63 | by (rule monofunI, unfold below) | 
| 17812 | 64 | |
| 65 | lemmas ch2ch_Rep = ch2ch_monofun [OF monofun_Rep] | |
| 66 | lemmas ub2ub_Rep = ub2ub_monofun [OF monofun_Rep] | |
| 67 | ||
| 68 | theorem typedef_chfin: | |
| 69 | fixes Abs :: "'a::chfin \<Rightarrow> 'b::po" | |
| 70 | 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 | 71 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 17812 | 72 |   shows "OFCLASS('b, chfin_class)"
 | 
| 25921 | 73 | 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 | 74 | apply (drule ch2ch_Rep [OF below]) | 
| 25921 | 75 | apply (drule chfin) | 
| 17812 | 76 | apply (unfold max_in_chain_def) | 
| 77 | apply (simp add: type_definition.Rep_inject [OF type]) | |
| 78 | done | |
| 79 | ||
| 16697 | 80 | subsection {* Proving a subtype is complete *}
 | 
| 81 | ||
| 82 | text {*
 | |
| 83 | A subtype of a cpo is itself a cpo if the ordering is | |
| 84 | defined in the standard way, and the defining subset | |
| 85 | is closed with respect to limits of chains. A set is | |
| 86 | closed if and only if membership in the set is an | |
| 87 | admissible predicate. | |
| 88 | *} | |
| 89 | ||
| 16918 | 90 | lemma Abs_inverse_lub_Rep: | 
| 16697 | 91 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::po" | 
| 92 | 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 | 93 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 94 | and adm: "adm (\<lambda>x. x \<in> A)" | 
| 16918 | 95 | shows "chain S \<Longrightarrow> Rep (Abs (\<Squnion>i. Rep (S i))) = (\<Squnion>i. Rep (S i))" | 
| 96 | 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 | 97 | apply (erule admD [OF adm ch2ch_Rep [OF below]]) | 
| 16697 | 98 | apply (rule type_definition.Rep [OF type]) | 
| 99 | done | |
| 100 | ||
| 16918 | 101 | theorem typedef_lub: | 
| 16697 | 102 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::po" | 
| 103 | 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 | 104 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 105 | and adm: "adm (\<lambda>x. x \<in> A)" | 
| 16918 | 106 | shows "chain S \<Longrightarrow> range S <<| Abs (\<Squnion>i. Rep (S i))" | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 107 | apply (frule ch2ch_Rep [OF below]) | 
| 16697 | 108 | apply (rule is_lubI) | 
| 109 | apply (rule ub_rangeI) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 110 | apply (simp only: below Abs_inverse_lub_Rep [OF type below adm]) | 
| 16918 | 111 | apply (erule is_ub_thelub) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 112 | apply (simp only: below Abs_inverse_lub_Rep [OF type below adm]) | 
| 16918 | 113 | apply (erule is_lub_thelub) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 114 | apply (erule ub2ub_Rep [OF below]) | 
| 16697 | 115 | done | 
| 116 | ||
| 16918 | 117 | lemmas typedef_thelub = typedef_lub [THEN thelubI, standard] | 
| 118 | ||
| 16697 | 119 | theorem typedef_cpo: | 
| 120 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::po" | |
| 121 | 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 | 122 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 123 | and adm: "adm (\<lambda>x. x \<in> A)" | 
| 124 |   shows "OFCLASS('b, cpo_class)"
 | |
| 16918 | 125 | proof | 
| 126 | fix S::"nat \<Rightarrow> 'b" assume "chain S" | |
| 127 | hence "range S <<| Abs (\<Squnion>i. Rep (S i))" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 128 | by (rule typedef_lub [OF type below adm]) | 
| 16918 | 129 | thus "\<exists>x. range S <<| x" .. | 
| 130 | qed | |
| 16697 | 131 | |
| 132 | subsubsection {* Continuity of @{term Rep} and @{term Abs} *}
 | |
| 133 | ||
| 134 | text {* For any sub-cpo, the @{term Rep} function is continuous. *}
 | |
| 135 | ||
| 136 | theorem typedef_cont_Rep: | |
| 137 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo" | |
| 138 | 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 | 139 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 140 | and adm: "adm (\<lambda>x. x \<in> A)" | 
| 141 | shows "cont Rep" | |
| 142 | apply (rule contI) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 143 | apply (simp only: typedef_thelub [OF type below adm]) | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 144 | apply (simp only: Abs_inverse_lub_Rep [OF type below adm]) | 
| 26027 | 145 | 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 | 146 | apply (erule ch2ch_Rep [OF below]) | 
| 16697 | 147 | done | 
| 148 | ||
| 149 | text {*
 | |
| 150 |   For a sub-cpo, we can make the @{term Abs} function continuous
 | |
| 151 | only if we restrict its domain to the defining subset by | |
| 152 | composing it with another continuous function. | |
| 153 | *} | |
| 154 | ||
| 16918 | 155 | theorem typedef_is_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 | 156 | assumes below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16918 | 157 | shows "range (\<lambda>i. Rep (S i)) <<| Rep x \<Longrightarrow> range S <<| x" | 
| 158 | apply (rule is_lubI) | |
| 159 | apply (rule ub_rangeI) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 160 | apply (subst below) | 
| 16918 | 161 | apply (erule is_ub_lub) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 162 | apply (subst below) | 
| 16918 | 163 | apply (erule is_lub_lub) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 164 | apply (erule ub2ub_Rep [OF below]) | 
| 16918 | 165 | done | 
| 166 | ||
| 16697 | 167 | theorem typedef_cont_Abs: | 
| 168 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo" | |
| 169 | fixes f :: "'c::cpo \<Rightarrow> 'a::cpo" | |
| 170 | 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 | 171 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16918 | 172 | and adm: "adm (\<lambda>x. x \<in> A)" (* not used *) | 
| 16697 | 173 | and f_in_A: "\<And>x. f x \<in> A" | 
| 174 | and cont_f: "cont f" | |
| 175 | shows "cont (\<lambda>x. Abs (f x))" | |
| 176 | apply (rule contI) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 177 | apply (rule typedef_is_lubI [OF below]) | 
| 16918 | 178 | apply (simp only: type_definition.Abs_inverse [OF type f_in_A]) | 
| 179 | apply (erule cont_f [THEN contE]) | |
| 16697 | 180 | done | 
| 181 | ||
| 17833 | 182 | subsection {* Proving subtype elements are compact *}
 | 
| 183 | ||
| 184 | theorem typedef_compact: | |
| 185 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo" | |
| 186 | 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 | 187 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 17833 | 188 | and adm: "adm (\<lambda>x. x \<in> A)" | 
| 189 | shows "compact (Rep k) \<Longrightarrow> compact k" | |
| 190 | proof (unfold compact_def) | |
| 191 | have cont_Rep: "cont 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 | 192 | by (rule typedef_cont_Rep [OF type below adm]) | 
| 17833 | 193 | assume "adm (\<lambda>x. \<not> Rep k \<sqsubseteq> x)" | 
| 194 | with cont_Rep have "adm (\<lambda>x. \<not> Rep k \<sqsubseteq> Rep x)" by (rule adm_subst) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 195 | thus "adm (\<lambda>x. \<not> k \<sqsubseteq> x)" by (unfold below) | 
| 17833 | 196 | qed | 
| 197 | ||
| 16697 | 198 | subsection {* Proving a subtype is pointed *}
 | 
| 199 | ||
| 200 | text {*
 | |
| 201 | A subtype of a cpo has a least element if and only if | |
| 202 | the defining subset has a least element. | |
| 203 | *} | |
| 204 | ||
| 16918 | 205 | theorem typedef_pcpo_generic: | 
| 16697 | 206 | fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo" | 
| 207 | 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 | 208 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 209 | and z_in_A: "z \<in> A" | 
| 210 | and z_least: "\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x" | |
| 211 |   shows "OFCLASS('b, pcpo_class)"
 | |
| 212 | apply (intro_classes) | |
| 213 | 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 | 214 | apply (unfold below) | 
| 16697 | 215 | apply (subst type_definition.Abs_inverse [OF type z_in_A]) | 
| 216 | apply (rule z_least [OF type_definition.Rep [OF type]]) | |
| 217 | done | |
| 218 | ||
| 219 | text {*
 | |
| 220 | As a special case, a subtype of a pcpo has a least element | |
| 221 |   if the defining subset contains @{term \<bottom>}.
 | |
| 222 | *} | |
| 223 | ||
| 16918 | 224 | theorem typedef_pcpo: | 
| 16697 | 225 | fixes Abs :: "'a::pcpo \<Rightarrow> 'b::cpo" | 
| 226 | 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 | 227 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 228 | and UU_in_A: "\<bottom> \<in> A" | 
| 229 |   shows "OFCLASS('b, pcpo_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 | 230 | by (rule typedef_pcpo_generic [OF type below UU_in_A], rule minimal) | 
| 16697 | 231 | |
| 232 | subsubsection {* Strictness of @{term Rep} and @{term Abs} *}
 | |
| 233 | ||
| 234 | text {*
 | |
| 235 |   For a sub-pcpo where @{term \<bottom>} is a member of the defining
 | |
| 236 |   subset, @{term Rep} and @{term Abs} are both strict.
 | |
| 237 | *} | |
| 238 | ||
| 239 | theorem typedef_Abs_strict: | |
| 240 | 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 | 241 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 242 | and UU_in_A: "\<bottom> \<in> A" | 
| 243 | shows "Abs \<bottom> = \<bottom>" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 244 | apply (rule UU_I, unfold below) | 
| 16697 | 245 | apply (simp add: type_definition.Abs_inverse [OF type UU_in_A]) | 
| 246 | done | |
| 247 | ||
| 248 | theorem typedef_Rep_strict: | |
| 249 | 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 | 250 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 251 | and UU_in_A: "\<bottom> \<in> A" | 
| 252 | shows "Rep \<bottom> = \<bottom>" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 253 | apply (rule typedef_Abs_strict [OF type below UU_in_A, THEN subst]) | 
| 16697 | 254 | apply (rule type_definition.Abs_inverse [OF type UU_in_A]) | 
| 255 | done | |
| 256 | ||
| 25926 | 257 | theorem typedef_Abs_strict_iff: | 
| 258 | 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 | 259 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 25926 | 260 | and UU_in_A: "\<bottom> \<in> A" | 
| 261 | shows "x \<in> A \<Longrightarrow> (Abs x = \<bottom>) = (x = \<bottom>)" | |
| 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 (rule typedef_Abs_strict [OF type below UU_in_A, THEN subst]) | 
| 25926 | 263 | apply (simp add: type_definition.Abs_inject [OF type] UU_in_A) | 
| 264 | done | |
| 265 | ||
| 266 | theorem typedef_Rep_strict_iff: | |
| 267 | 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 | 268 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 25926 | 269 | and UU_in_A: "\<bottom> \<in> A" | 
| 270 | shows "(Rep x = \<bottom>) = (x = \<bottom>)" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 271 | apply (rule typedef_Rep_strict [OF type below UU_in_A, THEN subst]) | 
| 25926 | 272 | apply (simp add: type_definition.Rep_inject [OF type]) | 
| 273 | done | |
| 274 | ||
| 16697 | 275 | theorem typedef_Abs_defined: | 
| 276 | 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 | 277 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 278 | and UU_in_A: "\<bottom> \<in> A" | 
| 279 | shows "\<lbrakk>x \<noteq> \<bottom>; x \<in> A\<rbrakk> \<Longrightarrow> Abs x \<noteq> \<bottom>" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 280 | by (simp add: typedef_Abs_strict_iff [OF type below UU_in_A]) | 
| 16697 | 281 | |
| 282 | theorem typedef_Rep_defined: | |
| 283 | 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 | 284 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 16697 | 285 | and UU_in_A: "\<bottom> \<in> A" | 
| 286 | shows "x \<noteq> \<bottom> \<Longrightarrow> Rep x \<noteq> \<bottom>" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 287 | by (simp add: typedef_Rep_strict_iff [OF type below UU_in_A]) | 
| 16697 | 288 | |
| 19519 | 289 | subsection {* Proving a subtype is flat *}
 | 
| 290 | ||
| 291 | theorem typedef_flat: | |
| 292 | fixes Abs :: "'a::flat \<Rightarrow> 'b::pcpo" | |
| 293 | 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 | 294 | and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" | 
| 19519 | 295 | and UU_in_A: "\<bottom> \<in> A" | 
| 296 |   shows "OFCLASS('b, flat_class)"
 | |
| 297 | 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 | 298 | apply (unfold below) | 
| 19519 | 299 | apply (simp add: type_definition.Rep_inject [OF type, symmetric]) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 300 | apply (simp add: typedef_Rep_strict [OF type below UU_in_A]) | 
| 19519 | 301 | apply (simp add: ax_flat) | 
| 302 | done | |
| 303 | ||
| 16697 | 304 | subsection {* HOLCF type definition package *}
 | 
| 305 | ||
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31076diff
changeset | 306 | use "Tools/pcpodef.ML" | 
| 16697 | 307 | |
| 308 | end |