| author | wenzelm |
| Mon, 25 Oct 2010 21:23:09 +0200 | |
| changeset 40133 | b61d52de66f0 |
| parent 40038 | 9d061b3d8f46 |
| child 40216 | 366309dfaf60 |
| 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 |
|
|
39987
8c2f449af35a
move all bifinite class instances to Bifinite.thy
huffman
parents:
39986
diff
changeset
|
8 |
imports Algebraic Bifinite Universal Ssum 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 |
||
| 39986 | 14 |
default_sort bifinite |
| 33588 | 15 |
|
16 |
subsection {* Representations of types *}
|
|
17 |
||
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
18 |
lemma emb_prj: "emb\<cdot>((prj\<cdot>x)::'a) = cast\<cdot>DEFL('a)\<cdot>x"
|
|
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
19 |
by (simp add: cast_DEFL) |
| 33588 | 20 |
|
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
21 |
lemma emb_prj_emb: |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
22 |
fixes x :: "'a" |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
23 |
assumes "DEFL('a) \<sqsubseteq> DEFL('b)"
|
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
24 |
shows "emb\<cdot>(prj\<cdot>(emb\<cdot>x) :: 'b) = emb\<cdot>x" |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
25 |
unfolding emb_prj |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
26 |
apply (rule cast.belowD) |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
27 |
apply (rule monofun_cfun_arg [OF assms]) |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
28 |
apply (simp add: cast_DEFL) |
| 33588 | 29 |
done |
30 |
||
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
31 |
lemma prj_emb_prj: |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
32 |
assumes "DEFL('a) \<sqsubseteq> DEFL('b)"
|
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
33 |
shows "prj\<cdot>(emb\<cdot>(prj\<cdot>x :: 'b)) = (prj\<cdot>x :: 'a)" |
| 33588 | 34 |
apply (rule emb_eq_iff [THEN iffD1]) |
35 |
apply (simp only: emb_prj) |
|
36 |
apply (rule deflation_below_comp1) |
|
37 |
apply (rule deflation_cast) |
|
38 |
apply (rule deflation_cast) |
|
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
39 |
apply (rule monofun_cfun_arg [OF assms]) |
| 33588 | 40 |
done |
41 |
||
| 33779 | 42 |
text {* Isomorphism lemmas used internally by the domain package: *}
|
43 |
||
44 |
lemma domain_abs_iso: |
|
45 |
fixes abs and rep |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
46 |
assumes DEFL: "DEFL('b) = DEFL('a)"
|
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
47 |
assumes abs_def: "(abs :: 'a \<rightarrow> 'b) \<equiv> prj oo emb" |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
48 |
assumes rep_def: "(rep :: 'b \<rightarrow> 'a) \<equiv> prj oo emb" |
| 33779 | 49 |
shows "rep\<cdot>(abs\<cdot>x) = x" |
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
50 |
unfolding abs_def rep_def |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
51 |
by (simp add: emb_prj_emb DEFL) |
| 33779 | 52 |
|
53 |
lemma domain_rep_iso: |
|
54 |
fixes abs and rep |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
55 |
assumes DEFL: "DEFL('b) = DEFL('a)"
|
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
56 |
assumes abs_def: "(abs :: 'a \<rightarrow> 'b) \<equiv> prj oo emb" |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
57 |
assumes rep_def: "(rep :: 'b \<rightarrow> 'a) \<equiv> prj oo emb" |
| 33779 | 58 |
shows "abs\<cdot>(rep\<cdot>x) = x" |
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
59 |
unfolding abs_def rep_def |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
60 |
by (simp add: emb_prj_emb DEFL) |
| 33779 | 61 |
|
|
40038
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
62 |
subsection {* Deflations as sets *}
|
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
63 |
|
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
64 |
definition defl_set :: "defl \<Rightarrow> udom set" |
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
65 |
where "defl_set A = {x. cast\<cdot>A\<cdot>x = x}"
|
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
66 |
|
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
67 |
lemma adm_defl_set: "adm (\<lambda>x. x \<in> defl_set A)" |
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
68 |
unfolding defl_set_def by simp |
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
69 |
|
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
70 |
lemma defl_set_bottom: "\<bottom> \<in> defl_set A" |
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
71 |
unfolding defl_set_def by simp |
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
72 |
|
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
73 |
lemma defl_set_cast [simp]: "cast\<cdot>A\<cdot>x \<in> defl_set A" |
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
74 |
unfolding defl_set_def by simp |
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
75 |
|
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
76 |
lemma defl_set_subset_iff: "defl_set A \<subseteq> defl_set B \<longleftrightarrow> A \<sqsubseteq> B" |
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
77 |
apply (simp add: defl_set_def subset_eq cast_below_cast [symmetric]) |
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
78 |
apply (auto simp add: cast.belowI cast.belowD) |
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
79 |
done |
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
80 |
|
| 33588 | 81 |
subsection {* Proving a subtype is representable *}
|
82 |
||
83 |
text {*
|
|
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
84 |
Temporarily relax type constraints for @{term emb} and @{term prj}.
|
| 33588 | 85 |
*} |
86 |
||
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
87 |
setup {*
|
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
88 |
fold Sign.add_const_constraint |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
89 |
[ (@{const_name defl}, SOME @{typ "'a::pcpo itself \<Rightarrow> defl"})
|
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
90 |
, (@{const_name emb}, SOME @{typ "'a::pcpo \<rightarrow> udom"})
|
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
91 |
, (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::pcpo"}) ]
|
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
92 |
*} |
| 33588 | 93 |
|
94 |
lemma typedef_rep_class: |
|
95 |
fixes Rep :: "'a::pcpo \<Rightarrow> udom" |
|
96 |
fixes Abs :: "udom \<Rightarrow> 'a::pcpo" |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
97 |
fixes t :: defl |
|
40038
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
98 |
assumes type: "type_definition Rep Abs (defl_set t)" |
| 33588 | 99 |
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
|
100 |
assumes emb: "emb \<equiv> (\<Lambda> x. Rep x)" |
|
331712879666
automate definition of representable domains from algebraic deflations
huffman
parents:
33589
diff
changeset
|
101 |
assumes prj: "prj \<equiv> (\<Lambda> x. Abs (cast\<cdot>t\<cdot>x))" |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
102 |
assumes defl: "defl \<equiv> (\<lambda> a::'a itself. t)" |
| 39986 | 103 |
shows "OFCLASS('a, bifinite_class)"
|
| 33588 | 104 |
proof |
105 |
have emb_beta: "\<And>x. emb\<cdot>x = Rep x" |
|
106 |
unfolding emb |
|
107 |
apply (rule beta_cfun) |
|
|
40038
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
108 |
apply (rule typedef_cont_Rep [OF type below adm_defl_set]) |
| 33588 | 109 |
done |
110 |
have prj_beta: "\<And>y. prj\<cdot>y = Abs (cast\<cdot>t\<cdot>y)" |
|
111 |
unfolding prj |
|
112 |
apply (rule beta_cfun) |
|
|
40038
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
113 |
apply (rule typedef_cont_Abs [OF type below adm_defl_set]) |
| 33588 | 114 |
apply simp_all |
115 |
done |
|
|
40038
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
116 |
have prj_emb: "\<And>x::'a. prj\<cdot>(emb\<cdot>x) = x" |
| 33588 | 117 |
using type_definition.Rep [OF type] |
|
40038
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
118 |
unfolding prj_beta emb_beta defl_set_def |
|
9d061b3d8f46
replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents:
40037
diff
changeset
|
119 |
by (simp add: type_definition.Rep_inverse [OF type]) |
| 33588 | 120 |
have emb_prj: "\<And>y. emb\<cdot>(prj\<cdot>y :: 'a) = cast\<cdot>t\<cdot>y" |
121 |
unfolding prj_beta emb_beta |
|
122 |
by (simp add: type_definition.Abs_inverse [OF type]) |
|
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
37770
diff
changeset
|
123 |
show "ep_pair (emb :: 'a \<rightarrow> udom) prj" |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
37770
diff
changeset
|
124 |
apply default |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
37770
diff
changeset
|
125 |
apply (simp add: prj_emb) |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
37770
diff
changeset
|
126 |
apply (simp add: emb_prj cast.below) |
| 33588 | 127 |
done |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
128 |
show "cast\<cdot>DEFL('a) = emb oo (prj :: udom \<rightarrow> 'a)"
|
|
40002
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents:
39989
diff
changeset
|
129 |
by (rule cfun_eqI, simp add: defl emb_prj) |
| 33588 | 130 |
qed |
131 |
||
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
132 |
lemma typedef_DEFL: |
|
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
133 |
assumes "defl \<equiv> (\<lambda>a::'a::pcpo itself. t)" |
|
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
134 |
shows "DEFL('a::pcpo) = t"
|
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
37770
diff
changeset
|
135 |
unfolding assms .. |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
37770
diff
changeset
|
136 |
|
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
137 |
text {* Restore original typing constraints. *}
|
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
37770
diff
changeset
|
138 |
|
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
139 |
setup {*
|
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
140 |
fold Sign.add_const_constraint |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
141 |
[ (@{const_name defl}, SOME @{typ "'a::bifinite itself \<Rightarrow> defl"})
|
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
142 |
, (@{const_name emb}, SOME @{typ "'a::bifinite \<rightarrow> udom"})
|
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
143 |
, (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::bifinite"}) ]
|
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
144 |
*} |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
37770
diff
changeset
|
145 |
|
|
33679
331712879666
automate definition of representable domains from algebraic deflations
huffman
parents:
33589
diff
changeset
|
146 |
use "Tools/repdef.ML" |
|
331712879666
automate definition of representable domains from algebraic deflations
huffman
parents:
33589
diff
changeset
|
147 |
|
| 33588 | 148 |
subsection {* Isomorphic deflations *}
|
149 |
||
150 |
definition |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
151 |
isodefl :: "('a \<rightarrow> 'a) \<Rightarrow> defl \<Rightarrow> bool"
|
| 33588 | 152 |
where |
153 |
"isodefl d t \<longleftrightarrow> cast\<cdot>t = emb oo d oo prj" |
|
154 |
||
155 |
lemma isodeflI: "(\<And>x. cast\<cdot>t\<cdot>x = emb\<cdot>(d\<cdot>(prj\<cdot>x))) \<Longrightarrow> isodefl d t" |
|
|
40002
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents:
39989
diff
changeset
|
156 |
unfolding isodefl_def by (simp add: cfun_eqI) |
| 33588 | 157 |
|
158 |
lemma cast_isodefl: "isodefl d t \<Longrightarrow> cast\<cdot>t = (\<Lambda> x. emb\<cdot>(d\<cdot>(prj\<cdot>x)))" |
|
|
40002
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents:
39989
diff
changeset
|
159 |
unfolding isodefl_def by (simp add: cfun_eqI) |
| 33588 | 160 |
|
161 |
lemma isodefl_strict: "isodefl d t \<Longrightarrow> d\<cdot>\<bottom> = \<bottom>" |
|
162 |
unfolding isodefl_def |
|
163 |
by (drule cfun_fun_cong [where x="\<bottom>"], simp) |
|
164 |
||
165 |
lemma isodefl_imp_deflation: |
|
| 39986 | 166 |
fixes d :: "'a \<rightarrow> 'a" |
| 33588 | 167 |
assumes "isodefl d t" shows "deflation d" |
168 |
proof |
|
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
37770
diff
changeset
|
169 |
note assms [unfolded isodefl_def, simp] |
| 33588 | 170 |
fix x :: 'a |
171 |
show "d\<cdot>(d\<cdot>x) = d\<cdot>x" |
|
172 |
using cast.idem [of t "emb\<cdot>x"] by simp |
|
173 |
show "d\<cdot>x \<sqsubseteq> x" |
|
174 |
using cast.below [of t "emb\<cdot>x"] by simp |
|
175 |
qed |
|
176 |
||
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
177 |
lemma isodefl_ID_DEFL: "isodefl (ID :: 'a \<rightarrow> 'a) DEFL('a)"
|
|
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
178 |
unfolding isodefl_def by (simp add: cast_DEFL) |
| 33588 | 179 |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
180 |
lemma isodefl_DEFL_imp_ID: "isodefl (d :: 'a \<rightarrow> 'a) DEFL('a) \<Longrightarrow> d = ID"
|
| 33588 | 181 |
unfolding isodefl_def |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
182 |
apply (simp add: cast_DEFL) |
|
40002
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents:
39989
diff
changeset
|
183 |
apply (simp add: cfun_eq_iff) |
| 33588 | 184 |
apply (rule allI) |
185 |
apply (drule_tac x="emb\<cdot>x" in spec) |
|
186 |
apply simp |
|
187 |
done |
|
188 |
||
189 |
lemma isodefl_bottom: "isodefl \<bottom> \<bottom>" |
|
|
40002
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents:
39989
diff
changeset
|
190 |
unfolding isodefl_def by (simp add: cfun_eq_iff) |
| 33588 | 191 |
|
192 |
lemma adm_isodefl: |
|
193 |
"cont f \<Longrightarrow> cont g \<Longrightarrow> adm (\<lambda>x. isodefl (f x) (g x))" |
|
194 |
unfolding isodefl_def by simp |
|
195 |
||
196 |
lemma isodefl_lub: |
|
197 |
assumes "chain d" and "chain t" |
|
198 |
assumes "\<And>i. isodefl (d i) (t i)" |
|
199 |
shows "isodefl (\<Squnion>i. d i) (\<Squnion>i. t i)" |
|
200 |
using prems unfolding isodefl_def |
|
201 |
by (simp add: contlub_cfun_arg contlub_cfun_fun) |
|
202 |
||
203 |
lemma isodefl_fix: |
|
204 |
assumes "\<And>d t. isodefl d t \<Longrightarrow> isodefl (f\<cdot>d) (g\<cdot>t)" |
|
205 |
shows "isodefl (fix\<cdot>f) (fix\<cdot>g)" |
|
206 |
unfolding fix_def2 |
|
207 |
apply (rule isodefl_lub, simp, simp) |
|
208 |
apply (induct_tac i) |
|
209 |
apply (simp add: isodefl_bottom) |
|
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
37770
diff
changeset
|
210 |
apply (simp add: assms) |
| 33588 | 211 |
done |
212 |
||
| 33779 | 213 |
lemma isodefl_abs_rep: |
214 |
fixes abs and rep and d |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
215 |
assumes DEFL: "DEFL('b) = DEFL('a)"
|
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
216 |
assumes abs_def: "(abs :: 'a \<rightarrow> 'b) \<equiv> prj oo emb" |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
217 |
assumes rep_def: "(rep :: 'b \<rightarrow> 'a) \<equiv> prj oo emb" |
| 33779 | 218 |
shows "isodefl d t \<Longrightarrow> isodefl (abs oo d oo rep) t" |
|
40037
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
219 |
unfolding isodefl_def |
|
81e6b89d8f58
eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents:
40002
diff
changeset
|
220 |
by (simp add: cfun_eq_iff assms prj_emb_prj emb_prj_emb) |
| 33779 | 221 |
|
| 33588 | 222 |
lemma isodefl_cfun: |
223 |
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow> |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
224 |
isodefl (cfun_map\<cdot>d1\<cdot>d2) (cfun_defl\<cdot>t1\<cdot>t2)" |
| 33588 | 225 |
apply (rule isodeflI) |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
226 |
apply (simp add: cast_cfun_defl cast_isodefl) |
| 33588 | 227 |
apply (simp add: emb_cfun_def prj_cfun_def) |
228 |
apply (simp add: cfun_map_map cfcomp1) |
|
229 |
done |
|
230 |
||
231 |
lemma isodefl_ssum: |
|
232 |
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow> |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
233 |
isodefl (ssum_map\<cdot>d1\<cdot>d2) (ssum_defl\<cdot>t1\<cdot>t2)" |
| 33588 | 234 |
apply (rule isodeflI) |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
235 |
apply (simp add: cast_ssum_defl cast_isodefl) |
| 33588 | 236 |
apply (simp add: emb_ssum_def prj_ssum_def) |
237 |
apply (simp add: ssum_map_map isodefl_strict) |
|
238 |
done |
|
239 |
||
240 |
lemma isodefl_sprod: |
|
241 |
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow> |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
242 |
isodefl (sprod_map\<cdot>d1\<cdot>d2) (sprod_defl\<cdot>t1\<cdot>t2)" |
| 33588 | 243 |
apply (rule isodeflI) |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
244 |
apply (simp add: cast_sprod_defl cast_isodefl) |
| 33588 | 245 |
apply (simp add: emb_sprod_def prj_sprod_def) |
246 |
apply (simp add: sprod_map_map isodefl_strict) |
|
247 |
done |
|
248 |
||
| 33786 | 249 |
lemma isodefl_cprod: |
250 |
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow> |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
251 |
isodefl (cprod_map\<cdot>d1\<cdot>d2) (prod_defl\<cdot>t1\<cdot>t2)" |
| 33786 | 252 |
apply (rule isodeflI) |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
253 |
apply (simp add: cast_prod_defl cast_isodefl) |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
37770
diff
changeset
|
254 |
apply (simp add: emb_prod_def prj_prod_def) |
| 33786 | 255 |
apply (simp add: cprod_map_map cfcomp1) |
256 |
done |
|
257 |
||
| 33588 | 258 |
lemma isodefl_u: |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
259 |
"isodefl d t \<Longrightarrow> isodefl (u_map\<cdot>d) (u_defl\<cdot>t)" |
| 33588 | 260 |
apply (rule isodeflI) |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
261 |
apply (simp add: cast_u_defl cast_isodefl) |
| 33588 | 262 |
apply (simp add: emb_u_def prj_u_def) |
263 |
apply (simp add: u_map_map) |
|
264 |
done |
|
265 |
||
|
33794
364bc92ba081
set up domain_isomorphism package in Representable.thy
huffman
parents:
33786
diff
changeset
|
266 |
subsection {* Constructing Domain Isomorphisms *}
|
|
364bc92ba081
set up domain_isomorphism package in Representable.thy
huffman
parents:
33786
diff
changeset
|
267 |
|
|
364bc92ba081
set up domain_isomorphism package in Representable.thy
huffman
parents:
33786
diff
changeset
|
268 |
use "Tools/Domain/domain_isomorphism.ML" |
|
364bc92ba081
set up domain_isomorphism package in Representable.thy
huffman
parents:
33786
diff
changeset
|
269 |
|
|
364bc92ba081
set up domain_isomorphism package in Representable.thy
huffman
parents:
33786
diff
changeset
|
270 |
setup {*
|
|
364bc92ba081
set up domain_isomorphism package in Representable.thy
huffman
parents:
33786
diff
changeset
|
271 |
fold Domain_Isomorphism.add_type_constructor |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
272 |
[(@{type_name cfun}, @{term cfun_defl}, @{const_name cfun_map}, @{thm DEFL_cfun},
|
|
35479
dffffe36344a
store deflation thms for map functions in theory data
huffman
parents:
35475
diff
changeset
|
273 |
@{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
|
274 |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
275 |
(@{type_name ssum}, @{term ssum_defl}, @{const_name ssum_map}, @{thm DEFL_ssum},
|
|
35479
dffffe36344a
store deflation thms for map functions in theory data
huffman
parents:
35475
diff
changeset
|
276 |
@{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
|
277 |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
278 |
(@{type_name sprod}, @{term sprod_defl}, @{const_name sprod_map}, @{thm DEFL_sprod},
|
|
35479
dffffe36344a
store deflation thms for map functions in theory data
huffman
parents:
35475
diff
changeset
|
279 |
@{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
|
280 |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
281 |
(@{type_name prod}, @{term cprod_defl}, @{const_name cprod_map}, @{thm DEFL_prod},
|
|
35479
dffffe36344a
store deflation thms for map functions in theory data
huffman
parents:
35475
diff
changeset
|
282 |
@{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
|
283 |
|
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39987
diff
changeset
|
284 |
(@{type_name "u"}, @{term u_defl}, @{const_name u_map}, @{thm DEFL_u},
|
|
35479
dffffe36344a
store deflation thms for map functions in theory data
huffman
parents:
35475
diff
changeset
|
285 |
@{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
|
286 |
*} |
|
364bc92ba081
set up domain_isomorphism package in Representable.thy
huffman
parents:
33786
diff
changeset
|
287 |
|
| 33588 | 288 |
end |