| author | wenzelm | 
| Tue, 02 Nov 2010 21:24:07 +0100 | |
| changeset 40309 | de842e206db2 | 
| parent 40089 | 8adc57fb8454 | 
| child 40321 | d065b195ec89 | 
| permissions | -rw-r--r-- | 
| 2640 | 1  | 
(* Title: HOLCF/Lift.thy  | 
| 12026 | 2  | 
Author: Olaf Mueller  | 
| 2640 | 3  | 
*)  | 
4  | 
||
| 
12338
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
12026 
diff
changeset
 | 
5  | 
header {* Lifting types of class type to flat pcpo's *}
 | 
| 12026 | 6  | 
|
| 15577 | 7  | 
theory Lift  | 
| 40086 | 8  | 
imports Discrete Up  | 
| 15577 | 9  | 
begin  | 
| 12026 | 10  | 
|
| 36452 | 11  | 
default_sort type  | 
| 12026 | 12  | 
|
| 40082 | 13  | 
pcpodef (open) 'a lift = "UNIV :: 'a discr u set"  | 
| 
29063
 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 
wenzelm 
parents: 
27311 
diff
changeset
 | 
14  | 
by simp_all  | 
| 12026 | 15  | 
|
| 16748 | 16  | 
lemmas inst_lift_pcpo = Abs_lift_strict [symmetric]  | 
| 12026 | 17  | 
|
| 
25131
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19764 
diff
changeset
 | 
18  | 
definition  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19764 
diff
changeset
 | 
19  | 
Def :: "'a \<Rightarrow> 'a lift" where  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19764 
diff
changeset
 | 
20  | 
"Def x = Abs_lift (up\<cdot>(Discr x))"  | 
| 12026 | 21  | 
|
22  | 
subsection {* Lift as a datatype *}
 | 
|
23  | 
||
| 16748 | 24  | 
lemma lift_induct: "\<lbrakk>P \<bottom>; \<And>x. P (Def x)\<rbrakk> \<Longrightarrow> P y"  | 
25  | 
apply (induct y)  | 
|
| 
16755
 
fd02f9d06e43
renamed upE1 to upE; added simp rule cont2cont_flift1
 
huffman 
parents: 
16749 
diff
changeset
 | 
26  | 
apply (rule_tac p=y in upE)  | 
| 16748 | 27  | 
apply (simp add: Abs_lift_strict)  | 
28  | 
apply (case_tac x)  | 
|
29  | 
apply (simp add: Def_def)  | 
|
30  | 
done  | 
|
| 12026 | 31  | 
|
| 
27104
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26963 
diff
changeset
 | 
32  | 
rep_datatype "\<bottom>\<Colon>'a lift" Def  | 
| 40082 | 33  | 
by (erule lift_induct) (simp_all add: Def_def Abs_lift_inject inst_lift_pcpo)  | 
| 12026 | 34  | 
|
| 
27104
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26963 
diff
changeset
 | 
35  | 
lemmas lift_distinct1 = lift.distinct(1)  | 
| 
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26963 
diff
changeset
 | 
36  | 
lemmas lift_distinct2 = lift.distinct(2)  | 
| 
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26963 
diff
changeset
 | 
37  | 
lemmas Def_not_UU = lift.distinct(2)  | 
| 
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26963 
diff
changeset
 | 
38  | 
lemmas Def_inject = lift.inject  | 
| 12026 | 39  | 
|
40  | 
||
| 16748 | 41  | 
text {* @{term UU} and @{term Def} *}
 | 
| 12026 | 42  | 
|
| 16748 | 43  | 
lemma not_Undef_is_Def: "(x \<noteq> \<bottom>) = (\<exists>y. x = Def y)"  | 
| 12026 | 44  | 
by (cases x) simp_all  | 
45  | 
||
| 
16630
 
83bf468b1dc7
added theorem lift_definedE; moved cont_if to Cont.thy
 
huffman 
parents: 
16555 
diff
changeset
 | 
46  | 
lemma lift_definedE: "\<lbrakk>x \<noteq> \<bottom>; \<And>a. x = Def a \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"  | 
| 
 
83bf468b1dc7
added theorem lift_definedE; moved cont_if to Cont.thy
 
huffman 
parents: 
16555 
diff
changeset
 | 
47  | 
by (cases x) simp_all  | 
| 
 
83bf468b1dc7
added theorem lift_definedE; moved cont_if to Cont.thy
 
huffman 
parents: 
16555 
diff
changeset
 | 
48  | 
|
| 12026 | 49  | 
text {*
 | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
29138 
diff
changeset
 | 
50  | 
  For @{term "x ~= UU"} in assumptions @{text defined} replaces @{text
 | 
| 12026 | 51  | 
  x} by @{text "Def a"} in conclusion. *}
 | 
52  | 
||
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
29138 
diff
changeset
 | 
53  | 
method_setup defined = {*
 | 
| 
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
29138 
diff
changeset
 | 
54  | 
Scan.succeed (fn ctxt => SIMPLE_METHOD'  | 
| 
32149
 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
31076 
diff
changeset
 | 
55  | 
    (etac @{thm lift_definedE} THEN' asm_simp_tac (simpset_of ctxt)))
 | 
| 
30607
 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 
wenzelm 
parents: 
29138 
diff
changeset
 | 
56  | 
*} ""  | 
| 12026 | 57  | 
|
| 16748 | 58  | 
lemma DefE: "Def x = \<bottom> \<Longrightarrow> R"  | 
59  | 
by simp  | 
|
| 12026 | 60  | 
|
| 16748 | 61  | 
lemma DefE2: "\<lbrakk>x = Def s; x = \<bottom>\<rbrakk> \<Longrightarrow> R"  | 
| 12026 | 62  | 
by simp  | 
63  | 
||
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
30607 
diff
changeset
 | 
64  | 
lemma Def_below_Def: "Def x \<sqsubseteq> Def y \<longleftrightarrow> x = y"  | 
| 40082 | 65  | 
by (simp add: below_lift_def Def_def Abs_lift_inverse)  | 
| 12026 | 66  | 
|
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
30607 
diff
changeset
 | 
67  | 
lemma Def_below_iff [simp]: "Def x \<sqsubseteq> y \<longleftrightarrow> Def x = y"  | 
| 
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
30607 
diff
changeset
 | 
68  | 
by (induct y, simp, simp add: Def_below_Def)  | 
| 12026 | 69  | 
|
70  | 
||
71  | 
subsection {* Lift is flat *}
 | 
|
72  | 
||
| 
12338
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
12026 
diff
changeset
 | 
73  | 
instance lift :: (type) flat  | 
| 27292 | 74  | 
proof  | 
75  | 
fix x y :: "'a lift"  | 
|
76  | 
assume "x \<sqsubseteq> y" thus "x = \<bottom> \<or> x = y"  | 
|
77  | 
by (induct x) auto  | 
|
78  | 
qed  | 
|
| 12026 | 79  | 
|
| 
40088
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
80  | 
subsection {* Continuity of @{const lift_case} *}
 | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
81  | 
|
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
82  | 
lemma lift_case_eq: "lift_case \<bottom> f x = fup\<cdot>(\<Lambda> y. f (undiscr y))\<cdot>(Rep_lift x)"  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
83  | 
apply (induct x, unfold lift.cases)  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
84  | 
apply (simp add: Rep_lift_strict)  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
85  | 
apply (simp add: Def_def Abs_lift_inverse)  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
86  | 
done  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
87  | 
|
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
88  | 
lemma cont2cont_lift_case [simp]:  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
89  | 
"\<lbrakk>\<And>y. cont (\<lambda>x. f x y); cont g\<rbrakk> \<Longrightarrow> cont (\<lambda>x. lift_case \<bottom> (f x) (g x))"  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
90  | 
unfolding lift_case_eq by (simp add: cont_Rep_lift [THEN cont_compose])  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
91  | 
|
| 
16695
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
92  | 
subsection {* Further operations *}
 | 
| 
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
93  | 
|
| 
25131
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19764 
diff
changeset
 | 
94  | 
definition  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19764 
diff
changeset
 | 
95  | 
  flift1 :: "('a \<Rightarrow> 'b::pcpo) \<Rightarrow> ('a lift \<rightarrow> 'b)"  (binder "FLIFT " 10)  where
 | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19764 
diff
changeset
 | 
96  | 
"flift1 = (\<lambda>f. (\<Lambda> x. lift_case \<bottom> f x))"  | 
| 
16695
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
97  | 
|
| 
25131
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19764 
diff
changeset
 | 
98  | 
definition  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19764 
diff
changeset
 | 
99  | 
  flift2 :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a lift \<rightarrow> 'b lift)" where
 | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19764 
diff
changeset
 | 
100  | 
"flift2 f = (FLIFT x. Def (f x))"  | 
| 
16695
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
101  | 
|
| 
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
102  | 
lemma flift1_Def [simp]: "flift1 f\<cdot>(Def x) = (f x)"  | 
| 
40088
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
103  | 
by (simp add: flift1_def)  | 
| 
16695
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
104  | 
|
| 
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
105  | 
lemma flift2_Def [simp]: "flift2 f\<cdot>(Def x) = Def (f x)"  | 
| 
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
106  | 
by (simp add: flift2_def)  | 
| 
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
107  | 
|
| 
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
108  | 
lemma flift1_strict [simp]: "flift1 f\<cdot>\<bottom> = \<bottom>"  | 
| 
40088
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
109  | 
by (simp add: flift1_def)  | 
| 
16695
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
110  | 
|
| 
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
111  | 
lemma flift2_strict [simp]: "flift2 f\<cdot>\<bottom> = \<bottom>"  | 
| 
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
112  | 
by (simp add: flift2_def)  | 
| 
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
113  | 
|
| 
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
114  | 
lemma flift2_defined [simp]: "x \<noteq> \<bottom> \<Longrightarrow> (flift2 f)\<cdot>x \<noteq> \<bottom>"  | 
| 
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
115  | 
by (erule lift_definedE, simp)  | 
| 
 
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
 
huffman 
parents: 
16630 
diff
changeset
 | 
116  | 
|
| 19520 | 117  | 
lemma flift2_defined_iff [simp]: "(flift2 f\<cdot>x = \<bottom>) = (x = \<bottom>)"  | 
118  | 
by (cases x, simp_all)  | 
|
119  | 
||
| 
40088
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
120  | 
lemma FLIFT_mono:  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
121  | 
"(\<And>x. f x \<sqsubseteq> g x) \<Longrightarrow> (FLIFT x. f x) \<sqsubseteq> (FLIFT x. g x)"  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
122  | 
by (rule cfun_belowI, case_tac x, simp_all)  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
123  | 
|
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
124  | 
lemma cont2cont_flift1 [simp, cont2cont]:  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
125  | 
"\<lbrakk>\<And>y. cont (\<lambda>x. f x y)\<rbrakk> \<Longrightarrow> cont (\<lambda>x. FLIFT y. f x y)"  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
126  | 
by (simp add: flift1_def cont2cont_LAM)  | 
| 
 
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
 
huffman 
parents: 
40086 
diff
changeset
 | 
127  | 
|
| 2640 | 128  | 
end  |