| author | bulwahn | 
| Tue, 07 Sep 2010 11:51:53 +0200 | |
| changeset 39189 | d183bf90dabd | 
| parent 36452 | d37c6eed8117 | 
| child 40026 | 8f8f18a88685 | 
| permissions | -rw-r--r-- | 
| 15741 | 1  | 
(* Title: HOLCF/Domain.thy  | 
2  | 
Author: Brian Huffman  | 
|
3  | 
*)  | 
|
4  | 
||
5  | 
header {* Domain package *}
 | 
|
6  | 
||
7  | 
theory Domain  | 
|
| 33800 | 8  | 
imports Ssum Sprod Up One Tr Fixrec Representable  | 
| 30910 | 9  | 
uses  | 
10  | 
  ("Tools/cont_consts.ML")
 | 
|
11  | 
  ("Tools/cont_proc.ML")
 | 
|
| 35444 | 12  | 
  ("Tools/Domain/domain_constructors.ML")
 | 
| 32126 | 13  | 
  ("Tools/Domain/domain_library.ML")
 | 
14  | 
  ("Tools/Domain/domain_axioms.ML")
 | 
|
15  | 
  ("Tools/Domain/domain_theorems.ML")
 | 
|
16  | 
  ("Tools/Domain/domain_extender.ML")
 | 
|
| 15741 | 17  | 
begin  | 
18  | 
||
| 36452 | 19  | 
default_sort pcpo  | 
| 15741 | 20  | 
|
| 23376 | 21  | 
|
| 15741 | 22  | 
subsection {* Casedist *}
 | 
23  | 
||
24  | 
lemma ex_one_defined_iff:  | 
|
25  | 
"(\<exists>x. P x \<and> x \<noteq> \<bottom>) = P ONE"  | 
|
26  | 
apply safe  | 
|
27  | 
apply (rule_tac p=x in oneE)  | 
|
28  | 
apply simp  | 
|
29  | 
apply simp  | 
|
30  | 
apply force  | 
|
| 23376 | 31  | 
done  | 
| 15741 | 32  | 
|
33  | 
lemma ex_up_defined_iff:  | 
|
34  | 
"(\<exists>x. P x \<and> x \<noteq> \<bottom>) = (\<exists>x. P (up\<cdot>x))"  | 
|
35  | 
apply safe  | 
|
| 16754 | 36  | 
apply (rule_tac p=x in upE)  | 
| 15741 | 37  | 
apply simp  | 
38  | 
apply fast  | 
|
| 16320 | 39  | 
apply (force intro!: up_defined)  | 
| 23376 | 40  | 
done  | 
| 15741 | 41  | 
|
42  | 
lemma ex_sprod_defined_iff:  | 
|
43  | 
"(\<exists>y. P y \<and> y \<noteq> \<bottom>) =  | 
|
44  | 
(\<exists>x y. (P (:x, y:) \<and> x \<noteq> \<bottom>) \<and> y \<noteq> \<bottom>)"  | 
|
45  | 
apply safe  | 
|
46  | 
apply (rule_tac p=y in sprodE)  | 
|
47  | 
apply simp  | 
|
48  | 
apply fast  | 
|
| 16217 | 49  | 
apply (force intro!: spair_defined)  | 
| 23376 | 50  | 
done  | 
| 15741 | 51  | 
|
52  | 
lemma ex_sprod_up_defined_iff:  | 
|
53  | 
"(\<exists>y. P y \<and> y \<noteq> \<bottom>) =  | 
|
54  | 
(\<exists>x y. P (:up\<cdot>x, y:) \<and> y \<noteq> \<bottom>)"  | 
|
55  | 
apply safe  | 
|
56  | 
apply (rule_tac p=y in sprodE)  | 
|
57  | 
apply simp  | 
|
| 16754 | 58  | 
apply (rule_tac p=x in upE)  | 
| 15741 | 59  | 
apply simp  | 
60  | 
apply fast  | 
|
| 16217 | 61  | 
apply (force intro!: spair_defined)  | 
| 23376 | 62  | 
done  | 
| 15741 | 63  | 
|
64  | 
lemma ex_ssum_defined_iff:  | 
|
65  | 
"(\<exists>x. P x \<and> x \<noteq> \<bottom>) =  | 
|
66  | 
((\<exists>x. P (sinl\<cdot>x) \<and> x \<noteq> \<bottom>) \<or>  | 
|
67  | 
(\<exists>x. P (sinr\<cdot>x) \<and> x \<noteq> \<bottom>))"  | 
|
68  | 
apply (rule iffI)  | 
|
69  | 
apply (erule exE)  | 
|
70  | 
apply (erule conjE)  | 
|
71  | 
apply (rule_tac p=x in ssumE)  | 
|
72  | 
apply simp  | 
|
73  | 
apply (rule disjI1, fast)  | 
|
74  | 
apply (rule disjI2, fast)  | 
|
75  | 
apply (erule disjE)  | 
|
| 17835 | 76  | 
apply force  | 
77  | 
apply force  | 
|
| 23376 | 78  | 
done  | 
| 15741 | 79  | 
|
80  | 
lemma exh_start: "p = \<bottom> \<or> (\<exists>x. p = x \<and> x \<noteq> \<bottom>)"  | 
|
| 23376 | 81  | 
by auto  | 
| 15741 | 82  | 
|
83  | 
lemmas ex_defined_iffs =  | 
|
84  | 
ex_ssum_defined_iff  | 
|
85  | 
ex_sprod_up_defined_iff  | 
|
86  | 
ex_sprod_defined_iff  | 
|
87  | 
ex_up_defined_iff  | 
|
88  | 
ex_one_defined_iff  | 
|
89  | 
||
90  | 
text {* Rules for turning exh into casedist *}
 | 
|
91  | 
||
92  | 
lemma exh_casedist0: "\<lbrakk>R; R \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" (* like make_elim *)  | 
|
| 23376 | 93  | 
by auto  | 
| 15741 | 94  | 
|
95  | 
lemma exh_casedist1: "((P \<or> Q \<Longrightarrow> R) \<Longrightarrow> S) \<equiv> (\<lbrakk>P \<Longrightarrow> R; Q \<Longrightarrow> R\<rbrakk> \<Longrightarrow> S)"  | 
|
| 23376 | 96  | 
by rule auto  | 
| 15741 | 97  | 
|
98  | 
lemma exh_casedist2: "(\<exists>x. P x \<Longrightarrow> Q) \<equiv> (\<And>x. P x \<Longrightarrow> Q)"  | 
|
| 23376 | 99  | 
by rule auto  | 
| 15741 | 100  | 
|
101  | 
lemma exh_casedist3: "(P \<and> Q \<Longrightarrow> R) \<equiv> (P \<Longrightarrow> Q \<Longrightarrow> R)"  | 
|
| 23376 | 102  | 
by rule auto  | 
| 15741 | 103  | 
|
104  | 
lemmas exh_casedists = exh_casedist1 exh_casedist2 exh_casedist3  | 
|
105  | 
||
| 30910 | 106  | 
|
| 31230 | 107  | 
subsection {* Combinators for building copy functions *}
 | 
108  | 
||
| 
33504
 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 
huffman 
parents: 
33400 
diff
changeset
 | 
109  | 
lemmas domain_map_stricts =  | 
| 
 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 
huffman 
parents: 
33400 
diff
changeset
 | 
110  | 
ssum_map_strict sprod_map_strict u_map_strict  | 
| 31230 | 111  | 
|
| 
33504
 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 
huffman 
parents: 
33400 
diff
changeset
 | 
112  | 
lemmas domain_map_simps =  | 
| 
 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 
huffman 
parents: 
33400 
diff
changeset
 | 
113  | 
ssum_map_sinl ssum_map_sinr sprod_map_spair u_map_up  | 
| 31230 | 114  | 
|
115  | 
||
| 30910 | 116  | 
subsection {* Installing the domain package *}
 | 
117  | 
||
| 
30911
 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 
huffman 
parents: 
30910 
diff
changeset
 | 
118  | 
lemmas con_strict_rules =  | 
| 
 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 
huffman 
parents: 
30910 
diff
changeset
 | 
119  | 
sinl_strict sinr_strict spair_strict1 spair_strict2  | 
| 
 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 
huffman 
parents: 
30910 
diff
changeset
 | 
120  | 
|
| 
 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 
huffman 
parents: 
30910 
diff
changeset
 | 
121  | 
lemmas con_defin_rules =  | 
| 
 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 
huffman 
parents: 
30910 
diff
changeset
 | 
122  | 
sinl_defined sinr_defined spair_defined up_defined ONE_defined  | 
| 
 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 
huffman 
parents: 
30910 
diff
changeset
 | 
123  | 
|
| 
 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 
huffman 
parents: 
30910 
diff
changeset
 | 
124  | 
lemmas con_defined_iff_rules =  | 
| 
 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 
huffman 
parents: 
30910 
diff
changeset
 | 
125  | 
sinl_defined_iff sinr_defined_iff spair_strict_iff up_defined ONE_defined  | 
| 
 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 
huffman 
parents: 
30910 
diff
changeset
 | 
126  | 
|
| 
35117
 
eeec2a320a77
change generated lemmas dist_eqs and dist_les to iff-style
 
huffman 
parents: 
33800 
diff
changeset
 | 
127  | 
lemmas con_below_iff_rules =  | 
| 
 
eeec2a320a77
change generated lemmas dist_eqs and dist_les to iff-style
 
huffman 
parents: 
33800 
diff
changeset
 | 
128  | 
sinl_below sinr_below sinl_below_sinr sinr_below_sinl con_defined_iff_rules  | 
| 
 
eeec2a320a77
change generated lemmas dist_eqs and dist_les to iff-style
 
huffman 
parents: 
33800 
diff
changeset
 | 
129  | 
|
| 
 
eeec2a320a77
change generated lemmas dist_eqs and dist_les to iff-style
 
huffman 
parents: 
33800 
diff
changeset
 | 
130  | 
lemmas con_eq_iff_rules =  | 
| 
 
eeec2a320a77
change generated lemmas dist_eqs and dist_les to iff-style
 
huffman 
parents: 
33800 
diff
changeset
 | 
131  | 
sinl_eq sinr_eq sinl_eq_sinr sinr_eq_sinl con_defined_iff_rules  | 
| 
 
eeec2a320a77
change generated lemmas dist_eqs and dist_les to iff-style
 
huffman 
parents: 
33800 
diff
changeset
 | 
132  | 
|
| 
35446
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
133  | 
lemmas sel_strict_rules =  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
134  | 
cfcomp2 sscase1 sfst_strict ssnd_strict fup1  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
135  | 
|
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
136  | 
lemma sel_app_extra_rules:  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
137  | 
"sscase\<cdot>ID\<cdot>\<bottom>\<cdot>(sinr\<cdot>x) = \<bottom>"  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
138  | 
"sscase\<cdot>ID\<cdot>\<bottom>\<cdot>(sinl\<cdot>x) = x"  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
139  | 
"sscase\<cdot>\<bottom>\<cdot>ID\<cdot>(sinl\<cdot>x) = \<bottom>"  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
140  | 
"sscase\<cdot>\<bottom>\<cdot>ID\<cdot>(sinr\<cdot>x) = x"  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
141  | 
"fup\<cdot>ID\<cdot>(up\<cdot>x) = x"  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
142  | 
by (cases "x = \<bottom>", simp, simp)+  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
143  | 
|
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
144  | 
lemmas sel_app_rules =  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
145  | 
sel_strict_rules sel_app_extra_rules  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
146  | 
ssnd_spair sfst_spair up_defined spair_defined  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
147  | 
|
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
148  | 
lemmas sel_defined_iff_rules =  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
149  | 
cfcomp2 sfst_defined_iff ssnd_defined_iff  | 
| 
 
b719dad322fa
rewrite domain package code for selector functions
 
huffman 
parents: 
35444 
diff
changeset
 | 
150  | 
|
| 
35494
 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 
huffman 
parents: 
35457 
diff
changeset
 | 
151  | 
lemmas take_con_rules =  | 
| 
36160
 
f84fa49a0b69
add rule deflation_ID to proof script for take + constructor rules
 
huffman 
parents: 
35653 
diff
changeset
 | 
152  | 
ssum_map_sinl' ssum_map_sinr' sprod_map_spair' u_map_up  | 
| 
 
f84fa49a0b69
add rule deflation_ID to proof script for take + constructor rules
 
huffman 
parents: 
35653 
diff
changeset
 | 
153  | 
deflation_strict deflation_ID ID1 cfcomp2  | 
| 
35494
 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 
huffman 
parents: 
35457 
diff
changeset
 | 
154  | 
|
| 30910 | 155  | 
use "Tools/cont_consts.ML"  | 
156  | 
use "Tools/cont_proc.ML"  | 
|
| 32126 | 157  | 
use "Tools/Domain/domain_library.ML"  | 
158  | 
use "Tools/Domain/domain_axioms.ML"  | 
|
| 
35457
 
d63655b88369
move proofs of casedist into domain_constructors.ML
 
huffman 
parents: 
35446 
diff
changeset
 | 
159  | 
use "Tools/Domain/domain_constructors.ML"  | 
| 32126 | 160  | 
use "Tools/Domain/domain_theorems.ML"  | 
161  | 
use "Tools/Domain/domain_extender.ML"  | 
|
| 30910 | 162  | 
|
| 15741 | 163  | 
end  |