| author | nipkow | 
| Tue, 14 Sep 2010 08:40:22 +0200 | |
| changeset 39314 | aecb239a2bbc | 
| parent 36452 | d37c6eed8117 | 
| child 39974 | b525988432e9 | 
| permissions | -rw-r--r-- | 
| 33591 | 1  | 
(* Title: HOLCF/ex/Domain_Proofs.thy  | 
2  | 
Author: Brian Huffman  | 
|
3  | 
*)  | 
|
4  | 
||
5  | 
header {* Internal domain package proofs done manually *}
 | 
|
6  | 
||
7  | 
theory Domain_Proofs  | 
|
8  | 
imports HOLCF  | 
|
9  | 
begin  | 
|
10  | 
||
| 36452 | 11  | 
default_sort rep  | 
| 33591 | 12  | 
|
13  | 
(*  | 
|
14  | 
||
15  | 
The definitions and proofs below are for the following recursive  | 
|
16  | 
datatypes:  | 
|
17  | 
||
18  | 
domain 'a foo = Foo1 | Foo2 (lazy 'a) (lazy "'a bar")  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
19  | 
and 'a bar = Bar (lazy "'a baz \<rightarrow> tr")  | 
| 
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
20  | 
and 'a baz = Baz (lazy "'a foo convex_pd \<rightarrow> tr")  | 
| 33591 | 21  | 
|
22  | 
*)  | 
|
23  | 
||
24  | 
(********************************************************************)  | 
|
25  | 
||
26  | 
subsection {* Step 1: Define the new type combinators *}
 | 
|
27  | 
||
28  | 
text {* Start with the one-step non-recursive version *}
 | 
|
29  | 
||
30  | 
definition  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
31  | 
foo_bar_baz_deflF ::  | 
| 33591 | 32  | 
"TypeRep \<rightarrow> TypeRep \<times> TypeRep \<times> TypeRep \<rightarrow> TypeRep \<times> TypeRep \<times> TypeRep"  | 
33  | 
where  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
34  | 
"foo_bar_baz_deflF = (\<Lambda> a. Abs_CFun (\<lambda>(t1, t2, t3).  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
35  | 
( ssum_defl\<cdot>REP(one)\<cdot>(sprod_defl\<cdot>(u_defl\<cdot>a)\<cdot>(u_defl\<cdot>t2))  | 
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
36  | 
, u_defl\<cdot>(cfun_defl\<cdot>t3\<cdot>REP(tr))  | 
| 
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
37  | 
, u_defl\<cdot>(cfun_defl\<cdot>(convex_defl\<cdot>t1)\<cdot>REP(tr)))))"  | 
| 33591 | 38  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
39  | 
lemma foo_bar_baz_deflF_beta:  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
40  | 
"foo_bar_baz_deflF\<cdot>a\<cdot>t =  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
41  | 
( ssum_defl\<cdot>REP(one)\<cdot>(sprod_defl\<cdot>(u_defl\<cdot>a)\<cdot>(u_defl\<cdot>(fst (snd t))))  | 
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
42  | 
, u_defl\<cdot>(cfun_defl\<cdot>(snd (snd t))\<cdot>REP(tr))  | 
| 
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
43  | 
, u_defl\<cdot>(cfun_defl\<cdot>(convex_defl\<cdot>(fst t))\<cdot>REP(tr)))"  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
44  | 
unfolding foo_bar_baz_deflF_def  | 
| 
33781
 
c7d32e726bb9
avoid using csplit; define copy functions exactly like the current domain package
 
huffman 
parents: 
33779 
diff
changeset
 | 
45  | 
by (simp add: split_def)  | 
| 33591 | 46  | 
|
47  | 
text {* Individual type combinators are projected from the fixed point. *}
 | 
|
48  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
49  | 
definition foo_defl :: "TypeRep \<rightarrow> TypeRep"  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
50  | 
where "foo_defl = (\<Lambda> a. fst (fix\<cdot>(foo_bar_baz_deflF\<cdot>a)))"  | 
| 33591 | 51  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
52  | 
definition bar_defl :: "TypeRep \<rightarrow> TypeRep"  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
53  | 
where "bar_defl = (\<Lambda> a. fst (snd (fix\<cdot>(foo_bar_baz_deflF\<cdot>a))))"  | 
| 33591 | 54  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
55  | 
definition baz_defl :: "TypeRep \<rightarrow> TypeRep"  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
56  | 
where "baz_defl = (\<Lambda> a. snd (snd (fix\<cdot>(foo_bar_baz_deflF\<cdot>a))))"  | 
| 33591 | 57  | 
|
| 36132 | 58  | 
lemma defl_apply_thms:  | 
59  | 
"foo_defl\<cdot>a = fst (fix\<cdot>(foo_bar_baz_deflF\<cdot>a))"  | 
|
60  | 
"bar_defl\<cdot>a = fst (snd (fix\<cdot>(foo_bar_baz_deflF\<cdot>a)))"  | 
|
61  | 
"baz_defl\<cdot>a = snd (snd (fix\<cdot>(foo_bar_baz_deflF\<cdot>a)))"  | 
|
62  | 
unfolding foo_defl_def bar_defl_def baz_defl_def by simp_all  | 
|
63  | 
||
| 33591 | 64  | 
text {* Unfold rules for each combinator. *}
 | 
65  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
66  | 
lemma foo_defl_unfold:  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
67  | 
"foo_defl\<cdot>a = ssum_defl\<cdot>REP(one)\<cdot>(sprod_defl\<cdot>(u_defl\<cdot>a)\<cdot>(u_defl\<cdot>(bar_defl\<cdot>a)))"  | 
| 36132 | 68  | 
unfolding defl_apply_thms by (subst fix_eq, simp add: foo_bar_baz_deflF_beta)  | 
| 33591 | 69  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
70  | 
lemma bar_defl_unfold: "bar_defl\<cdot>a = u_defl\<cdot>(cfun_defl\<cdot>(baz_defl\<cdot>a)\<cdot>REP(tr))"  | 
| 36132 | 71  | 
unfolding defl_apply_thms by (subst fix_eq, simp add: foo_bar_baz_deflF_beta)  | 
| 33591 | 72  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
73  | 
lemma baz_defl_unfold: "baz_defl\<cdot>a = u_defl\<cdot>(cfun_defl\<cdot>(convex_defl\<cdot>(foo_defl\<cdot>a))\<cdot>REP(tr))"  | 
| 36132 | 74  | 
unfolding defl_apply_thms by (subst fix_eq, simp add: foo_bar_baz_deflF_beta)  | 
| 33591 | 75  | 
|
76  | 
text "The automation for the previous steps will be quite similar to  | 
|
77  | 
how the fixrec package works."  | 
|
78  | 
||
79  | 
(********************************************************************)  | 
|
80  | 
||
81  | 
subsection {* Step 2: Define types, prove class instances *}
 | 
|
82  | 
||
83  | 
text {* Use @{text pcpodef} with the appropriate type combinator. *}
 | 
|
84  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
85  | 
pcpodef (open) 'a foo = "{x. x ::: foo_defl\<cdot>REP('a)}"
 | 
| 33591 | 86  | 
by (simp_all add: adm_in_deflation)  | 
87  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
88  | 
pcpodef (open) 'a bar = "{x. x ::: bar_defl\<cdot>REP('a)}"
 | 
| 33591 | 89  | 
by (simp_all add: adm_in_deflation)  | 
90  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
91  | 
pcpodef (open) 'a baz = "{x. x ::: baz_defl\<cdot>REP('a)}"
 | 
| 33591 | 92  | 
by (simp_all add: adm_in_deflation)  | 
93  | 
||
94  | 
text {* Prove rep instance using lemma @{text typedef_rep_class}. *}
 | 
|
95  | 
||
96  | 
instantiation foo :: (rep) rep  | 
|
97  | 
begin  | 
|
98  | 
||
99  | 
definition emb_foo :: "'a foo \<rightarrow> udom"  | 
|
| 
33679
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33591 
diff
changeset
 | 
100  | 
where "emb_foo \<equiv> (\<Lambda> x. Rep_foo x)"  | 
| 33591 | 101  | 
|
102  | 
definition prj_foo :: "udom \<rightarrow> 'a foo"  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
103  | 
where "prj_foo \<equiv> (\<Lambda> y. Abs_foo (cast\<cdot>(foo_defl\<cdot>REP('a))\<cdot>y))"
 | 
| 33591 | 104  | 
|
105  | 
definition approx_foo :: "nat \<Rightarrow> 'a foo \<rightarrow> 'a foo"  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
106  | 
where "approx_foo \<equiv> repdef_approx Rep_foo Abs_foo (foo_defl\<cdot>REP('a))"
 | 
| 33591 | 107  | 
|
108  | 
instance  | 
|
109  | 
apply (rule typedef_rep_class)  | 
|
110  | 
apply (rule type_definition_foo)  | 
|
111  | 
apply (rule below_foo_def)  | 
|
112  | 
apply (rule emb_foo_def)  | 
|
113  | 
apply (rule prj_foo_def)  | 
|
114  | 
apply (rule approx_foo_def)  | 
|
115  | 
done  | 
|
116  | 
||
117  | 
end  | 
|
118  | 
||
119  | 
instantiation bar :: (rep) rep  | 
|
120  | 
begin  | 
|
121  | 
||
122  | 
definition emb_bar :: "'a bar \<rightarrow> udom"  | 
|
| 
33679
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33591 
diff
changeset
 | 
123  | 
where "emb_bar \<equiv> (\<Lambda> x. Rep_bar x)"  | 
| 33591 | 124  | 
|
125  | 
definition prj_bar :: "udom \<rightarrow> 'a bar"  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
126  | 
where "prj_bar \<equiv> (\<Lambda> y. Abs_bar (cast\<cdot>(bar_defl\<cdot>REP('a))\<cdot>y))"
 | 
| 33591 | 127  | 
|
128  | 
definition approx_bar :: "nat \<Rightarrow> 'a bar \<rightarrow> 'a bar"  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
129  | 
where "approx_bar \<equiv> repdef_approx Rep_bar Abs_bar (bar_defl\<cdot>REP('a))"
 | 
| 33591 | 130  | 
|
131  | 
instance  | 
|
132  | 
apply (rule typedef_rep_class)  | 
|
133  | 
apply (rule type_definition_bar)  | 
|
134  | 
apply (rule below_bar_def)  | 
|
135  | 
apply (rule emb_bar_def)  | 
|
136  | 
apply (rule prj_bar_def)  | 
|
137  | 
apply (rule approx_bar_def)  | 
|
138  | 
done  | 
|
139  | 
||
140  | 
end  | 
|
141  | 
||
142  | 
instantiation baz :: (rep) rep  | 
|
143  | 
begin  | 
|
144  | 
||
145  | 
definition emb_baz :: "'a baz \<rightarrow> udom"  | 
|
| 
33679
 
331712879666
automate definition of representable domains from algebraic deflations
 
huffman 
parents: 
33591 
diff
changeset
 | 
146  | 
where "emb_baz \<equiv> (\<Lambda> x. Rep_baz x)"  | 
| 33591 | 147  | 
|
148  | 
definition prj_baz :: "udom \<rightarrow> 'a baz"  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
149  | 
where "prj_baz \<equiv> (\<Lambda> y. Abs_baz (cast\<cdot>(baz_defl\<cdot>REP('a))\<cdot>y))"
 | 
| 33591 | 150  | 
|
151  | 
definition approx_baz :: "nat \<Rightarrow> 'a baz \<rightarrow> 'a baz"  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
152  | 
where "approx_baz \<equiv> repdef_approx Rep_baz Abs_baz (baz_defl\<cdot>REP('a))"
 | 
| 33591 | 153  | 
|
154  | 
instance  | 
|
155  | 
apply (rule typedef_rep_class)  | 
|
156  | 
apply (rule type_definition_baz)  | 
|
157  | 
apply (rule below_baz_def)  | 
|
158  | 
apply (rule emb_baz_def)  | 
|
159  | 
apply (rule prj_baz_def)  | 
|
160  | 
apply (rule approx_baz_def)  | 
|
161  | 
done  | 
|
162  | 
||
163  | 
end  | 
|
164  | 
||
165  | 
text {* Prove REP rules using lemma @{text typedef_REP}. *}
 | 
|
166  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
167  | 
lemma REP_foo: "REP('a foo) = foo_defl\<cdot>REP('a)"
 | 
| 33591 | 168  | 
apply (rule typedef_REP)  | 
169  | 
apply (rule type_definition_foo)  | 
|
170  | 
apply (rule below_foo_def)  | 
|
171  | 
apply (rule emb_foo_def)  | 
|
172  | 
apply (rule prj_foo_def)  | 
|
173  | 
done  | 
|
174  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
175  | 
lemma REP_bar: "REP('a bar) = bar_defl\<cdot>REP('a)"
 | 
| 33591 | 176  | 
apply (rule typedef_REP)  | 
177  | 
apply (rule type_definition_bar)  | 
|
178  | 
apply (rule below_bar_def)  | 
|
179  | 
apply (rule emb_bar_def)  | 
|
180  | 
apply (rule prj_bar_def)  | 
|
181  | 
done  | 
|
182  | 
||
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
183  | 
lemma REP_baz: "REP('a baz) = baz_defl\<cdot>REP('a)"
 | 
| 33591 | 184  | 
apply (rule typedef_REP)  | 
185  | 
apply (rule type_definition_baz)  | 
|
186  | 
apply (rule below_baz_def)  | 
|
187  | 
apply (rule emb_baz_def)  | 
|
188  | 
apply (rule prj_baz_def)  | 
|
189  | 
done  | 
|
190  | 
||
191  | 
text {* Prove REP equations using type combinator unfold lemmas. *}
 | 
|
192  | 
||
193  | 
lemma REP_foo': "REP('a foo) = REP(one \<oplus> 'a\<^sub>\<bottom> \<otimes> ('a bar)\<^sub>\<bottom>)"
 | 
|
194  | 
unfolding REP_foo REP_bar REP_baz REP_simps  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
195  | 
by (rule foo_defl_unfold)  | 
| 33591 | 196  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
197  | 
lemma REP_bar': "REP('a bar) = REP(('a baz \<rightarrow> tr)\<^sub>\<bottom>)"
 | 
| 33591 | 198  | 
unfolding REP_foo REP_bar REP_baz REP_simps  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
199  | 
by (rule bar_defl_unfold)  | 
| 33591 | 200  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
201  | 
lemma REP_baz': "REP('a baz) = REP(('a foo convex_pd \<rightarrow> tr)\<^sub>\<bottom>)"
 | 
| 35493 | 202  | 
unfolding REP_foo REP_bar REP_baz REP_simps REP_convex  | 
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
203  | 
by (rule baz_defl_unfold)  | 
| 33591 | 204  | 
|
205  | 
(********************************************************************)  | 
|
206  | 
||
207  | 
subsection {* Step 3: Define rep and abs functions *}
 | 
|
208  | 
||
209  | 
text {* Define them all using @{text coerce}! *}
 | 
|
210  | 
||
211  | 
definition foo_rep :: "'a foo \<rightarrow> one \<oplus> ('a\<^sub>\<bottom> \<otimes> ('a bar)\<^sub>\<bottom>)"
 | 
|
| 33779 | 212  | 
where "foo_rep \<equiv> coerce"  | 
| 33591 | 213  | 
|
214  | 
definition foo_abs :: "one \<oplus> ('a\<^sub>\<bottom> \<otimes> ('a bar)\<^sub>\<bottom>) \<rightarrow> 'a foo"
 | 
|
| 33779 | 215  | 
where "foo_abs \<equiv> coerce"  | 
| 33591 | 216  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
217  | 
definition bar_rep :: "'a bar \<rightarrow> ('a baz \<rightarrow> tr)\<^sub>\<bottom>"
 | 
| 33779 | 218  | 
where "bar_rep \<equiv> coerce"  | 
| 33591 | 219  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
220  | 
definition bar_abs :: "('a baz \<rightarrow> tr)\<^sub>\<bottom> \<rightarrow> 'a bar"
 | 
| 33779 | 221  | 
where "bar_abs \<equiv> coerce"  | 
| 33591 | 222  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
223  | 
definition baz_rep :: "'a baz \<rightarrow> ('a foo convex_pd \<rightarrow> tr)\<^sub>\<bottom>"
 | 
| 33779 | 224  | 
where "baz_rep \<equiv> coerce"  | 
| 33591 | 225  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
226  | 
definition baz_abs :: "('a foo convex_pd \<rightarrow> tr)\<^sub>\<bottom> \<rightarrow> 'a baz"
 | 
| 33779 | 227  | 
where "baz_abs \<equiv> coerce"  | 
228  | 
||
229  | 
text {* Prove isomorphism rules. *}
 | 
|
230  | 
||
231  | 
lemma foo_abs_iso: "foo_rep\<cdot>(foo_abs\<cdot>x) = x"  | 
|
232  | 
by (rule domain_abs_iso [OF REP_foo' foo_abs_def foo_rep_def])  | 
|
233  | 
||
234  | 
lemma foo_rep_iso: "foo_abs\<cdot>(foo_rep\<cdot>x) = x"  | 
|
235  | 
by (rule domain_rep_iso [OF REP_foo' foo_abs_def foo_rep_def])  | 
|
236  | 
||
237  | 
lemma bar_abs_iso: "bar_rep\<cdot>(bar_abs\<cdot>x) = x"  | 
|
238  | 
by (rule domain_abs_iso [OF REP_bar' bar_abs_def bar_rep_def])  | 
|
239  | 
||
240  | 
lemma bar_rep_iso: "bar_abs\<cdot>(bar_rep\<cdot>x) = x"  | 
|
241  | 
by (rule domain_rep_iso [OF REP_bar' bar_abs_def bar_rep_def])  | 
|
242  | 
||
243  | 
lemma baz_abs_iso: "baz_rep\<cdot>(baz_abs\<cdot>x) = x"  | 
|
244  | 
by (rule domain_abs_iso [OF REP_baz' baz_abs_def baz_rep_def])  | 
|
245  | 
||
246  | 
lemma baz_rep_iso: "baz_abs\<cdot>(baz_rep\<cdot>x) = x"  | 
|
247  | 
by (rule domain_rep_iso [OF REP_baz' baz_abs_def baz_rep_def])  | 
|
| 33591 | 248  | 
|
249  | 
text {* Prove isodefl rules using @{text isodefl_coerce}. *}
 | 
|
250  | 
||
251  | 
lemma isodefl_foo_abs:  | 
|
252  | 
"isodefl d t \<Longrightarrow> isodefl (foo_abs oo d oo foo_rep) t"  | 
|
| 33779 | 253  | 
by (rule isodefl_abs_rep [OF REP_foo' foo_abs_def foo_rep_def])  | 
| 33591 | 254  | 
|
255  | 
lemma isodefl_bar_abs:  | 
|
256  | 
"isodefl d t \<Longrightarrow> isodefl (bar_abs oo d oo bar_rep) t"  | 
|
| 33779 | 257  | 
by (rule isodefl_abs_rep [OF REP_bar' bar_abs_def bar_rep_def])  | 
| 33591 | 258  | 
|
259  | 
lemma isodefl_baz_abs:  | 
|
260  | 
"isodefl d t \<Longrightarrow> isodefl (baz_abs oo d oo baz_rep) t"  | 
|
| 33779 | 261  | 
by (rule isodefl_abs_rep [OF REP_baz' baz_abs_def baz_rep_def])  | 
| 33591 | 262  | 
|
263  | 
(********************************************************************)  | 
|
264  | 
||
265  | 
subsection {* Step 4: Define map functions, prove isodefl property *}
 | 
|
266  | 
||
267  | 
text {* Start with the one-step non-recursive version. *}
 | 
|
268  | 
||
269  | 
text {* Note that the type of the map function depends on which
 | 
|
270  | 
variables are used in positive and negative positions. *}  | 
|
271  | 
||
272  | 
definition  | 
|
273  | 
foo_bar_baz_mapF ::  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
274  | 
    "('a \<rightarrow> 'b) \<rightarrow>
 | 
| 
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
275  | 
     ('a foo \<rightarrow> 'b foo) \<times> ('a bar \<rightarrow> 'b bar) \<times> ('b baz \<rightarrow> 'a baz) \<rightarrow>
 | 
| 
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
276  | 
     ('a foo \<rightarrow> 'b foo) \<times> ('a bar \<rightarrow> 'b bar) \<times> ('b baz \<rightarrow> 'a baz)"
 | 
| 33591 | 277  | 
where  | 
| 
33781
 
c7d32e726bb9
avoid using csplit; define copy functions exactly like the current domain package
 
huffman 
parents: 
33779 
diff
changeset
 | 
278  | 
"foo_bar_baz_mapF = (\<Lambda> f. Abs_CFun (\<lambda>(d1, d2, d3).  | 
| 33591 | 279  | 
(  | 
280  | 
foo_abs oo  | 
|
281  | 
ssum_map\<cdot>ID\<cdot>(sprod_map\<cdot>(u_map\<cdot>f)\<cdot>(u_map\<cdot>d2))  | 
|
282  | 
oo foo_rep  | 
|
283  | 
,  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
284  | 
bar_abs oo u_map\<cdot>(cfun_map\<cdot>d3\<cdot>ID) oo bar_rep  | 
| 33591 | 285  | 
,  | 
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
286  | 
baz_abs oo u_map\<cdot>(cfun_map\<cdot>(convex_map\<cdot>d1)\<cdot>ID) oo baz_rep  | 
| 
33781
 
c7d32e726bb9
avoid using csplit; define copy functions exactly like the current domain package
 
huffman 
parents: 
33779 
diff
changeset
 | 
287  | 
)))"  | 
| 33591 | 288  | 
|
289  | 
lemma foo_bar_baz_mapF_beta:  | 
|
290  | 
"foo_bar_baz_mapF\<cdot>f\<cdot>d =  | 
|
291  | 
(  | 
|
292  | 
foo_abs oo  | 
|
293  | 
ssum_map\<cdot>ID\<cdot>(sprod_map\<cdot>(u_map\<cdot>f)\<cdot>(u_map\<cdot>(fst (snd d))))  | 
|
294  | 
oo foo_rep  | 
|
295  | 
,  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
296  | 
bar_abs oo u_map\<cdot>(cfun_map\<cdot>(snd (snd d))\<cdot>ID) oo bar_rep  | 
| 33591 | 297  | 
,  | 
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
298  | 
baz_abs oo u_map\<cdot>(cfun_map\<cdot>(convex_map\<cdot>(fst d))\<cdot>ID) oo baz_rep  | 
| 33591 | 299  | 
)"  | 
300  | 
unfolding foo_bar_baz_mapF_def  | 
|
| 
33781
 
c7d32e726bb9
avoid using csplit; define copy functions exactly like the current domain package
 
huffman 
parents: 
33779 
diff
changeset
 | 
301  | 
by (simp add: split_def)  | 
| 33591 | 302  | 
|
303  | 
text {* Individual map functions are projected from the fixed point. *}
 | 
|
304  | 
||
305  | 
definition foo_map :: "('a \<rightarrow> 'b) \<rightarrow> ('a foo \<rightarrow> 'b foo)"
 | 
|
306  | 
where "foo_map = (\<Lambda> f. fst (fix\<cdot>(foo_bar_baz_mapF\<cdot>f)))"  | 
|
307  | 
||
308  | 
definition bar_map :: "('a \<rightarrow> 'b) \<rightarrow> ('a bar \<rightarrow> 'b bar)"
 | 
|
309  | 
where "bar_map = (\<Lambda> f. fst (snd (fix\<cdot>(foo_bar_baz_mapF\<cdot>f))))"  | 
|
310  | 
||
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
311  | 
definition baz_map :: "('a \<rightarrow> 'b) \<rightarrow> ('b baz \<rightarrow> 'a baz)"
 | 
| 33591 | 312  | 
where "baz_map = (\<Lambda> f. snd (snd (fix\<cdot>(foo_bar_baz_mapF\<cdot>f))))"  | 
313  | 
||
| 36132 | 314  | 
lemma map_apply_thms:  | 
315  | 
"foo_map\<cdot>f = fst (fix\<cdot>(foo_bar_baz_mapF\<cdot>f))"  | 
|
316  | 
"bar_map\<cdot>f = fst (snd (fix\<cdot>(foo_bar_baz_mapF\<cdot>f)))"  | 
|
317  | 
"baz_map\<cdot>f = snd (snd (fix\<cdot>(foo_bar_baz_mapF\<cdot>f)))"  | 
|
318  | 
unfolding foo_map_def bar_map_def baz_map_def by simp_all  | 
|
319  | 
||
| 33591 | 320  | 
text {* Prove isodefl rules for all map functions simultaneously. *}
 | 
321  | 
||
322  | 
lemma isodefl_foo_bar_baz:  | 
|
323  | 
assumes isodefl_d: "isodefl d t"  | 
|
324  | 
shows  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
325  | 
"isodefl (foo_map\<cdot>d) (foo_defl\<cdot>t) \<and>  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
326  | 
isodefl (bar_map\<cdot>d) (bar_defl\<cdot>t) \<and>  | 
| 
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
327  | 
isodefl (baz_map\<cdot>d) (baz_defl\<cdot>t)"  | 
| 36132 | 328  | 
unfolding map_apply_thms defl_apply_thms  | 
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
329  | 
apply (rule parallel_fix_ind)  | 
| 33591 | 330  | 
apply (intro adm_conj adm_isodefl cont2cont_fst cont2cont_snd cont_id)  | 
331  | 
apply (simp only: fst_strict snd_strict isodefl_bottom simp_thms)  | 
|
332  | 
apply (simp only: foo_bar_baz_mapF_beta  | 
|
| 
33784
 
7e434813752f
change naming convention for deflation combinators
 
huffman 
parents: 
33781 
diff
changeset
 | 
333  | 
foo_bar_baz_deflF_beta  | 
| 33591 | 334  | 
fst_conv snd_conv)  | 
335  | 
apply (elim conjE)  | 
|
336  | 
apply (intro  | 
|
337  | 
conjI  | 
|
338  | 
isodefl_foo_abs  | 
|
339  | 
isodefl_bar_abs  | 
|
340  | 
isodefl_baz_abs  | 
|
| 
33787
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
341  | 
isodefl_ssum isodefl_sprod isodefl_ID_REP  | 
| 
 
71a675065128
change example to use recursion with continuous function space
 
huffman 
parents: 
33784 
diff
changeset
 | 
342  | 
isodefl_u isodefl_convex isodefl_cfun  | 
| 33591 | 343  | 
isodefl_d  | 
344  | 
)  | 
|
345  | 
apply assumption+  | 
|
346  | 
done  | 
|
347  | 
||
348  | 
lemmas isodefl_foo = isodefl_foo_bar_baz [THEN conjunct1]  | 
|
349  | 
lemmas isodefl_bar = isodefl_foo_bar_baz [THEN conjunct2, THEN conjunct1]  | 
|
350  | 
lemmas isodefl_baz = isodefl_foo_bar_baz [THEN conjunct2, THEN conjunct2]  | 
|
351  | 
||
352  | 
text {* Prove map ID lemmas, using isodefl_REP_imp_ID *}
 | 
|
353  | 
||
354  | 
lemma foo_map_ID: "foo_map\<cdot>ID = ID"  | 
|
355  | 
apply (rule isodefl_REP_imp_ID)  | 
|
356  | 
apply (subst REP_foo)  | 
|
357  | 
apply (rule isodefl_foo)  | 
|
358  | 
apply (rule isodefl_ID_REP)  | 
|
359  | 
done  | 
|
360  | 
||
361  | 
lemma bar_map_ID: "bar_map\<cdot>ID = ID"  | 
|
362  | 
apply (rule isodefl_REP_imp_ID)  | 
|
363  | 
apply (subst REP_bar)  | 
|
364  | 
apply (rule isodefl_bar)  | 
|
365  | 
apply (rule isodefl_ID_REP)  | 
|
366  | 
done  | 
|
367  | 
||
368  | 
lemma baz_map_ID: "baz_map\<cdot>ID = ID"  | 
|
369  | 
apply (rule isodefl_REP_imp_ID)  | 
|
370  | 
apply (subst REP_baz)  | 
|
371  | 
apply (rule isodefl_baz)  | 
|
372  | 
apply (rule isodefl_ID_REP)  | 
|
373  | 
done  | 
|
374  | 
||
375  | 
(********************************************************************)  | 
|
376  | 
||
| 36132 | 377  | 
subsection {* Step 5: Define take functions, prove lub-take lemmas *}
 | 
| 
33781
 
c7d32e726bb9
avoid using csplit; define copy functions exactly like the current domain package
 
huffman 
parents: 
33779 
diff
changeset
 | 
378  | 
|
| 
 
c7d32e726bb9
avoid using csplit; define copy functions exactly like the current domain package
 
huffman 
parents: 
33779 
diff
changeset
 | 
379  | 
definition  | 
| 36132 | 380  | 
foo_bar_baz_takeF ::  | 
| 
33781
 
c7d32e726bb9
avoid using csplit; define copy functions exactly like the current domain package
 
huffman 
parents: 
33779 
diff
changeset
 | 
381  | 
    "('a foo \<rightarrow> 'a foo) \<times> ('a bar \<rightarrow> 'a bar) \<times> ('a baz \<rightarrow> 'a baz) \<rightarrow>
 | 
| 
 
c7d32e726bb9
avoid using csplit; define copy functions exactly like the current domain package
 
huffman 
parents: 
33779 
diff
changeset
 | 
382  | 
     ('a foo \<rightarrow> 'a foo) \<times> ('a bar \<rightarrow> 'a bar) \<times> ('a baz \<rightarrow> 'a baz)"
 | 
| 
 
c7d32e726bb9
avoid using csplit; define copy functions exactly like the current domain package
 
huffman 
parents: 
33779 
diff
changeset
 | 
383  | 
where  | 
| 36132 | 384  | 
"foo_bar_baz_takeF = (\<Lambda> p.  | 
385  | 
( foo_abs oo  | 
|
386  | 
ssum_map\<cdot>ID\<cdot>(sprod_map\<cdot>(u_map\<cdot>ID)\<cdot>(u_map\<cdot>(fst (snd p))))  | 
|
387  | 
oo foo_rep  | 
|
388  | 
, bar_abs oo  | 
|
389  | 
u_map\<cdot>(cfun_map\<cdot>(snd (snd p))\<cdot>ID) oo bar_rep  | 
|
390  | 
, baz_abs oo  | 
|
391  | 
u_map\<cdot>(cfun_map\<cdot>(convex_map\<cdot>(fst p))\<cdot>ID) oo baz_rep  | 
|
392  | 
))"  | 
|
393  | 
||
394  | 
lemma foo_bar_baz_takeF_beta:  | 
|
395  | 
"foo_bar_baz_takeF\<cdot>p =  | 
|
396  | 
( foo_abs oo  | 
|
397  | 
ssum_map\<cdot>ID\<cdot>(sprod_map\<cdot>(u_map\<cdot>ID)\<cdot>(u_map\<cdot>(fst (snd p))))  | 
|
398  | 
oo foo_rep  | 
|
399  | 
, bar_abs oo  | 
|
400  | 
u_map\<cdot>(cfun_map\<cdot>(snd (snd p))\<cdot>ID) oo bar_rep  | 
|
401  | 
, baz_abs oo  | 
|
402  | 
u_map\<cdot>(cfun_map\<cdot>(convex_map\<cdot>(fst p))\<cdot>ID) oo baz_rep  | 
|
403  | 
)"  | 
|
404  | 
unfolding foo_bar_baz_takeF_def by (rule beta_cfun, simp)  | 
|
405  | 
||
406  | 
definition  | 
|
407  | 
foo_take :: "nat \<Rightarrow> 'a foo \<rightarrow> 'a foo"  | 
|
408  | 
where  | 
|
409  | 
"foo_take = (\<lambda>n. fst (iterate n\<cdot>foo_bar_baz_takeF\<cdot>\<bottom>))"  | 
|
410  | 
||
411  | 
definition  | 
|
412  | 
bar_take :: "nat \<Rightarrow> 'a bar \<rightarrow> 'a bar"  | 
|
413  | 
where  | 
|
414  | 
"bar_take = (\<lambda>n. fst (snd (iterate n\<cdot>foo_bar_baz_takeF\<cdot>\<bottom>)))"  | 
|
| 33591 | 415  | 
|
| 36132 | 416  | 
definition  | 
417  | 
baz_take :: "nat \<Rightarrow> 'a baz \<rightarrow> 'a baz"  | 
|
418  | 
where  | 
|
419  | 
"baz_take = (\<lambda>n. snd (snd (iterate n\<cdot>foo_bar_baz_takeF\<cdot>\<bottom>)))"  | 
|
420  | 
||
421  | 
lemma chain_take_thms: "chain foo_take" "chain bar_take" "chain baz_take"  | 
|
422  | 
unfolding foo_take_def bar_take_def baz_take_def  | 
|
423  | 
by (intro ch2ch_fst ch2ch_snd chain_iterate)+  | 
|
424  | 
||
425  | 
lemma take_0_thms: "foo_take 0 = \<bottom>" "bar_take 0 = \<bottom>" "baz_take 0 = \<bottom>"  | 
|
426  | 
unfolding foo_take_def bar_take_def baz_take_def  | 
|
427  | 
by (simp only: iterate_0 fst_strict snd_strict)+  | 
|
428  | 
||
429  | 
lemma take_Suc_thms:  | 
|
430  | 
"foo_take (Suc n) =  | 
|
431  | 
foo_abs oo ssum_map\<cdot>ID\<cdot>(sprod_map\<cdot>(u_map\<cdot>ID)\<cdot>(u_map\<cdot>(bar_take n))) oo foo_rep"  | 
|
432  | 
"bar_take (Suc n) =  | 
|
433  | 
bar_abs oo u_map\<cdot>(cfun_map\<cdot>(baz_take n)\<cdot>ID) oo bar_rep"  | 
|
434  | 
"baz_take (Suc n) =  | 
|
435  | 
baz_abs oo u_map\<cdot>(cfun_map\<cdot>(convex_map\<cdot>(foo_take n))\<cdot>ID) oo baz_rep"  | 
|
436  | 
unfolding foo_take_def bar_take_def baz_take_def  | 
|
437  | 
by (simp only: iterate_Suc foo_bar_baz_takeF_beta fst_conv snd_conv)+  | 
|
438  | 
||
439  | 
lemma lub_take_lemma:  | 
|
440  | 
"(\<Squnion>n. foo_take n, \<Squnion>n. bar_take n, \<Squnion>n. baz_take n)  | 
|
441  | 
= (foo_map\<cdot>(ID::'a \<rightarrow> 'a), bar_map\<cdot>(ID::'a \<rightarrow> 'a), baz_map\<cdot>(ID::'a \<rightarrow> 'a))"  | 
|
442  | 
apply (simp only: thelub_Pair [symmetric] ch2ch_Pair chain_take_thms)  | 
|
443  | 
apply (simp only: map_apply_thms pair_collapse)  | 
|
444  | 
apply (simp only: fix_def2)  | 
|
445  | 
apply (rule lub_eq)  | 
|
446  | 
apply (rule nat.induct)  | 
|
447  | 
apply (simp only: iterate_0 Pair_strict take_0_thms)  | 
|
448  | 
apply (simp only: iterate_Suc Pair_fst_snd_eq fst_conv snd_conv  | 
|
449  | 
foo_bar_baz_mapF_beta take_Suc_thms simp_thms)  | 
|
| 
33781
 
c7d32e726bb9
avoid using csplit; define copy functions exactly like the current domain package
 
huffman 
parents: 
33779 
diff
changeset
 | 
450  | 
done  | 
| 33591 | 451  | 
|
| 36132 | 452  | 
lemma lub_foo_take: "(\<Squnion>n. foo_take n) = ID"  | 
453  | 
apply (rule trans [OF _ foo_map_ID])  | 
|
454  | 
using lub_take_lemma  | 
|
455  | 
apply (elim Pair_inject)  | 
|
456  | 
apply assumption  | 
|
457  | 
done  | 
|
| 33591 | 458  | 
|
| 36132 | 459  | 
lemma lub_bar_take: "(\<Squnion>n. bar_take n) = ID"  | 
460  | 
apply (rule trans [OF _ bar_map_ID])  | 
|
461  | 
using lub_take_lemma  | 
|
462  | 
apply (elim Pair_inject)  | 
|
463  | 
apply assumption  | 
|
464  | 
done  | 
|
| 33591 | 465  | 
|
| 36132 | 466  | 
lemma lub_baz_take: "(\<Squnion>n. baz_take n) = ID"  | 
467  | 
apply (rule trans [OF _ baz_map_ID])  | 
|
468  | 
using lub_take_lemma  | 
|
469  | 
apply (elim Pair_inject)  | 
|
470  | 
apply assumption  | 
|
471  | 
done  | 
|
| 33591 | 472  | 
|
473  | 
end  |