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