| author | bulwahn | 
| Tue, 07 Sep 2010 11:51:53 +0200 | |
| changeset 39189 | d183bf90dabd | 
| parent 37770 | cddb3106adb8 | 
| child 39974 | b525988432e9 | 
| permissions | -rw-r--r-- | 
| 33589 | 1  | 
(* Title: HOLCF/Representable.thy  | 
2  | 
Author: Brian Huffman  | 
|
3  | 
*)  | 
|
4  | 
||
| 33588 | 5  | 
header {* Representable Types *}
 | 
6  | 
||
7  | 
theory Representable  | 
|
| 
35652
 
05ca920cd94b
move take-proofs stuff into new theory Domain_Aux.thy
 
huffman 
parents: 
35596 
diff
changeset
 | 
8  | 
imports Algebraic Universal Ssum Sprod One Fixrec Domain_Aux  | 
| 
33794
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
9  | 
uses  | 
| 
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
10  | 
  ("Tools/repdef.ML")
 | 
| 
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
11  | 
  ("Tools/Domain/domain_isomorphism.ML")
 | 
| 33588 | 12  | 
begin  | 
13  | 
||
14  | 
subsection {* Class of representable types *}
 | 
|
15  | 
||
16  | 
text "Overloaded embedding and projection functions between  | 
|
17  | 
a representable type and the universal domain."  | 
|
18  | 
||
19  | 
class rep = bifinite +  | 
|
20  | 
fixes emb :: "'a::pcpo \<rightarrow> udom"  | 
|
21  | 
fixes prj :: "udom \<rightarrow> 'a::pcpo"  | 
|
22  | 
assumes ep_pair_emb_prj: "ep_pair emb prj"  | 
|
23  | 
||
| 
37770
 
cddb3106adb8
avoid explicit mandatory prefix markers when prefixes are mandatory implicitly
 
haftmann 
parents: 
37678 
diff
changeset
 | 
24  | 
interpretation rep:  | 
| 33588 | 25  | 
pcpo_ep_pair  | 
26  | 
"emb :: 'a::rep \<rightarrow> udom"  | 
|
27  | 
"prj :: udom \<rightarrow> 'a::rep"  | 
|
28  | 
unfolding pcpo_ep_pair_def  | 
|
29  | 
by (rule ep_pair_emb_prj)  | 
|
30  | 
||
31  | 
lemmas emb_inverse = rep.e_inverse  | 
|
32  | 
lemmas emb_prj_below = rep.e_p_below  | 
|
33  | 
lemmas emb_eq_iff = rep.e_eq_iff  | 
|
34  | 
lemmas emb_strict = rep.e_strict  | 
|
35  | 
lemmas prj_strict = rep.p_strict  | 
|
36  | 
||
37  | 
||
| 
35900
 
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
 
huffman 
parents: 
35652 
diff
changeset
 | 
38  | 
subsection {* Making \emph{rep} the default class *}
 | 
| 33588 | 39  | 
|
40  | 
text {*
 | 
|
41  | 
From now on, free type variables are assumed to be in class  | 
|
42  | 
  @{term rep}, unless specified otherwise.
 | 
|
43  | 
*}  | 
|
44  | 
||
| 36452 | 45  | 
default_sort rep  | 
| 33588 | 46  | 
|
47  | 
subsection {* Representations of types *}
 | 
|
48  | 
||
49  | 
text "A TypeRep is an algebraic deflation over the universe of values."  | 
|
50  | 
||
51  | 
types TypeRep = "udom alg_defl"  | 
|
| 35431 | 52  | 
translations (type) "TypeRep" \<leftharpoondown> (type) "udom alg_defl"  | 
| 33588 | 53  | 
|
54  | 
definition  | 
|
55  | 
Rep_of :: "'a::rep itself \<Rightarrow> TypeRep"  | 
|
56  | 
where  | 
|
57  | 
  "Rep_of TYPE('a::rep) =
 | 
|
58  | 
(\<Squnion>i. alg_defl_principal (Abs_fin_defl  | 
|
59  | 
(emb oo (approx i :: 'a \<rightarrow> 'a) oo prj)))"  | 
|
60  | 
||
61  | 
syntax "_REP" :: "type \<Rightarrow> TypeRep"  ("(1REP/(1'(_')))")
 | 
|
| 35431 | 62  | 
translations "REP('t)" \<rightleftharpoons> "CONST Rep_of TYPE('t)"
 | 
| 33588 | 63  | 
|
64  | 
lemma cast_REP:  | 
|
65  | 
  "cast\<cdot>REP('a::rep) = (emb::'a \<rightarrow> udom) oo (prj::udom \<rightarrow> 'a)"
 | 
|
66  | 
proof -  | 
|
67  | 
let ?a = "\<lambda>i. emb oo approx i oo (prj::udom \<rightarrow> 'a)"  | 
|
68  | 
have a: "\<And>i. finite_deflation (?a i)"  | 
|
69  | 
apply (rule rep.finite_deflation_e_d_p)  | 
|
70  | 
apply (rule finite_deflation_approx)  | 
|
71  | 
done  | 
|
72  | 
show ?thesis  | 
|
73  | 
unfolding Rep_of_def  | 
|
74  | 
apply (subst contlub_cfun_arg)  | 
|
75  | 
apply (rule chainI)  | 
|
76  | 
apply (rule alg_defl.principal_mono)  | 
|
77  | 
apply (rule Abs_fin_defl_mono [OF a a])  | 
|
78  | 
apply (rule chainE, simp)  | 
|
79  | 
apply (subst cast_alg_defl_principal)  | 
|
80  | 
apply (simp add: Abs_fin_defl_inverse a)  | 
|
81  | 
apply (simp add: expand_cfun_eq lub_distribs)  | 
|
82  | 
done  | 
|
83  | 
qed  | 
|
84  | 
||
85  | 
lemma emb_prj: "emb\<cdot>((prj\<cdot>x)::'a::rep) = cast\<cdot>REP('a)\<cdot>x"
 | 
|
86  | 
by (simp add: cast_REP)  | 
|
87  | 
||
88  | 
lemma in_REP_iff:  | 
|
89  | 
  "x ::: REP('a::rep) \<longleftrightarrow> emb\<cdot>((prj\<cdot>x)::'a) = x"
 | 
|
90  | 
by (simp add: in_deflation_def cast_REP)  | 
|
91  | 
||
92  | 
lemma prj_inverse:  | 
|
93  | 
  "x ::: REP('a::rep) \<Longrightarrow> emb\<cdot>((prj\<cdot>x)::'a) = x"
 | 
|
94  | 
by (simp only: in_REP_iff)  | 
|
95  | 
||
96  | 
lemma emb_in_REP [simp]:  | 
|
97  | 
  "emb\<cdot>(x::'a::rep) ::: REP('a)"
 | 
|
98  | 
by (simp add: in_REP_iff)  | 
|
99  | 
||
100  | 
subsection {* Coerce operator *}
 | 
|
101  | 
||
102  | 
definition coerce :: "'a \<rightarrow> 'b"  | 
|
103  | 
where "coerce = prj oo emb"  | 
|
104  | 
||
105  | 
lemma beta_coerce: "coerce\<cdot>x = prj\<cdot>(emb\<cdot>x)"  | 
|
106  | 
by (simp add: coerce_def)  | 
|
107  | 
||
108  | 
lemma prj_emb: "prj\<cdot>(emb\<cdot>x) = coerce\<cdot>x"  | 
|
109  | 
by (simp add: coerce_def)  | 
|
110  | 
||
111  | 
lemma coerce_strict [simp]: "coerce\<cdot>\<bottom> = \<bottom>"  | 
|
112  | 
by (simp add: coerce_def)  | 
|
113  | 
||
114  | 
lemma coerce_eq_ID [simp]: "(coerce :: 'a \<rightarrow> 'a) = ID"  | 
|
115  | 
by (rule ext_cfun, simp add: beta_coerce)  | 
|
116  | 
||
117  | 
lemma emb_coerce:  | 
|
118  | 
  "REP('a) \<sqsubseteq> REP('b)
 | 
|
119  | 
\<Longrightarrow> emb\<cdot>((coerce::'a \<rightarrow> 'b)\<cdot>x) = emb\<cdot>x"  | 
|
120  | 
apply (simp add: beta_coerce)  | 
|
121  | 
apply (rule prj_inverse)  | 
|
122  | 
apply (erule subdeflationD)  | 
|
123  | 
apply (rule emb_in_REP)  | 
|
124  | 
done  | 
|
125  | 
||
126  | 
lemma coerce_prj:  | 
|
127  | 
  "REP('a) \<sqsubseteq> REP('b)
 | 
|
128  | 
\<Longrightarrow> (coerce::'b \<rightarrow> 'a)\<cdot>(prj\<cdot>x) = prj\<cdot>x"  | 
|
129  | 
apply (simp add: coerce_def)  | 
|
130  | 
apply (rule emb_eq_iff [THEN iffD1])  | 
|
131  | 
apply (simp only: emb_prj)  | 
|
132  | 
apply (rule deflation_below_comp1)  | 
|
133  | 
apply (rule deflation_cast)  | 
|
134  | 
apply (rule deflation_cast)  | 
|
135  | 
apply (erule monofun_cfun_arg)  | 
|
136  | 
done  | 
|
137  | 
||
138  | 
lemma coerce_coerce [simp]:  | 
|
139  | 
  "REP('a) \<sqsubseteq> REP('b)
 | 
|
140  | 
\<Longrightarrow> coerce\<cdot>((coerce::'a \<rightarrow> 'b)\<cdot>x) = coerce\<cdot>x"  | 
|
141  | 
by (simp add: beta_coerce prj_inverse subdeflationD)  | 
|
142  | 
||
143  | 
lemma coerce_inverse:  | 
|
144  | 
  "emb\<cdot>(x::'a) ::: REP('b) \<Longrightarrow> coerce\<cdot>(coerce\<cdot>x :: 'b) = x"
 | 
|
145  | 
by (simp only: beta_coerce prj_inverse emb_inverse)  | 
|
146  | 
||
147  | 
lemma coerce_type:  | 
|
148  | 
  "REP('a) \<sqsubseteq> REP('b)
 | 
|
149  | 
   \<Longrightarrow> emb\<cdot>((coerce::'a \<rightarrow> 'b)\<cdot>x) ::: REP('a)"
 | 
|
150  | 
by (simp add: beta_coerce prj_inverse subdeflationD)  | 
|
151  | 
||
152  | 
lemma ep_pair_coerce:  | 
|
153  | 
  "REP('a) \<sqsubseteq> REP('b)
 | 
|
154  | 
\<Longrightarrow> ep_pair (coerce::'a \<rightarrow> 'b) (coerce::'b \<rightarrow> 'a)"  | 
|
155  | 
apply (rule ep_pair.intro)  | 
|
156  | 
apply simp  | 
|
157  | 
apply (simp only: beta_coerce)  | 
|
158  | 
apply (rule below_trans)  | 
|
159  | 
apply (rule monofun_cfun_arg)  | 
|
160  | 
apply (rule emb_prj_below)  | 
|
161  | 
apply simp  | 
|
162  | 
done  | 
|
163  | 
||
| 33779 | 164  | 
text {* Isomorphism lemmas used internally by the domain package: *}
 | 
165  | 
||
166  | 
lemma domain_abs_iso:  | 
|
167  | 
fixes abs and rep  | 
|
168  | 
  assumes REP: "REP('b) = REP('a)"
 | 
|
169  | 
assumes abs_def: "abs \<equiv> (coerce :: 'a \<rightarrow> 'b)"  | 
|
170  | 
assumes rep_def: "rep \<equiv> (coerce :: 'b \<rightarrow> 'a)"  | 
|
171  | 
shows "rep\<cdot>(abs\<cdot>x) = x"  | 
|
172  | 
unfolding abs_def rep_def by (simp add: REP)  | 
|
173  | 
||
174  | 
lemma domain_rep_iso:  | 
|
175  | 
fixes abs and rep  | 
|
176  | 
  assumes REP: "REP('b) = REP('a)"
 | 
|
177  | 
assumes abs_def: "abs \<equiv> (coerce :: 'a \<rightarrow> 'b)"  | 
|
178  | 
assumes rep_def: "rep \<equiv> (coerce :: 'b \<rightarrow> 'a)"  | 
|
179  | 
shows "abs\<cdot>(rep\<cdot>x) = x"  | 
|
180  | 
unfolding abs_def rep_def by (simp add: REP [symmetric])  | 
|
181  | 
||
182  | 
||
| 33588 | 183  | 
subsection {* Proving a subtype is representable *}
 | 
184  | 
||
185  | 
text {*
 | 
|
186  | 
  Temporarily relax type constraints for @{term "approx"},
 | 
|
187  | 
  @{term emb}, and @{term prj}.
 | 
|
188  | 
*}  | 
|
189  | 
||
190  | 
setup {* Sign.add_const_constraint
 | 
|
191  | 
  (@{const_name "approx"}, SOME @{typ "nat \<Rightarrow> 'a::cpo \<rightarrow> 'a"}) *}
 | 
|
192  | 
||
193  | 
setup {* Sign.add_const_constraint
 | 
|
194  | 
  (@{const_name emb}, SOME @{typ "'a::pcpo \<rightarrow> udom"}) *}
 | 
|
195  | 
||
196  | 
setup {* Sign.add_const_constraint
 | 
|
197  | 
  (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::pcpo"}) *}
 | 
|
198  | 
||
| 
33679
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
199  | 
definition  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
200  | 
repdef_approx ::  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
201  | 
    "('a::pcpo \<Rightarrow> udom) \<Rightarrow> (udom \<Rightarrow> 'a) \<Rightarrow> udom alg_defl \<Rightarrow> nat \<Rightarrow> 'a \<rightarrow> 'a"
 | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
202  | 
where  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
203  | 
"repdef_approx Rep Abs t = (\<lambda>i. \<Lambda> x. Abs (cast\<cdot>(approx i\<cdot>t)\<cdot>(Rep x)))"  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
204  | 
|
| 33588 | 205  | 
lemma typedef_rep_class:  | 
206  | 
fixes Rep :: "'a::pcpo \<Rightarrow> udom"  | 
|
207  | 
fixes Abs :: "udom \<Rightarrow> 'a::pcpo"  | 
|
208  | 
fixes t :: TypeRep  | 
|
209  | 
  assumes type: "type_definition Rep Abs {x. x ::: t}"
 | 
|
210  | 
assumes below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"  | 
|
| 
33679
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
211  | 
assumes emb: "emb \<equiv> (\<Lambda> x. Rep x)"  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
212  | 
assumes prj: "prj \<equiv> (\<Lambda> x. Abs (cast\<cdot>t\<cdot>x))"  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
213  | 
assumes approx: "(approx :: nat \<Rightarrow> 'a \<rightarrow> 'a) \<equiv> repdef_approx Rep Abs t"  | 
| 33588 | 214  | 
  shows "OFCLASS('a, rep_class)"
 | 
215  | 
proof  | 
|
216  | 
  have adm: "adm (\<lambda>x. x \<in> {x. x ::: t})"
 | 
|
217  | 
by (simp add: adm_in_deflation)  | 
|
218  | 
have emb_beta: "\<And>x. emb\<cdot>x = Rep x"  | 
|
219  | 
unfolding emb  | 
|
220  | 
apply (rule beta_cfun)  | 
|
221  | 
apply (rule typedef_cont_Rep [OF type below adm])  | 
|
222  | 
done  | 
|
223  | 
have prj_beta: "\<And>y. prj\<cdot>y = Abs (cast\<cdot>t\<cdot>y)"  | 
|
224  | 
unfolding prj  | 
|
225  | 
apply (rule beta_cfun)  | 
|
226  | 
apply (rule typedef_cont_Abs [OF type below adm])  | 
|
227  | 
apply simp_all  | 
|
228  | 
done  | 
|
| 
33679
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
229  | 
have cast_cast_approx:  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
230  | 
"\<And>i x. cast\<cdot>t\<cdot>(cast\<cdot>(approx i\<cdot>t)\<cdot>x) = cast\<cdot>(approx i\<cdot>t)\<cdot>x"  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
231  | 
apply (rule cast_fixed)  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
232  | 
apply (rule subdeflationD)  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
233  | 
apply (rule approx.below)  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
234  | 
apply (rule cast_in_deflation)  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
235  | 
done  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
236  | 
have approx':  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
237  | 
"approx = (\<lambda>i. \<Lambda>(x::'a). prj\<cdot>(cast\<cdot>(approx i\<cdot>t)\<cdot>(emb\<cdot>x)))"  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
238  | 
unfolding approx repdef_approx_def  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
239  | 
apply (subst cast_cast_approx [symmetric])  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
240  | 
apply (simp add: prj_beta [symmetric] emb_beta [symmetric])  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
241  | 
done  | 
| 33588 | 242  | 
have emb_in_deflation: "\<And>x::'a. emb\<cdot>x ::: t"  | 
243  | 
using type_definition.Rep [OF type]  | 
|
244  | 
by (simp add: emb_beta)  | 
|
245  | 
have prj_emb: "\<And>x::'a. prj\<cdot>(emb\<cdot>x) = x"  | 
|
246  | 
unfolding prj_beta  | 
|
247  | 
apply (simp add: cast_fixed [OF emb_in_deflation])  | 
|
248  | 
apply (simp add: emb_beta type_definition.Rep_inverse [OF type])  | 
|
249  | 
done  | 
|
250  | 
have emb_prj: "\<And>y. emb\<cdot>(prj\<cdot>y :: 'a) = cast\<cdot>t\<cdot>y"  | 
|
251  | 
unfolding prj_beta emb_beta  | 
|
252  | 
by (simp add: type_definition.Abs_inverse [OF type])  | 
|
253  | 
show "ep_pair (emb :: 'a \<rightarrow> udom) prj"  | 
|
254  | 
apply default  | 
|
255  | 
apply (simp add: prj_emb)  | 
|
256  | 
apply (simp add: emb_prj cast.below)  | 
|
257  | 
done  | 
|
258  | 
show "chain (approx :: nat \<Rightarrow> 'a \<rightarrow> 'a)"  | 
|
| 
33679
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
259  | 
unfolding approx' by simp  | 
| 33588 | 260  | 
show "\<And>x::'a. (\<Squnion>i. approx i\<cdot>x) = x"  | 
| 
33679
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
261  | 
unfolding approx'  | 
| 33588 | 262  | 
apply (simp add: lub_distribs)  | 
263  | 
apply (subst cast_fixed [OF emb_in_deflation])  | 
|
264  | 
apply (rule prj_emb)  | 
|
265  | 
done  | 
|
266  | 
show "\<And>(i::nat) (x::'a). approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"  | 
|
| 
33679
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
267  | 
unfolding approx'  | 
| 33588 | 268  | 
apply simp  | 
269  | 
apply (simp add: emb_prj)  | 
|
270  | 
apply (simp add: cast_cast_approx)  | 
|
271  | 
done  | 
|
272  | 
  show "\<And>i::nat. finite {x::'a. approx i\<cdot>x = x}"
 | 
|
273  | 
    apply (rule_tac B="(\<lambda>x. prj\<cdot>x) ` {x. cast\<cdot>(approx i\<cdot>t)\<cdot>x = x}"
 | 
|
274  | 
in finite_subset)  | 
|
| 
33679
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
275  | 
apply (clarsimp simp add: approx')  | 
| 33588 | 276  | 
apply (drule_tac f="\<lambda>x. emb\<cdot>x" in arg_cong)  | 
277  | 
apply (rule image_eqI)  | 
|
278  | 
apply (rule prj_emb [symmetric])  | 
|
279  | 
apply (simp add: emb_prj)  | 
|
280  | 
apply (simp add: cast_cast_approx)  | 
|
281  | 
apply (rule finite_imageI)  | 
|
282  | 
apply (simp add: cast_approx_fixed_iff)  | 
|
283  | 
apply (simp add: Collect_conj_eq)  | 
|
284  | 
apply (simp add: finite_fixes_approx)  | 
|
285  | 
done  | 
|
286  | 
qed  | 
|
287  | 
||
288  | 
text {* Restore original typing constraints *}
 | 
|
289  | 
||
290  | 
setup {* Sign.add_const_constraint
 | 
|
291  | 
  (@{const_name "approx"}, SOME @{typ "nat \<Rightarrow> 'a::profinite \<rightarrow> 'a"}) *}
 | 
|
292  | 
||
293  | 
setup {* Sign.add_const_constraint
 | 
|
294  | 
  (@{const_name emb}, SOME @{typ "'a::rep \<rightarrow> udom"}) *}
 | 
|
295  | 
||
296  | 
setup {* Sign.add_const_constraint
 | 
|
297  | 
  (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::rep"}) *}
 | 
|
298  | 
||
299  | 
lemma typedef_REP:  | 
|
300  | 
fixes Rep :: "'a::rep \<Rightarrow> udom"  | 
|
301  | 
fixes Abs :: "udom \<Rightarrow> 'a::rep"  | 
|
302  | 
fixes t :: TypeRep  | 
|
303  | 
  assumes type: "type_definition Rep Abs {x. x ::: t}"
 | 
|
304  | 
assumes below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"  | 
|
| 
33679
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
305  | 
assumes emb: "emb \<equiv> (\<Lambda> x. Rep x)"  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
306  | 
assumes prj: "prj \<equiv> (\<Lambda> x. Abs (cast\<cdot>t\<cdot>x))"  | 
| 33588 | 307  | 
  shows "REP('a) = t"
 | 
308  | 
proof -  | 
|
309  | 
  have adm: "adm (\<lambda>x. x \<in> {x. x ::: t})"
 | 
|
310  | 
by (simp add: adm_in_deflation)  | 
|
311  | 
have emb_beta: "\<And>x. emb\<cdot>x = Rep x"  | 
|
312  | 
unfolding emb  | 
|
313  | 
apply (rule beta_cfun)  | 
|
314  | 
apply (rule typedef_cont_Rep [OF type below adm])  | 
|
315  | 
done  | 
|
316  | 
have prj_beta: "\<And>y. prj\<cdot>y = Abs (cast\<cdot>t\<cdot>y)"  | 
|
317  | 
unfolding prj  | 
|
318  | 
apply (rule beta_cfun)  | 
|
319  | 
apply (rule typedef_cont_Abs [OF type below adm])  | 
|
320  | 
apply simp_all  | 
|
321  | 
done  | 
|
322  | 
have emb_in_deflation: "\<And>x::'a. emb\<cdot>x ::: t"  | 
|
323  | 
using type_definition.Rep [OF type]  | 
|
324  | 
by (simp add: emb_beta)  | 
|
325  | 
have prj_emb: "\<And>x::'a. prj\<cdot>(emb\<cdot>x) = x"  | 
|
326  | 
unfolding prj_beta  | 
|
327  | 
apply (simp add: cast_fixed [OF emb_in_deflation])  | 
|
328  | 
apply (simp add: emb_beta type_definition.Rep_inverse [OF type])  | 
|
329  | 
done  | 
|
330  | 
have emb_prj: "\<And>y. emb\<cdot>(prj\<cdot>y :: 'a) = cast\<cdot>t\<cdot>y"  | 
|
331  | 
unfolding prj_beta emb_beta  | 
|
332  | 
by (simp add: type_definition.Abs_inverse [OF type])  | 
|
333  | 
  show "REP('a) = t"
 | 
|
334  | 
apply (rule cast_eq_imp_eq, rule ext_cfun)  | 
|
335  | 
apply (simp add: cast_REP emb_prj)  | 
|
336  | 
done  | 
|
337  | 
qed  | 
|
338  | 
||
| 
33679
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
339  | 
lemma adm_mem_Collect_in_deflation: "adm (\<lambda>x. x \<in> {x. x ::: A})"
 | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
340  | 
unfolding mem_Collect_eq by (rule adm_in_deflation)  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
341  | 
|
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
342  | 
use "Tools/repdef.ML"  | 
| 
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33589 
diff
changeset
 | 
343  | 
|
| 33588 | 344  | 
|
| 
35900
 
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
 
huffman 
parents: 
35652 
diff
changeset
 | 
345  | 
subsection {* Instances of class \emph{rep} *}
 | 
| 33588 | 346  | 
|
347  | 
subsubsection {* Universal Domain *}
 | 
|
348  | 
||
349  | 
text "The Universal Domain itself is trivially representable."  | 
|
350  | 
||
351  | 
instantiation udom :: rep  | 
|
352  | 
begin  | 
|
353  | 
||
354  | 
definition emb_udom_def [simp]: "emb = (ID :: udom \<rightarrow> udom)"  | 
|
355  | 
definition prj_udom_def [simp]: "prj = (ID :: udom \<rightarrow> udom)"  | 
|
356  | 
||
357  | 
instance  | 
|
358  | 
apply (intro_classes)  | 
|
359  | 
apply (simp_all add: ep_pair.intro)  | 
|
360  | 
done  | 
|
361  | 
||
362  | 
end  | 
|
363  | 
||
364  | 
subsubsection {* Lifted types *}
 | 
|
365  | 
||
366  | 
instantiation lift :: (countable) rep  | 
|
367  | 
begin  | 
|
368  | 
||
369  | 
definition emb_lift_def:  | 
|
370  | 
"emb = udom_emb oo (FLIFT x. Def (to_nat x))"  | 
|
371  | 
||
372  | 
definition prj_lift_def:  | 
|
373  | 
"prj = (FLIFT n. if (\<exists>x::'a::countable. n = to_nat x)  | 
|
374  | 
then Def (THE x::'a. n = to_nat x) else \<bottom>) oo udom_prj"  | 
|
375  | 
||
376  | 
instance  | 
|
377  | 
apply (intro_classes, unfold emb_lift_def prj_lift_def)  | 
|
378  | 
apply (rule ep_pair_comp [OF _ ep_pair_udom])  | 
|
379  | 
apply (rule ep_pair.intro)  | 
|
380  | 
apply (case_tac x, simp, simp)  | 
|
381  | 
apply (case_tac y, simp, clarsimp)  | 
|
382  | 
done  | 
|
383  | 
||
384  | 
end  | 
|
385  | 
||
386  | 
subsubsection {* Representable type constructors *}
 | 
|
387  | 
||
388  | 
text "Functions between representable types are representable."  | 
|
389  | 
||
| 35525 | 390  | 
instantiation cfun :: (rep, rep) rep  | 
| 33588 | 391  | 
begin  | 
392  | 
||
393  | 
definition emb_cfun_def: "emb = udom_emb oo cfun_map\<cdot>prj\<cdot>emb"  | 
|
394  | 
definition prj_cfun_def: "prj = cfun_map\<cdot>emb\<cdot>prj oo udom_prj"  | 
|
395  | 
||
396  | 
instance  | 
|
397  | 
apply (intro_classes, unfold emb_cfun_def prj_cfun_def)  | 
|
398  | 
apply (intro ep_pair_comp ep_pair_cfun_map ep_pair_emb_prj ep_pair_udom)  | 
|
399  | 
done  | 
|
400  | 
||
401  | 
end  | 
|
402  | 
||
403  | 
text "Strict products of representable types are representable."  | 
|
404  | 
||
| 35525 | 405  | 
instantiation sprod :: (rep, rep) rep  | 
| 33588 | 406  | 
begin  | 
407  | 
||
408  | 
definition emb_sprod_def: "emb = udom_emb oo sprod_map\<cdot>emb\<cdot>emb"  | 
|
409  | 
definition prj_sprod_def: "prj = sprod_map\<cdot>prj\<cdot>prj oo udom_prj"  | 
|
410  | 
||
411  | 
instance  | 
|
412  | 
apply (intro_classes, unfold emb_sprod_def prj_sprod_def)  | 
|
413  | 
apply (intro ep_pair_comp ep_pair_sprod_map ep_pair_emb_prj ep_pair_udom)  | 
|
414  | 
done  | 
|
415  | 
||
416  | 
end  | 
|
417  | 
||
418  | 
text "Strict sums of representable types are representable."  | 
|
419  | 
||
| 35525 | 420  | 
instantiation ssum :: (rep, rep) rep  | 
| 33588 | 421  | 
begin  | 
422  | 
||
423  | 
definition emb_ssum_def: "emb = udom_emb oo ssum_map\<cdot>emb\<cdot>emb"  | 
|
424  | 
definition prj_ssum_def: "prj = ssum_map\<cdot>prj\<cdot>prj oo udom_prj"  | 
|
425  | 
||
426  | 
instance  | 
|
427  | 
apply (intro_classes, unfold emb_ssum_def prj_ssum_def)  | 
|
428  | 
apply (intro ep_pair_comp ep_pair_ssum_map ep_pair_emb_prj ep_pair_udom)  | 
|
429  | 
done  | 
|
430  | 
||
431  | 
end  | 
|
432  | 
||
433  | 
text "Up of a representable type is representable."  | 
|
434  | 
||
435  | 
instantiation "u" :: (rep) rep  | 
|
436  | 
begin  | 
|
437  | 
||
438  | 
definition emb_u_def: "emb = udom_emb oo u_map\<cdot>emb"  | 
|
439  | 
definition prj_u_def: "prj = u_map\<cdot>prj oo udom_prj"  | 
|
440  | 
||
441  | 
instance  | 
|
442  | 
apply (intro_classes, unfold emb_u_def prj_u_def)  | 
|
443  | 
apply (intro ep_pair_comp ep_pair_u_map ep_pair_emb_prj ep_pair_udom)  | 
|
444  | 
done  | 
|
445  | 
||
446  | 
end  | 
|
447  | 
||
448  | 
text "Cartesian products of representable types are representable."  | 
|
449  | 
||
| 
37678
 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 
haftmann 
parents: 
36452 
diff
changeset
 | 
450  | 
instantiation prod :: (rep, rep) rep  | 
| 33588 | 451  | 
begin  | 
452  | 
||
453  | 
definition emb_cprod_def: "emb = udom_emb oo cprod_map\<cdot>emb\<cdot>emb"  | 
|
454  | 
definition prj_cprod_def: "prj = cprod_map\<cdot>prj\<cdot>prj oo udom_prj"  | 
|
455  | 
||
456  | 
instance  | 
|
457  | 
apply (intro_classes, unfold emb_cprod_def prj_cprod_def)  | 
|
458  | 
apply (intro ep_pair_comp ep_pair_cprod_map ep_pair_emb_prj ep_pair_udom)  | 
|
459  | 
done  | 
|
460  | 
||
461  | 
end  | 
|
462  | 
||
463  | 
subsection {* Type combinators *}
 | 
|
464  | 
||
465  | 
definition  | 
|
466  | 
TypeRep_fun1 ::  | 
|
467  | 
    "((udom \<rightarrow> udom) \<rightarrow> ('a \<rightarrow> 'a))
 | 
|
468  | 
\<Rightarrow> (TypeRep \<rightarrow> TypeRep)"  | 
|
469  | 
where  | 
|
470  | 
"TypeRep_fun1 f =  | 
|
471  | 
alg_defl.basis_fun (\<lambda>a.  | 
|
472  | 
alg_defl_principal (  | 
|
473  | 
Abs_fin_defl (udom_emb oo f\<cdot>(Rep_fin_defl a) oo udom_prj)))"  | 
|
474  | 
||
475  | 
definition  | 
|
476  | 
TypeRep_fun2 ::  | 
|
477  | 
    "((udom \<rightarrow> udom) \<rightarrow> (udom \<rightarrow> udom) \<rightarrow> ('a \<rightarrow> 'a))
 | 
|
478  | 
\<Rightarrow> (TypeRep \<rightarrow> TypeRep \<rightarrow> TypeRep)"  | 
|
479  | 
where  | 
|
480  | 
"TypeRep_fun2 f =  | 
|
481  | 
alg_defl.basis_fun (\<lambda>a.  | 
|
482  | 
alg_defl.basis_fun (\<lambda>b.  | 
|
483  | 
alg_defl_principal (  | 
|
484  | 
Abs_fin_defl (udom_emb oo  | 
|
485  | 
f\<cdot>(Rep_fin_defl a)\<cdot>(Rep_fin_defl b) oo udom_prj))))"  | 
|
486  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
487  | 
definition "cfun_defl = TypeRep_fun2 cfun_map"  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
488  | 
definition "ssum_defl = TypeRep_fun2 ssum_map"  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
489  | 
definition "sprod_defl = TypeRep_fun2 sprod_map"  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
490  | 
definition "cprod_defl = TypeRep_fun2 cprod_map"  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
491  | 
definition "u_defl = TypeRep_fun1 u_map"  | 
| 33588 | 492  | 
|
493  | 
lemma Rep_fin_defl_mono: "a \<sqsubseteq> b \<Longrightarrow> Rep_fin_defl a \<sqsubseteq> Rep_fin_defl b"  | 
|
494  | 
unfolding below_fin_defl_def .  | 
|
495  | 
||
496  | 
lemma cast_TypeRep_fun1:  | 
|
497  | 
assumes f: "\<And>a. finite_deflation a \<Longrightarrow> finite_deflation (f\<cdot>a)"  | 
|
498  | 
shows "cast\<cdot>(TypeRep_fun1 f\<cdot>A) = udom_emb oo f\<cdot>(cast\<cdot>A) oo udom_prj"  | 
|
499  | 
proof -  | 
|
500  | 
have 1: "\<And>a. finite_deflation (udom_emb oo f\<cdot>(Rep_fin_defl a) oo udom_prj)"  | 
|
501  | 
apply (rule ep_pair.finite_deflation_e_d_p [OF ep_pair_udom])  | 
|
502  | 
apply (rule f, rule finite_deflation_Rep_fin_defl)  | 
|
503  | 
done  | 
|
504  | 
show ?thesis  | 
|
505  | 
by (induct A rule: alg_defl.principal_induct, simp)  | 
|
506  | 
(simp only: TypeRep_fun1_def  | 
|
507  | 
alg_defl.basis_fun_principal  | 
|
508  | 
alg_defl.basis_fun_mono  | 
|
509  | 
alg_defl.principal_mono  | 
|
510  | 
Abs_fin_defl_mono [OF 1 1]  | 
|
511  | 
monofun_cfun below_refl  | 
|
512  | 
Rep_fin_defl_mono  | 
|
513  | 
cast_alg_defl_principal  | 
|
514  | 
Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1])  | 
|
515  | 
qed  | 
|
516  | 
||
517  | 
lemma cast_TypeRep_fun2:  | 
|
518  | 
assumes f: "\<And>a b. finite_deflation a \<Longrightarrow> finite_deflation b \<Longrightarrow>  | 
|
519  | 
finite_deflation (f\<cdot>a\<cdot>b)"  | 
|
| 
35901
 
12f09bf2c77f
fix LaTeX overfull hbox warnings in HOLCF document
 
huffman 
parents: 
35900 
diff
changeset
 | 
520  | 
shows "cast\<cdot>(TypeRep_fun2 f\<cdot>A\<cdot>B) =  | 
| 
 
12f09bf2c77f
fix LaTeX overfull hbox warnings in HOLCF document
 
huffman 
parents: 
35900 
diff
changeset
 | 
521  | 
udom_emb oo f\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"  | 
| 33588 | 522  | 
proof -  | 
523  | 
have 1: "\<And>a b. finite_deflation  | 
|
524  | 
(udom_emb oo f\<cdot>(Rep_fin_defl a)\<cdot>(Rep_fin_defl b) oo udom_prj)"  | 
|
525  | 
apply (rule ep_pair.finite_deflation_e_d_p [OF ep_pair_udom])  | 
|
526  | 
apply (rule f, (rule finite_deflation_Rep_fin_defl)+)  | 
|
527  | 
done  | 
|
528  | 
show ?thesis  | 
|
529  | 
by (induct A B rule: alg_defl.principal_induct2, simp, simp)  | 
|
530  | 
(simp only: TypeRep_fun2_def  | 
|
531  | 
alg_defl.basis_fun_principal  | 
|
532  | 
alg_defl.basis_fun_mono  | 
|
533  | 
alg_defl.principal_mono  | 
|
534  | 
Abs_fin_defl_mono [OF 1 1]  | 
|
535  | 
monofun_cfun below_refl  | 
|
536  | 
Rep_fin_defl_mono  | 
|
537  | 
cast_alg_defl_principal  | 
|
538  | 
Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1])  | 
|
539  | 
qed  | 
|
540  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
541  | 
lemma cast_cfun_defl:  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
542  | 
"cast\<cdot>(cfun_defl\<cdot>A\<cdot>B) = udom_emb oo cfun_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
543  | 
unfolding cfun_defl_def  | 
| 33588 | 544  | 
apply (rule cast_TypeRep_fun2)  | 
545  | 
apply (erule (1) finite_deflation_cfun_map)  | 
|
546  | 
done  | 
|
547  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
548  | 
lemma cast_ssum_defl:  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
549  | 
"cast\<cdot>(ssum_defl\<cdot>A\<cdot>B) = udom_emb oo ssum_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
550  | 
unfolding ssum_defl_def  | 
| 33588 | 551  | 
apply (rule cast_TypeRep_fun2)  | 
552  | 
apply (erule (1) finite_deflation_ssum_map)  | 
|
553  | 
done  | 
|
554  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
555  | 
lemma cast_sprod_defl:  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
556  | 
"cast\<cdot>(sprod_defl\<cdot>A\<cdot>B) = udom_emb oo sprod_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
557  | 
unfolding sprod_defl_def  | 
| 33588 | 558  | 
apply (rule cast_TypeRep_fun2)  | 
559  | 
apply (erule (1) finite_deflation_sprod_map)  | 
|
560  | 
done  | 
|
561  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
562  | 
lemma cast_cprod_defl:  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
563  | 
"cast\<cdot>(cprod_defl\<cdot>A\<cdot>B) = udom_emb oo cprod_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
564  | 
unfolding cprod_defl_def  | 
| 33588 | 565  | 
apply (rule cast_TypeRep_fun2)  | 
566  | 
apply (erule (1) finite_deflation_cprod_map)  | 
|
567  | 
done  | 
|
568  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
569  | 
lemma cast_u_defl:  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
570  | 
"cast\<cdot>(u_defl\<cdot>A) = udom_emb oo u_map\<cdot>(cast\<cdot>A) oo udom_prj"  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
571  | 
unfolding u_defl_def  | 
| 33588 | 572  | 
apply (rule cast_TypeRep_fun1)  | 
573  | 
apply (erule finite_deflation_u_map)  | 
|
574  | 
done  | 
|
575  | 
||
576  | 
text {* REP of type constructor = type combinator *}
 | 
|
577  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
578  | 
lemma REP_cfun: "REP('a \<rightarrow> 'b) = cfun_defl\<cdot>REP('a)\<cdot>REP('b)"
 | 
| 33588 | 579  | 
apply (rule cast_eq_imp_eq, rule ext_cfun)  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
580  | 
apply (simp add: cast_REP cast_cfun_defl)  | 
| 33588 | 581  | 
apply (simp add: cfun_map_def)  | 
582  | 
apply (simp only: prj_cfun_def emb_cfun_def)  | 
|
583  | 
apply (simp add: expand_cfun_eq ep_pair.e_eq_iff [OF ep_pair_udom])  | 
|
584  | 
done  | 
|
585  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
586  | 
lemma REP_ssum: "REP('a \<oplus> 'b) = ssum_defl\<cdot>REP('a)\<cdot>REP('b)"
 | 
| 33588 | 587  | 
apply (rule cast_eq_imp_eq, rule ext_cfun)  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
588  | 
apply (simp add: cast_REP cast_ssum_defl)  | 
| 33588 | 589  | 
apply (simp add: prj_ssum_def)  | 
590  | 
apply (simp add: emb_ssum_def)  | 
|
591  | 
apply (simp add: ssum_map_map cfcomp1)  | 
|
592  | 
done  | 
|
593  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
594  | 
lemma REP_sprod: "REP('a \<otimes> 'b) = sprod_defl\<cdot>REP('a)\<cdot>REP('b)"
 | 
| 33588 | 595  | 
apply (rule cast_eq_imp_eq, rule ext_cfun)  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
596  | 
apply (simp add: cast_REP cast_sprod_defl)  | 
| 33588 | 597  | 
apply (simp add: prj_sprod_def)  | 
598  | 
apply (simp add: emb_sprod_def)  | 
|
599  | 
apply (simp add: sprod_map_map cfcomp1)  | 
|
600  | 
done  | 
|
601  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
602  | 
lemma REP_cprod: "REP('a \<times> 'b) = cprod_defl\<cdot>REP('a)\<cdot>REP('b)"
 | 
| 33588 | 603  | 
apply (rule cast_eq_imp_eq, rule ext_cfun)  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
604  | 
apply (simp add: cast_REP cast_cprod_defl)  | 
| 33588 | 605  | 
apply (simp add: prj_cprod_def)  | 
606  | 
apply (simp add: emb_cprod_def)  | 
|
607  | 
apply (simp add: cprod_map_map cfcomp1)  | 
|
608  | 
done  | 
|
609  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
610  | 
lemma REP_up: "REP('a u) = u_defl\<cdot>REP('a)"
 | 
| 33588 | 611  | 
apply (rule cast_eq_imp_eq, rule ext_cfun)  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
612  | 
apply (simp add: cast_REP cast_u_defl)  | 
| 33588 | 613  | 
apply (simp add: prj_u_def)  | 
614  | 
apply (simp add: emb_u_def)  | 
|
615  | 
apply (simp add: u_map_map cfcomp1)  | 
|
616  | 
done  | 
|
617  | 
||
618  | 
lemmas REP_simps =  | 
|
619  | 
REP_cfun  | 
|
620  | 
REP_ssum  | 
|
621  | 
REP_sprod  | 
|
622  | 
REP_cprod  | 
|
623  | 
REP_up  | 
|
624  | 
||
625  | 
subsection {* Isomorphic deflations *}
 | 
|
626  | 
||
627  | 
definition  | 
|
628  | 
  isodefl :: "('a::rep \<rightarrow> 'a) \<Rightarrow> udom alg_defl \<Rightarrow> bool"
 | 
|
629  | 
where  | 
|
630  | 
"isodefl d t \<longleftrightarrow> cast\<cdot>t = emb oo d oo prj"  | 
|
631  | 
||
632  | 
lemma isodeflI: "(\<And>x. cast\<cdot>t\<cdot>x = emb\<cdot>(d\<cdot>(prj\<cdot>x))) \<Longrightarrow> isodefl d t"  | 
|
633  | 
unfolding isodefl_def by (simp add: ext_cfun)  | 
|
634  | 
||
635  | 
lemma cast_isodefl: "isodefl d t \<Longrightarrow> cast\<cdot>t = (\<Lambda> x. emb\<cdot>(d\<cdot>(prj\<cdot>x)))"  | 
|
636  | 
unfolding isodefl_def by (simp add: ext_cfun)  | 
|
637  | 
||
638  | 
lemma isodefl_strict: "isodefl d t \<Longrightarrow> d\<cdot>\<bottom> = \<bottom>"  | 
|
639  | 
unfolding isodefl_def  | 
|
640  | 
by (drule cfun_fun_cong [where x="\<bottom>"], simp)  | 
|
641  | 
||
642  | 
lemma isodefl_imp_deflation:  | 
|
643  | 
fixes d :: "'a::rep \<rightarrow> 'a"  | 
|
644  | 
assumes "isodefl d t" shows "deflation d"  | 
|
645  | 
proof  | 
|
646  | 
note prems [unfolded isodefl_def, simp]  | 
|
647  | 
fix x :: 'a  | 
|
648  | 
show "d\<cdot>(d\<cdot>x) = d\<cdot>x"  | 
|
649  | 
using cast.idem [of t "emb\<cdot>x"] by simp  | 
|
650  | 
show "d\<cdot>x \<sqsubseteq> x"  | 
|
651  | 
using cast.below [of t "emb\<cdot>x"] by simp  | 
|
652  | 
qed  | 
|
653  | 
||
654  | 
lemma isodefl_ID_REP: "isodefl (ID :: 'a \<rightarrow> 'a) REP('a)"
 | 
|
655  | 
unfolding isodefl_def by (simp add: cast_REP)  | 
|
656  | 
||
657  | 
lemma isodefl_REP_imp_ID: "isodefl (d :: 'a \<rightarrow> 'a) REP('a) \<Longrightarrow> d = ID"
 | 
|
658  | 
unfolding isodefl_def  | 
|
659  | 
apply (simp add: cast_REP)  | 
|
660  | 
apply (simp add: expand_cfun_eq)  | 
|
661  | 
apply (rule allI)  | 
|
662  | 
apply (drule_tac x="emb\<cdot>x" in spec)  | 
|
663  | 
apply simp  | 
|
664  | 
done  | 
|
665  | 
||
666  | 
lemma isodefl_bottom: "isodefl \<bottom> \<bottom>"  | 
|
667  | 
unfolding isodefl_def by (simp add: expand_cfun_eq)  | 
|
668  | 
||
669  | 
lemma adm_isodefl:  | 
|
670  | 
"cont f \<Longrightarrow> cont g \<Longrightarrow> adm (\<lambda>x. isodefl (f x) (g x))"  | 
|
671  | 
unfolding isodefl_def by simp  | 
|
672  | 
||
673  | 
lemma isodefl_lub:  | 
|
674  | 
assumes "chain d" and "chain t"  | 
|
675  | 
assumes "\<And>i. isodefl (d i) (t i)"  | 
|
676  | 
shows "isodefl (\<Squnion>i. d i) (\<Squnion>i. t i)"  | 
|
677  | 
using prems unfolding isodefl_def  | 
|
678  | 
by (simp add: contlub_cfun_arg contlub_cfun_fun)  | 
|
679  | 
||
680  | 
lemma isodefl_fix:  | 
|
681  | 
assumes "\<And>d t. isodefl d t \<Longrightarrow> isodefl (f\<cdot>d) (g\<cdot>t)"  | 
|
682  | 
shows "isodefl (fix\<cdot>f) (fix\<cdot>g)"  | 
|
683  | 
unfolding fix_def2  | 
|
684  | 
apply (rule isodefl_lub, simp, simp)  | 
|
685  | 
apply (induct_tac i)  | 
|
686  | 
apply (simp add: isodefl_bottom)  | 
|
687  | 
apply (simp add: prems)  | 
|
688  | 
done  | 
|
689  | 
||
690  | 
lemma isodefl_coerce:  | 
|
691  | 
fixes d :: "'a \<rightarrow> 'a"  | 
|
692  | 
  assumes REP: "REP('b) = REP('a)"
 | 
|
693  | 
shows "isodefl d t \<Longrightarrow> isodefl (coerce oo d oo coerce :: 'b \<rightarrow> 'b) t"  | 
|
694  | 
unfolding isodefl_def  | 
|
695  | 
apply (simp add: expand_cfun_eq)  | 
|
696  | 
apply (simp add: emb_coerce coerce_prj REP)  | 
|
697  | 
done  | 
|
698  | 
||
| 33779 | 699  | 
lemma isodefl_abs_rep:  | 
700  | 
fixes abs and rep and d  | 
|
701  | 
  assumes REP: "REP('b) = REP('a)"
 | 
|
702  | 
assumes abs_def: "abs \<equiv> (coerce :: 'a \<rightarrow> 'b)"  | 
|
703  | 
assumes rep_def: "rep \<equiv> (coerce :: 'b \<rightarrow> 'a)"  | 
|
704  | 
shows "isodefl d t \<Longrightarrow> isodefl (abs oo d oo rep) t"  | 
|
705  | 
unfolding abs_def rep_def using REP by (rule isodefl_coerce)  | 
|
706  | 
||
| 33588 | 707  | 
lemma isodefl_cfun:  | 
708  | 
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
709  | 
isodefl (cfun_map\<cdot>d1\<cdot>d2) (cfun_defl\<cdot>t1\<cdot>t2)"  | 
| 33588 | 710  | 
apply (rule isodeflI)  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
711  | 
apply (simp add: cast_cfun_defl cast_isodefl)  | 
| 33588 | 712  | 
apply (simp add: emb_cfun_def prj_cfun_def)  | 
713  | 
apply (simp add: cfun_map_map cfcomp1)  | 
|
714  | 
done  | 
|
715  | 
||
716  | 
lemma isodefl_ssum:  | 
|
717  | 
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
718  | 
isodefl (ssum_map\<cdot>d1\<cdot>d2) (ssum_defl\<cdot>t1\<cdot>t2)"  | 
| 33588 | 719  | 
apply (rule isodeflI)  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
720  | 
apply (simp add: cast_ssum_defl cast_isodefl)  | 
| 33588 | 721  | 
apply (simp add: emb_ssum_def prj_ssum_def)  | 
722  | 
apply (simp add: ssum_map_map isodefl_strict)  | 
|
723  | 
done  | 
|
724  | 
||
725  | 
lemma isodefl_sprod:  | 
|
726  | 
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
727  | 
isodefl (sprod_map\<cdot>d1\<cdot>d2) (sprod_defl\<cdot>t1\<cdot>t2)"  | 
| 33588 | 728  | 
apply (rule isodeflI)  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
729  | 
apply (simp add: cast_sprod_defl cast_isodefl)  | 
| 33588 | 730  | 
apply (simp add: emb_sprod_def prj_sprod_def)  | 
731  | 
apply (simp add: sprod_map_map isodefl_strict)  | 
|
732  | 
done  | 
|
733  | 
||
| 33786 | 734  | 
lemma isodefl_cprod:  | 
735  | 
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>  | 
|
736  | 
isodefl (cprod_map\<cdot>d1\<cdot>d2) (cprod_defl\<cdot>t1\<cdot>t2)"  | 
|
737  | 
apply (rule isodeflI)  | 
|
738  | 
apply (simp add: cast_cprod_defl cast_isodefl)  | 
|
739  | 
apply (simp add: emb_cprod_def prj_cprod_def)  | 
|
740  | 
apply (simp add: cprod_map_map cfcomp1)  | 
|
741  | 
done  | 
|
742  | 
||
| 33588 | 743  | 
lemma isodefl_u:  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
744  | 
"isodefl d t \<Longrightarrow> isodefl (u_map\<cdot>d) (u_defl\<cdot>t)"  | 
| 33588 | 745  | 
apply (rule isodeflI)  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33779 
diff
changeset
 | 
746  | 
apply (simp add: cast_u_defl cast_isodefl)  | 
| 33588 | 747  | 
apply (simp add: emb_u_def prj_u_def)  | 
748  | 
apply (simp add: u_map_map)  | 
|
749  | 
done  | 
|
750  | 
||
| 
33794
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
751  | 
subsection {* Constructing Domain Isomorphisms *}
 | 
| 
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
752  | 
|
| 
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
753  | 
use "Tools/Domain/domain_isomorphism.ML"  | 
| 
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
754  | 
|
| 
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
755  | 
setup {*
 | 
| 
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
756  | 
fold Domain_Isomorphism.add_type_constructor  | 
| 35525 | 757  | 
    [(@{type_name cfun}, @{term cfun_defl}, @{const_name cfun_map}, @{thm REP_cfun},
 | 
| 
35479
 
dffffe36344a
store deflation thms for map functions in theory data
 
huffman 
parents: 
35475 
diff
changeset
 | 
758  | 
        @{thm isodefl_cfun}, @{thm cfun_map_ID}, @{thm deflation_cfun_map}),
 | 
| 
33794
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
759  | 
|
| 35525 | 760  | 
     (@{type_name ssum}, @{term ssum_defl}, @{const_name ssum_map}, @{thm REP_ssum},
 | 
| 
35479
 
dffffe36344a
store deflation thms for map functions in theory data
 
huffman 
parents: 
35475 
diff
changeset
 | 
761  | 
        @{thm isodefl_ssum}, @{thm ssum_map_ID}, @{thm deflation_ssum_map}),
 | 
| 
33794
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
762  | 
|
| 35525 | 763  | 
     (@{type_name sprod}, @{term sprod_defl}, @{const_name sprod_map}, @{thm REP_sprod},
 | 
| 
35479
 
dffffe36344a
store deflation thms for map functions in theory data
 
huffman 
parents: 
35475 
diff
changeset
 | 
764  | 
        @{thm isodefl_sprod}, @{thm sprod_map_ID}, @{thm deflation_sprod_map}),
 | 
| 
33794
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
765  | 
|
| 
37678
 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 
haftmann 
parents: 
36452 
diff
changeset
 | 
766  | 
     (@{type_name prod}, @{term cprod_defl}, @{const_name cprod_map}, @{thm REP_cprod},
 | 
| 
35479
 
dffffe36344a
store deflation thms for map functions in theory data
 
huffman 
parents: 
35475 
diff
changeset
 | 
767  | 
        @{thm isodefl_cprod}, @{thm cprod_map_ID}, @{thm deflation_cprod_map}),
 | 
| 
33794
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
768  | 
|
| 
35479
 
dffffe36344a
store deflation thms for map functions in theory data
 
huffman 
parents: 
35475 
diff
changeset
 | 
769  | 
     (@{type_name "u"}, @{term u_defl}, @{const_name u_map}, @{thm REP_up},
 | 
| 
 
dffffe36344a
store deflation thms for map functions in theory data
 
huffman 
parents: 
35475 
diff
changeset
 | 
770  | 
        @{thm isodefl_u}, @{thm u_map_ID}, @{thm deflation_u_map})]
 | 
| 
33794
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
771  | 
*}  | 
| 
 
364bc92ba081
set up domain_isomorphism package in Representable.thy
 
huffman 
parents: 
33786 
diff
changeset
 | 
772  | 
|
| 33588 | 773  | 
end  |