author | Fabian Huch <huch@in.tum.de> |
Thu, 30 Nov 2023 13:44:51 +0100 | |
changeset 79084 | dd689c4ab688 |
parent 72835 | 66ca5016b008 |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
42151 | 1 |
(* Title: HOL/HOLCF/Lift.thy |
72835 | 2 |
Author: Olaf Müller |
2640 | 3 |
*) |
4 |
||
62175 | 5 |
section \<open>Lifting types of class type to flat pcpo's\<close> |
12026 | 6 |
|
15577 | 7 |
theory Lift |
40086 | 8 |
imports Discrete Up |
15577 | 9 |
begin |
12026 | 10 |
|
36452 | 11 |
default_sort type |
12026 | 12 |
|
49759
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
huffman
parents:
42814
diff
changeset
|
13 |
pcpodef '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 |
|
62175 | 22 |
subsection \<open>Lift as a datatype\<close> |
12026 | 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 |
|
61076 | 32 |
old_rep_datatype "\<bottom>::'a lift" Def |
40082 | 33 |
by (erule lift_induct) (simp_all add: Def_def Abs_lift_inject inst_lift_pcpo) |
12026 | 34 |
|
69597 | 35 |
text \<open>\<^term>\<open>bottom\<close> and \<^term>\<open>Def\<close>\<close> |
12026 | 36 |
|
16748 | 37 |
lemma not_Undef_is_Def: "(x \<noteq> \<bottom>) = (\<exists>y. x = Def y)" |
12026 | 38 |
by (cases x) simp_all |
39 |
||
16630
83bf468b1dc7
added theorem lift_definedE; moved cont_if to Cont.thy
huffman
parents:
16555
diff
changeset
|
40 |
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
|
41 |
by (cases x) simp_all |
83bf468b1dc7
added theorem lift_definedE; moved cont_if to Cont.thy
huffman
parents:
16555
diff
changeset
|
42 |
|
62175 | 43 |
text \<open> |
69597 | 44 |
For \<^term>\<open>x ~= \<bottom>\<close> in assumptions \<open>defined\<close> replaces \<open>x\<close> by \<open>Def a\<close> in conclusion.\<close> |
12026 | 45 |
|
62175 | 46 |
method_setup defined = \<open> |
30607
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents:
29138
diff
changeset
|
47 |
Scan.succeed (fn ctxt => SIMPLE_METHOD' |
60754 | 48 |
(eresolve_tac ctxt @{thms lift_definedE} THEN' asm_simp_tac ctxt)) |
62175 | 49 |
\<close> |
12026 | 50 |
|
16748 | 51 |
lemma DefE: "Def x = \<bottom> \<Longrightarrow> R" |
52 |
by simp |
|
12026 | 53 |
|
16748 | 54 |
lemma DefE2: "\<lbrakk>x = Def s; x = \<bottom>\<rbrakk> \<Longrightarrow> R" |
12026 | 55 |
by simp |
56 |
||
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
|
57 |
lemma Def_below_Def: "Def x \<sqsubseteq> Def y \<longleftrightarrow> x = y" |
40082 | 58 |
by (simp add: below_lift_def Def_def Abs_lift_inverse) |
12026 | 59 |
|
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
|
60 |
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
|
61 |
by (induct y, simp, simp add: Def_below_Def) |
12026 | 62 |
|
63 |
||
62175 | 64 |
subsection \<open>Lift is flat\<close> |
12026 | 65 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12026
diff
changeset
|
66 |
instance lift :: (type) flat |
27292 | 67 |
proof |
68 |
fix x y :: "'a lift" |
|
69 |
assume "x \<sqsubseteq> y" thus "x = \<bottom> \<or> x = y" |
|
70 |
by (induct x) auto |
|
71 |
qed |
|
12026 | 72 |
|
69597 | 73 |
subsection \<open>Continuity of \<^const>\<open>case_lift\<close>\<close> |
40088
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
74 |
|
55417
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents:
51717
diff
changeset
|
75 |
lemma case_lift_eq: "case_lift \<bottom> f x = fup\<cdot>(\<Lambda> y. f (undiscr y))\<cdot>(Rep_lift x)" |
55642
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents:
55417
diff
changeset
|
76 |
apply (induct x, unfold lift.case) |
40088
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
77 |
apply (simp add: Rep_lift_strict) |
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
78 |
apply (simp add: Def_def Abs_lift_inverse) |
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
79 |
done |
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
80 |
|
55417
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents:
51717
diff
changeset
|
81 |
lemma cont2cont_case_lift [simp]: |
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents:
51717
diff
changeset
|
82 |
"\<lbrakk>\<And>y. cont (\<lambda>x. f x y); cont g\<rbrakk> \<Longrightarrow> cont (\<lambda>x. case_lift \<bottom> (f x) (g x))" |
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents:
51717
diff
changeset
|
83 |
unfolding case_lift_eq by (simp add: cont_Rep_lift) |
40088
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
84 |
|
62175 | 85 |
subsection \<open>Further operations\<close> |
16695
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
86 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
19764
diff
changeset
|
87 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
19764
diff
changeset
|
88 |
flift1 :: "('a \<Rightarrow> 'b::pcpo) \<Rightarrow> ('a lift \<rightarrow> 'b)" (binder "FLIFT " 10) where |
55417
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents:
51717
diff
changeset
|
89 |
"flift1 = (\<lambda>f. (\<Lambda> x. case_lift \<bottom> f x))" |
16695
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
90 |
|
40323
4cce7c708402
add 'LAM (Def x). t' as alternative syntax for 'FLIFT x. t'
huffman
parents:
40321
diff
changeset
|
91 |
translations |
4cce7c708402
add 'LAM (Def x). t' as alternative syntax for 'FLIFT x. t'
huffman
parents:
40321
diff
changeset
|
92 |
"\<Lambda>(XCONST Def x). t" => "CONST flift1 (\<lambda>x. t)" |
4cce7c708402
add 'LAM (Def x). t' as alternative syntax for 'FLIFT x. t'
huffman
parents:
40321
diff
changeset
|
93 |
"\<Lambda>(CONST Def x). FLIFT y. t" <= "FLIFT x y. t" |
4cce7c708402
add 'LAM (Def x). t' as alternative syntax for 'FLIFT x. t'
huffman
parents:
40321
diff
changeset
|
94 |
"\<Lambda>(CONST Def x). t" <= "FLIFT x. t" |
4cce7c708402
add 'LAM (Def x). t' as alternative syntax for 'FLIFT x. t'
huffman
parents:
40321
diff
changeset
|
95 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
19764
diff
changeset
|
96 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
19764
diff
changeset
|
97 |
flift2 :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a lift \<rightarrow> 'b lift)" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
19764
diff
changeset
|
98 |
"flift2 f = (FLIFT x. Def (f x))" |
16695
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
99 |
|
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
100 |
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
|
101 |
by (simp add: flift1_def) |
16695
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
102 |
|
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
103 |
lemma flift2_Def [simp]: "flift2 f\<cdot>(Def x) = Def (f x)" |
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
104 |
by (simp add: flift2_def) |
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
105 |
|
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
106 |
lemma flift1_strict [simp]: "flift1 f\<cdot>\<bottom> = \<bottom>" |
40088
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
107 |
by (simp add: flift1_def) |
16695
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
108 |
|
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
109 |
lemma flift2_strict [simp]: "flift2 f\<cdot>\<bottom> = \<bottom>" |
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
110 |
by (simp add: flift2_def) |
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
111 |
|
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
112 |
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
|
113 |
by (erule lift_definedE, simp) |
dc8c868e910b
simplified definitions of flift1, flift2, liftpair;
huffman
parents:
16630
diff
changeset
|
114 |
|
40321
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
huffman
parents:
40089
diff
changeset
|
115 |
lemma flift2_bottom_iff [simp]: "(flift2 f\<cdot>x = \<bottom>) = (x = \<bottom>)" |
19520 | 116 |
by (cases x, simp_all) |
117 |
||
40088
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
118 |
lemma FLIFT_mono: |
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
119 |
"(\<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
|
120 |
by (rule cfun_belowI, case_tac x, simp_all) |
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
121 |
|
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
122 |
lemma cont2cont_flift1 [simp, cont2cont]: |
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
123 |
"\<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
|
124 |
by (simp add: flift1_def cont2cont_LAM) |
49b9d301fadb
simplify proofs about flift; remove unneeded lemmas
huffman
parents:
40086
diff
changeset
|
125 |
|
2640 | 126 |
end |