| author | desharna | 
| Mon, 19 Dec 2022 08:01:31 +0100 | |
| changeset 76684 | 3eda063a20a4 | 
| parent 69913 | ca515cf61651 | 
| child 81573 | 972fecd8907a | 
| permissions | -rw-r--r-- | 
| 42151 | 1  | 
(* Title: HOL/HOLCF/Domain.thy  | 
| 15741 | 2  | 
Author: Brian Huffman  | 
3  | 
*)  | 
|
4  | 
||
| 62175 | 5  | 
section \<open>Domain package\<close>  | 
| 15741 | 6  | 
|
7  | 
theory Domain  | 
|
| 41285 | 8  | 
imports Representable Domain_Aux  | 
| 
46950
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46947 
diff
changeset
 | 
9  | 
keywords  | 
| 63432 | 10  | 
"lazy" "unsafe" and  | 
| 69913 | 11  | 
"domaindef" "domain" :: thy_defn and  | 
12  | 
"domain_isomorphism" :: thy_decl  | 
|
| 15741 | 13  | 
begin  | 
14  | 
||
| 40504 | 15  | 
default_sort "domain"  | 
16  | 
||
| 62175 | 17  | 
subsection \<open>Representations of types\<close>  | 
| 40504 | 18  | 
|
19  | 
lemma emb_prj: "emb\<cdot>((prj\<cdot>x)::'a) = cast\<cdot>DEFL('a)\<cdot>x"
 | 
|
20  | 
by (simp add: cast_DEFL)  | 
|
21  | 
||
22  | 
lemma emb_prj_emb:  | 
|
23  | 
fixes x :: "'a"  | 
|
24  | 
  assumes "DEFL('a) \<sqsubseteq> DEFL('b)"
 | 
|
25  | 
shows "emb\<cdot>(prj\<cdot>(emb\<cdot>x) :: 'b) = emb\<cdot>x"  | 
|
26  | 
unfolding emb_prj  | 
|
27  | 
apply (rule cast.belowD)  | 
|
28  | 
apply (rule monofun_cfun_arg [OF assms])  | 
|
29  | 
apply (simp add: cast_DEFL)  | 
|
30  | 
done  | 
|
31  | 
||
32  | 
lemma prj_emb_prj:  | 
|
33  | 
  assumes "DEFL('a) \<sqsubseteq> DEFL('b)"
 | 
|
34  | 
shows "prj\<cdot>(emb\<cdot>(prj\<cdot>x :: 'b)) = (prj\<cdot>x :: 'a)"  | 
|
35  | 
apply (rule emb_eq_iff [THEN iffD1])  | 
|
36  | 
apply (simp only: emb_prj)  | 
|
37  | 
apply (rule deflation_below_comp1)  | 
|
38  | 
apply (rule deflation_cast)  | 
|
39  | 
apply (rule deflation_cast)  | 
|
40  | 
apply (rule monofun_cfun_arg [OF assms])  | 
|
41  | 
done  | 
|
42  | 
||
| 62175 | 43  | 
text \<open>Isomorphism lemmas used internally by the domain package:\<close>  | 
| 40504 | 44  | 
|
45  | 
lemma domain_abs_iso:  | 
|
46  | 
fixes abs and rep  | 
|
47  | 
  assumes DEFL: "DEFL('b) = DEFL('a)"
 | 
|
48  | 
assumes abs_def: "(abs :: 'a \<rightarrow> 'b) \<equiv> prj oo emb"  | 
|
49  | 
assumes rep_def: "(rep :: 'b \<rightarrow> 'a) \<equiv> prj oo emb"  | 
|
50  | 
shows "rep\<cdot>(abs\<cdot>x) = x"  | 
|
51  | 
unfolding abs_def rep_def  | 
|
52  | 
by (simp add: emb_prj_emb DEFL)  | 
|
53  | 
||
54  | 
lemma domain_rep_iso:  | 
|
55  | 
fixes abs and rep  | 
|
56  | 
  assumes DEFL: "DEFL('b) = DEFL('a)"
 | 
|
57  | 
assumes abs_def: "(abs :: 'a \<rightarrow> 'b) \<equiv> prj oo emb"  | 
|
58  | 
assumes rep_def: "(rep :: 'b \<rightarrow> 'a) \<equiv> prj oo emb"  | 
|
59  | 
shows "abs\<cdot>(rep\<cdot>x) = x"  | 
|
60  | 
unfolding abs_def rep_def  | 
|
61  | 
by (simp add: emb_prj_emb DEFL)  | 
|
62  | 
||
| 62175 | 63  | 
subsection \<open>Deflations as sets\<close>  | 
| 40504 | 64  | 
|
| 
41287
 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 
huffman 
parents: 
41285 
diff
changeset
 | 
65  | 
definition defl_set :: "'a::bifinite defl \<Rightarrow> 'a set"  | 
| 40504 | 66  | 
where "defl_set A = {x. cast\<cdot>A\<cdot>x = x}"
 | 
67  | 
||
68  | 
lemma adm_defl_set: "adm (\<lambda>x. x \<in> defl_set A)"  | 
|
69  | 
unfolding defl_set_def by simp  | 
|
70  | 
||
71  | 
lemma defl_set_bottom: "\<bottom> \<in> defl_set A"  | 
|
72  | 
unfolding defl_set_def by simp  | 
|
73  | 
||
74  | 
lemma defl_set_cast [simp]: "cast\<cdot>A\<cdot>x \<in> defl_set A"  | 
|
75  | 
unfolding defl_set_def by simp  | 
|
76  | 
||
77  | 
lemma defl_set_subset_iff: "defl_set A \<subseteq> defl_set B \<longleftrightarrow> A \<sqsubseteq> B"  | 
|
78  | 
apply (simp add: defl_set_def subset_eq cast_below_cast [symmetric])  | 
|
79  | 
apply (auto simp add: cast.belowI cast.belowD)  | 
|
80  | 
done  | 
|
81  | 
||
| 62175 | 82  | 
subsection \<open>Proving a subtype is representable\<close>  | 
| 40504 | 83  | 
|
| 62175 | 84  | 
text \<open>Temporarily relax type constraints.\<close>  | 
| 40504 | 85  | 
|
| 62175 | 86  | 
setup \<open>  | 
| 40504 | 87  | 
fold Sign.add_const_constraint  | 
| 69597 | 88  | 
[ (\<^const_name>\<open>defl\<close>, SOME \<^typ>\<open>'a::pcpo itself \<Rightarrow> udom defl\<close>)  | 
89  | 
, (\<^const_name>\<open>emb\<close>, SOME \<^typ>\<open>'a::pcpo \<rightarrow> udom\<close>)  | 
|
90  | 
, (\<^const_name>\<open>prj\<close>, SOME \<^typ>\<open>udom \<rightarrow> 'a::pcpo\<close>)  | 
|
91  | 
, (\<^const_name>\<open>liftdefl\<close>, SOME \<^typ>\<open>'a::pcpo itself \<Rightarrow> udom u defl\<close>)  | 
|
92  | 
, (\<^const_name>\<open>liftemb\<close>, SOME \<^typ>\<open>'a::pcpo u \<rightarrow> udom u\<close>)  | 
|
93  | 
, (\<^const_name>\<open>liftprj\<close>, SOME \<^typ>\<open>udom u \<rightarrow> 'a::pcpo u\<close>) ]  | 
|
| 62175 | 94  | 
\<close>  | 
| 40504 | 95  | 
|
| 
41292
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
96  | 
lemma typedef_domain_class:  | 
| 40504 | 97  | 
fixes Rep :: "'a::pcpo \<Rightarrow> udom"  | 
98  | 
fixes Abs :: "udom \<Rightarrow> 'a::pcpo"  | 
|
| 
41287
 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 
huffman 
parents: 
41285 
diff
changeset
 | 
99  | 
fixes t :: "udom defl"  | 
| 40504 | 100  | 
assumes type: "type_definition Rep Abs (defl_set t)"  | 
| 67399 | 101  | 
assumes below: "(\<sqsubseteq>) \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"  | 
| 40504 | 102  | 
assumes emb: "emb \<equiv> (\<Lambda> x. Rep x)"  | 
103  | 
assumes prj: "prj \<equiv> (\<Lambda> x. Abs (cast\<cdot>t\<cdot>x))"  | 
|
104  | 
assumes defl: "defl \<equiv> (\<lambda> a::'a itself. t)"  | 
|
| 
41292
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
105  | 
assumes liftemb: "(liftemb :: 'a u \<rightarrow> udom u) \<equiv> u_map\<cdot>emb"  | 
| 
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
106  | 
assumes liftprj: "(liftprj :: udom u \<rightarrow> 'a u) \<equiv> u_map\<cdot>prj"  | 
| 41436 | 107  | 
  assumes liftdefl: "(liftdefl :: 'a itself \<Rightarrow> _) \<equiv> (\<lambda>t. liftdefl_of\<cdot>DEFL('a))"
 | 
| 
41292
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
108  | 
  shows "OFCLASS('a, domain_class)"
 | 
| 
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
109  | 
proof  | 
| 40504 | 110  | 
have emb_beta: "\<And>x. emb\<cdot>x = Rep x"  | 
111  | 
unfolding emb  | 
|
112  | 
apply (rule beta_cfun)  | 
|
| 
40834
 
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
 
huffman 
parents: 
40830 
diff
changeset
 | 
113  | 
apply (rule typedef_cont_Rep [OF type below adm_defl_set cont_id])  | 
| 40504 | 114  | 
done  | 
115  | 
have prj_beta: "\<And>y. prj\<cdot>y = Abs (cast\<cdot>t\<cdot>y)"  | 
|
116  | 
unfolding prj  | 
|
117  | 
apply (rule beta_cfun)  | 
|
118  | 
apply (rule typedef_cont_Abs [OF type below adm_defl_set])  | 
|
119  | 
apply simp_all  | 
|
120  | 
done  | 
|
121  | 
have prj_emb: "\<And>x::'a. prj\<cdot>(emb\<cdot>x) = x"  | 
|
122  | 
using type_definition.Rep [OF type]  | 
|
123  | 
unfolding prj_beta emb_beta defl_set_def  | 
|
124  | 
by (simp add: type_definition.Rep_inverse [OF type])  | 
|
125  | 
have emb_prj: "\<And>y. emb\<cdot>(prj\<cdot>y :: 'a) = cast\<cdot>t\<cdot>y"  | 
|
126  | 
unfolding prj_beta emb_beta  | 
|
127  | 
by (simp add: type_definition.Abs_inverse [OF type])  | 
|
128  | 
show "ep_pair (emb :: 'a \<rightarrow> udom) prj"  | 
|
| 61169 | 129  | 
apply standard  | 
| 40504 | 130  | 
apply (simp add: prj_emb)  | 
131  | 
apply (simp add: emb_prj cast.below)  | 
|
132  | 
done  | 
|
133  | 
  show "cast\<cdot>DEFL('a) = emb oo (prj :: udom \<rightarrow> 'a)"
 | 
|
134  | 
by (rule cfun_eqI, simp add: defl emb_prj)  | 
|
| 
41292
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
135  | 
qed (simp_all only: liftemb liftprj liftdefl)  | 
| 40504 | 136  | 
|
137  | 
lemma typedef_DEFL:  | 
|
138  | 
assumes "defl \<equiv> (\<lambda>a::'a::pcpo itself. t)"  | 
|
139  | 
  shows "DEFL('a::pcpo) = t"
 | 
|
140  | 
unfolding assms ..  | 
|
141  | 
||
| 62175 | 142  | 
text \<open>Restore original typing constraints.\<close>  | 
| 40504 | 143  | 
|
| 62175 | 144  | 
setup \<open>  | 
| 40504 | 145  | 
fold Sign.add_const_constraint  | 
| 69597 | 146  | 
[(\<^const_name>\<open>defl\<close>, SOME \<^typ>\<open>'a::domain itself \<Rightarrow> udom defl\<close>),  | 
147  | 
(\<^const_name>\<open>emb\<close>, SOME \<^typ>\<open>'a::domain \<rightarrow> udom\<close>),  | 
|
148  | 
(\<^const_name>\<open>prj\<close>, SOME \<^typ>\<open>udom \<rightarrow> 'a::domain\<close>),  | 
|
149  | 
(\<^const_name>\<open>liftdefl\<close>, SOME \<^typ>\<open>'a::predomain itself \<Rightarrow> udom u defl\<close>),  | 
|
150  | 
(\<^const_name>\<open>liftemb\<close>, SOME \<^typ>\<open>'a::predomain u \<rightarrow> udom u\<close>),  | 
|
151  | 
(\<^const_name>\<open>liftprj\<close>, SOME \<^typ>\<open>udom u \<rightarrow> 'a::predomain u\<close>)]  | 
|
| 62175 | 152  | 
\<close>  | 
| 40504 | 153  | 
|
| 69605 | 154  | 
ML_file \<open>Tools/domaindef.ML\<close>  | 
| 40504 | 155  | 
|
| 62175 | 156  | 
subsection \<open>Isomorphic deflations\<close>  | 
| 40504 | 157  | 
|
| 
41292
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
158  | 
definition isodefl :: "('a \<rightarrow> 'a) \<Rightarrow> udom defl \<Rightarrow> bool"
 | 
| 
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
159  | 
where "isodefl d t \<longleftrightarrow> cast\<cdot>t = emb oo d oo prj"  | 
| 
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
160  | 
|
| 
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
161  | 
definition isodefl' :: "('a::predomain \<rightarrow> 'a) \<Rightarrow> udom u defl \<Rightarrow> bool"
 | 
| 
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
162  | 
where "isodefl' d t \<longleftrightarrow> cast\<cdot>t = liftemb oo u_map\<cdot>d oo liftprj"  | 
| 40504 | 163  | 
|
164  | 
lemma isodeflI: "(\<And>x. cast\<cdot>t\<cdot>x = emb\<cdot>(d\<cdot>(prj\<cdot>x))) \<Longrightarrow> isodefl d t"  | 
|
165  | 
unfolding isodefl_def by (simp add: cfun_eqI)  | 
|
166  | 
||
167  | 
lemma cast_isodefl: "isodefl d t \<Longrightarrow> cast\<cdot>t = (\<Lambda> x. emb\<cdot>(d\<cdot>(prj\<cdot>x)))"  | 
|
168  | 
unfolding isodefl_def by (simp add: cfun_eqI)  | 
|
169  | 
||
170  | 
lemma isodefl_strict: "isodefl d t \<Longrightarrow> d\<cdot>\<bottom> = \<bottom>"  | 
|
171  | 
unfolding isodefl_def  | 
|
172  | 
by (drule cfun_fun_cong [where x="\<bottom>"], simp)  | 
|
173  | 
||
174  | 
lemma isodefl_imp_deflation:  | 
|
175  | 
fixes d :: "'a \<rightarrow> 'a"  | 
|
176  | 
assumes "isodefl d t" shows "deflation d"  | 
|
177  | 
proof  | 
|
178  | 
note assms [unfolded isodefl_def, simp]  | 
|
179  | 
fix x :: 'a  | 
|
180  | 
show "d\<cdot>(d\<cdot>x) = d\<cdot>x"  | 
|
181  | 
using cast.idem [of t "emb\<cdot>x"] by simp  | 
|
182  | 
show "d\<cdot>x \<sqsubseteq> x"  | 
|
183  | 
using cast.below [of t "emb\<cdot>x"] by simp  | 
|
184  | 
qed  | 
|
185  | 
||
186  | 
lemma isodefl_ID_DEFL: "isodefl (ID :: 'a \<rightarrow> 'a) DEFL('a)"
 | 
|
187  | 
unfolding isodefl_def by (simp add: cast_DEFL)  | 
|
188  | 
||
189  | 
lemma isodefl_LIFTDEFL:  | 
|
| 
41292
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
190  | 
  "isodefl' (ID :: 'a \<rightarrow> 'a) LIFTDEFL('a::predomain)"
 | 
| 
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
191  | 
unfolding isodefl'_def by (simp add: cast_liftdefl u_map_ID)  | 
| 40504 | 192  | 
|
193  | 
lemma isodefl_DEFL_imp_ID: "isodefl (d :: 'a \<rightarrow> 'a) DEFL('a) \<Longrightarrow> d = ID"
 | 
|
194  | 
unfolding isodefl_def  | 
|
195  | 
apply (simp add: cast_DEFL)  | 
|
196  | 
apply (simp add: cfun_eq_iff)  | 
|
197  | 
apply (rule allI)  | 
|
198  | 
apply (drule_tac x="emb\<cdot>x" in spec)  | 
|
199  | 
apply simp  | 
|
200  | 
done  | 
|
201  | 
||
202  | 
lemma isodefl_bottom: "isodefl \<bottom> \<bottom>"  | 
|
203  | 
unfolding isodefl_def by (simp add: cfun_eq_iff)  | 
|
204  | 
||
205  | 
lemma adm_isodefl:  | 
|
206  | 
"cont f \<Longrightarrow> cont g \<Longrightarrow> adm (\<lambda>x. isodefl (f x) (g x))"  | 
|
207  | 
unfolding isodefl_def by simp  | 
|
208  | 
||
209  | 
lemma isodefl_lub:  | 
|
210  | 
assumes "chain d" and "chain t"  | 
|
211  | 
assumes "\<And>i. isodefl (d i) (t i)"  | 
|
212  | 
shows "isodefl (\<Squnion>i. d i) (\<Squnion>i. t i)"  | 
|
| 41529 | 213  | 
using assms unfolding isodefl_def  | 
| 40504 | 214  | 
by (simp add: contlub_cfun_arg contlub_cfun_fun)  | 
215  | 
||
216  | 
lemma isodefl_fix:  | 
|
217  | 
assumes "\<And>d t. isodefl d t \<Longrightarrow> isodefl (f\<cdot>d) (g\<cdot>t)"  | 
|
218  | 
shows "isodefl (fix\<cdot>f) (fix\<cdot>g)"  | 
|
219  | 
unfolding fix_def2  | 
|
220  | 
apply (rule isodefl_lub, simp, simp)  | 
|
221  | 
apply (induct_tac i)  | 
|
222  | 
apply (simp add: isodefl_bottom)  | 
|
223  | 
apply (simp add: assms)  | 
|
224  | 
done  | 
|
225  | 
||
226  | 
lemma isodefl_abs_rep:  | 
|
227  | 
fixes abs and rep and d  | 
|
228  | 
  assumes DEFL: "DEFL('b) = DEFL('a)"
 | 
|
229  | 
assumes abs_def: "(abs :: 'a \<rightarrow> 'b) \<equiv> prj oo emb"  | 
|
230  | 
assumes rep_def: "(rep :: 'b \<rightarrow> 'a) \<equiv> prj oo emb"  | 
|
231  | 
shows "isodefl d t \<Longrightarrow> isodefl (abs oo d oo rep) t"  | 
|
232  | 
unfolding isodefl_def  | 
|
233  | 
by (simp add: cfun_eq_iff assms prj_emb_prj emb_prj_emb)  | 
|
234  | 
||
| 41436 | 235  | 
lemma isodefl'_liftdefl_of: "isodefl d t \<Longrightarrow> isodefl' d (liftdefl_of\<cdot>t)"  | 
| 
41292
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
236  | 
unfolding isodefl_def isodefl'_def  | 
| 41436 | 237  | 
by (simp add: cast_liftdefl_of u_map_oo liftemb_eq liftprj_eq)  | 
| 
41292
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
238  | 
|
| 
40592
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
239  | 
lemma isodefl_sfun:  | 
| 40504 | 240  | 
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>  | 
| 
40592
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
241  | 
isodefl (sfun_map\<cdot>d1\<cdot>d2) (sfun_defl\<cdot>t1\<cdot>t2)"  | 
| 40504 | 242  | 
apply (rule isodeflI)  | 
| 
40592
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
243  | 
apply (simp add: cast_sfun_defl cast_isodefl)  | 
| 
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
244  | 
apply (simp add: emb_sfun_def prj_sfun_def)  | 
| 
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
245  | 
apply (simp add: sfun_map_map isodefl_strict)  | 
| 40504 | 246  | 
done  | 
247  | 
||
248  | 
lemma isodefl_ssum:  | 
|
249  | 
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>  | 
|
250  | 
isodefl (ssum_map\<cdot>d1\<cdot>d2) (ssum_defl\<cdot>t1\<cdot>t2)"  | 
|
251  | 
apply (rule isodeflI)  | 
|
252  | 
apply (simp add: cast_ssum_defl cast_isodefl)  | 
|
253  | 
apply (simp add: emb_ssum_def prj_ssum_def)  | 
|
254  | 
apply (simp add: ssum_map_map isodefl_strict)  | 
|
255  | 
done  | 
|
256  | 
||
257  | 
lemma isodefl_sprod:  | 
|
258  | 
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>  | 
|
259  | 
isodefl (sprod_map\<cdot>d1\<cdot>d2) (sprod_defl\<cdot>t1\<cdot>t2)"  | 
|
260  | 
apply (rule isodeflI)  | 
|
261  | 
apply (simp add: cast_sprod_defl cast_isodefl)  | 
|
262  | 
apply (simp add: emb_sprod_def prj_sprod_def)  | 
|
263  | 
apply (simp add: sprod_map_map isodefl_strict)  | 
|
264  | 
done  | 
|
265  | 
||
| 41297 | 266  | 
lemma isodefl_prod:  | 
| 40504 | 267  | 
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>  | 
| 41297 | 268  | 
isodefl (prod_map\<cdot>d1\<cdot>d2) (prod_defl\<cdot>t1\<cdot>t2)"  | 
| 40504 | 269  | 
apply (rule isodeflI)  | 
270  | 
apply (simp add: cast_prod_defl cast_isodefl)  | 
|
271  | 
apply (simp add: emb_prod_def prj_prod_def)  | 
|
| 41297 | 272  | 
apply (simp add: prod_map_map cfcomp1)  | 
| 40504 | 273  | 
done  | 
274  | 
||
275  | 
lemma isodefl_u:  | 
|
| 41437 | 276  | 
"isodefl d t \<Longrightarrow> isodefl (u_map\<cdot>d) (u_defl\<cdot>t)"  | 
| 40504 | 277  | 
apply (rule isodeflI)  | 
| 41437 | 278  | 
apply (simp add: cast_u_defl cast_isodefl)  | 
279  | 
apply (simp add: emb_u_def prj_u_def liftemb_eq liftprj_eq u_map_map)  | 
|
280  | 
done  | 
|
281  | 
||
282  | 
lemma isodefl_u_liftdefl:  | 
|
283  | 
"isodefl' d t \<Longrightarrow> isodefl (u_map\<cdot>d) (u_liftdefl\<cdot>t)"  | 
|
284  | 
apply (rule isodeflI)  | 
|
285  | 
apply (simp add: cast_u_liftdefl isodefl'_def)  | 
|
| 40504 | 286  | 
apply (simp add: emb_u_def prj_u_def liftemb_eq liftprj_eq)  | 
287  | 
done  | 
|
288  | 
||
289  | 
lemma encode_prod_u_map:  | 
|
| 41297 | 290  | 
"encode_prod_u\<cdot>(u_map\<cdot>(prod_map\<cdot>f\<cdot>g)\<cdot>(decode_prod_u\<cdot>x))  | 
| 40504 | 291  | 
= sprod_map\<cdot>(u_map\<cdot>f)\<cdot>(u_map\<cdot>g)\<cdot>x"  | 
292  | 
unfolding encode_prod_u_def decode_prod_u_def  | 
|
293  | 
apply (case_tac x, simp, rename_tac a b)  | 
|
294  | 
apply (case_tac a, simp, case_tac b, simp, simp)  | 
|
295  | 
done  | 
|
296  | 
||
| 41297 | 297  | 
lemma isodefl_prod_u:  | 
| 
41292
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
298  | 
assumes "isodefl' d1 t1" and "isodefl' d2 t2"  | 
| 41297 | 299  | 
shows "isodefl' (prod_map\<cdot>d1\<cdot>d2) (prod_liftdefl\<cdot>t1\<cdot>t2)"  | 
| 
41292
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
300  | 
using assms unfolding isodefl'_def  | 
| 
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
301  | 
unfolding liftemb_prod_def liftprj_prod_def  | 
| 
 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 
huffman 
parents: 
41290 
diff
changeset
 | 
302  | 
by (simp add: cast_prod_liftdefl cfcomp1 encode_prod_u_map sprod_map_map)  | 
| 40504 | 303  | 
|
| 
40592
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
304  | 
lemma encode_cfun_map:  | 
| 
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
305  | 
"encode_cfun\<cdot>(cfun_map\<cdot>f\<cdot>g\<cdot>(decode_cfun\<cdot>x))  | 
| 
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
306  | 
= sfun_map\<cdot>(u_map\<cdot>f)\<cdot>g\<cdot>x"  | 
| 
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
307  | 
unfolding encode_cfun_def decode_cfun_def  | 
| 
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
308  | 
apply (simp add: sfun_eq_iff cfun_map_def sfun_map_def)  | 
| 
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
309  | 
apply (rule cfun_eqI, rename_tac y, case_tac y, simp_all)  | 
| 
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
310  | 
done  | 
| 
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
311  | 
|
| 
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
312  | 
lemma isodefl_cfun:  | 
| 40830 | 313  | 
assumes "isodefl (u_map\<cdot>d1) t1" and "isodefl d2 t2"  | 
314  | 
shows "isodefl (cfun_map\<cdot>d1\<cdot>d2) (sfun_defl\<cdot>t1\<cdot>t2)"  | 
|
315  | 
using isodefl_sfun [OF assms] unfolding isodefl_def  | 
|
316  | 
by (simp add: emb_cfun_def prj_cfun_def cfcomp1 encode_cfun_map)  | 
|
| 
40592
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
317  | 
|
| 62175 | 318  | 
subsection \<open>Setting up the domain package\<close>  | 
| 40504 | 319  | 
|
| 
57945
 
cacb00a569e0
prefer 'named_theorems' over Named_Thms, with subtle change of semantics due to visual order vs. internal reverse order;
 
wenzelm 
parents: 
48891 
diff
changeset
 | 
320  | 
named_theorems domain_defl_simps "theorems like DEFL('a t) = t_defl$DEFL('a)"
 | 
| 59028 | 321  | 
and domain_isodefl "theorems like isodefl d t ==> isodefl (foo_map$d) (foo_defl$t)"  | 
| 
57945
 
cacb00a569e0
prefer 'named_theorems' over Named_Thms, with subtle change of semantics due to visual order vs. internal reverse order;
 
wenzelm 
parents: 
48891 
diff
changeset
 | 
322  | 
|
| 69605 | 323  | 
ML_file \<open>Tools/Domain/domain_isomorphism.ML\<close>  | 
324  | 
ML_file \<open>Tools/Domain/domain_axioms.ML\<close>  | 
|
325  | 
ML_file \<open>Tools/Domain/domain.ML\<close>  | 
|
| 40504 | 326  | 
|
327  | 
lemmas [domain_defl_simps] =  | 
|
| 
40592
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
328  | 
DEFL_cfun DEFL_sfun DEFL_ssum DEFL_sprod DEFL_prod DEFL_u  | 
| 41437 | 329  | 
liftdefl_eq LIFTDEFL_prod u_liftdefl_liftdefl_of  | 
| 40504 | 330  | 
|
331  | 
lemmas [domain_map_ID] =  | 
|
| 41297 | 332  | 
cfun_map_ID sfun_map_ID ssum_map_ID sprod_map_ID prod_map_ID u_map_ID  | 
| 40504 | 333  | 
|
334  | 
lemmas [domain_isodefl] =  | 
|
| 
40592
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
335  | 
isodefl_u isodefl_sfun isodefl_ssum isodefl_sprod  | 
| 41436 | 336  | 
isodefl_cfun isodefl_prod isodefl_prod_u isodefl'_liftdefl_of  | 
| 41437 | 337  | 
isodefl_u_liftdefl  | 
| 40504 | 338  | 
|
339  | 
lemmas [domain_deflation] =  | 
|
| 
40592
 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 
huffman 
parents: 
40575 
diff
changeset
 | 
340  | 
deflation_cfun_map deflation_sfun_map deflation_ssum_map  | 
| 41297 | 341  | 
deflation_sprod_map deflation_prod_map deflation_u_map  | 
| 40504 | 342  | 
|
| 62175 | 343  | 
setup \<open>  | 
| 
40737
 
2037021f034f
remove map function names from domain package theory data
 
huffman 
parents: 
40592 
diff
changeset
 | 
344  | 
fold Domain_Take_Proofs.add_rec_type  | 
| 68357 | 345  | 
[(\<^type_name>\<open>cfun\<close>, [true, true]),  | 
346  | 
(\<^type_name>\<open>sfun\<close>, [true, true]),  | 
|
347  | 
(\<^type_name>\<open>ssum\<close>, [true, true]),  | 
|
348  | 
(\<^type_name>\<open>sprod\<close>, [true, true]),  | 
|
349  | 
(\<^type_name>\<open>prod\<close>, [true, true]),  | 
|
350  | 
(\<^type_name>\<open>u\<close>, [true])]  | 
|
| 62175 | 351  | 
\<close>  | 
| 40504 | 352  | 
|
| 15741 | 353  | 
end  |