author | blanchet |
Tue, 12 Nov 2013 13:47:24 +0100 | |
changeset 54396 | 8baee6b04a7c |
parent 53952 | b2781a3ce958 |
child 55083 | 0a689157e3ce |
permissions | -rw-r--r-- |
47308 | 1 |
(* Title: HOL/Lifting.thy |
2 |
Author: Brian Huffman and Ondrej Kuncar |
|
3 |
Author: Cezary Kaliszyk and Christian Urban |
|
4 |
*) |
|
5 |
||
6 |
header {* Lifting package *} |
|
7 |
||
8 |
theory Lifting |
|
51112 | 9 |
imports Equiv_Relations Transfer |
47308 | 10 |
keywords |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
11 |
"parametric" and |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
53151
diff
changeset
|
12 |
"print_quot_maps" "print_quotients" :: diag and |
47308 | 13 |
"lift_definition" :: thy_goal and |
53651 | 14 |
"setup_lifting" "lifting_forget" "lifting_update" :: thy_decl |
47308 | 15 |
begin |
16 |
||
47325 | 17 |
subsection {* Function map *} |
47308 | 18 |
|
53011
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52307
diff
changeset
|
19 |
context |
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52307
diff
changeset
|
20 |
begin |
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52307
diff
changeset
|
21 |
interpretation lifting_syntax . |
47308 | 22 |
|
23 |
lemma map_fun_id: |
|
24 |
"(id ---> id) = id" |
|
25 |
by (simp add: fun_eq_iff) |
|
26 |
||
51994 | 27 |
subsection {* Other predicates on relations *} |
28 |
||
29 |
definition left_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" |
|
30 |
where "left_total R \<longleftrightarrow> (\<forall>x. \<exists>y. R x y)" |
|
31 |
||
32 |
lemma left_totalI: |
|
33 |
"(\<And>x. \<exists>y. R x y) \<Longrightarrow> left_total R" |
|
34 |
unfolding left_total_def by blast |
|
35 |
||
36 |
lemma left_totalE: |
|
37 |
assumes "left_total R" |
|
38 |
obtains "(\<And>x. \<exists>y. R x y)" |
|
39 |
using assms unfolding left_total_def by blast |
|
40 |
||
53952 | 41 |
lemma bi_total_iff: "bi_total A = (right_total A \<and> left_total A)" |
42 |
unfolding left_total_def right_total_def bi_total_def by blast |
|
43 |
||
53927 | 44 |
lemma bi_total_conv_left_right: "bi_total R \<longleftrightarrow> left_total R \<and> right_total R" |
45 |
by(simp add: left_total_def right_total_def bi_total_def) |
|
46 |
||
51994 | 47 |
definition left_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" |
48 |
where "left_unique R \<longleftrightarrow> (\<forall>x y z. R x z \<longrightarrow> R y z \<longrightarrow> x = y)" |
|
49 |
||
53952 | 50 |
lemma left_unique_transfer [transfer_rule]: |
51 |
assumes [transfer_rule]: "right_total A" |
|
52 |
assumes [transfer_rule]: "right_total B" |
|
53 |
assumes [transfer_rule]: "bi_unique A" |
|
54 |
shows "((A ===> B ===> op=) ===> implies) left_unique left_unique" |
|
55 |
using assms unfolding left_unique_def[abs_def] right_total_def bi_unique_def fun_rel_def |
|
56 |
by metis |
|
57 |
||
58 |
lemma bi_unique_iff: "bi_unique A = (right_unique A \<and> left_unique A)" |
|
59 |
unfolding left_unique_def right_unique_def bi_unique_def by blast |
|
60 |
||
53927 | 61 |
lemma bi_unique_conv_left_right: "bi_unique R \<longleftrightarrow> left_unique R \<and> right_unique R" |
62 |
by(auto simp add: left_unique_def right_unique_def bi_unique_def) |
|
63 |
||
64 |
lemma left_uniqueI: "(\<And>x y z. \<lbrakk> A x z; A y z \<rbrakk> \<Longrightarrow> x = y) \<Longrightarrow> left_unique A" |
|
65 |
unfolding left_unique_def by blast |
|
66 |
||
67 |
lemma left_uniqueD: "\<lbrakk> left_unique A; A x z; A y z \<rbrakk> \<Longrightarrow> x = y" |
|
68 |
unfolding left_unique_def by blast |
|
69 |
||
52036
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
70 |
lemma left_total_fun: |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
71 |
"\<lbrakk>left_unique A; left_total B\<rbrakk> \<Longrightarrow> left_total (A ===> B)" |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
72 |
unfolding left_total_def fun_rel_def |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
73 |
apply (rule allI, rename_tac f) |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
74 |
apply (rule_tac x="\<lambda>y. SOME z. B (f (THE x. A x y)) z" in exI) |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
75 |
apply clarify |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
76 |
apply (subgoal_tac "(THE x. A x y) = x", simp) |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
77 |
apply (rule someI_ex) |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
78 |
apply (simp) |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
79 |
apply (rule the_equality) |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
80 |
apply assumption |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
81 |
apply (simp add: left_unique_def) |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
82 |
done |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
83 |
|
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
84 |
lemma left_unique_fun: |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
85 |
"\<lbrakk>left_total A; left_unique B\<rbrakk> \<Longrightarrow> left_unique (A ===> B)" |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
86 |
unfolding left_total_def left_unique_def fun_rel_def |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
87 |
by (clarify, rule ext, fast) |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
88 |
|
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
89 |
lemma left_total_eq: "left_total op=" unfolding left_total_def by blast |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
90 |
|
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
91 |
lemma left_unique_eq: "left_unique op=" unfolding left_unique_def by blast |
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
92 |
|
53944 | 93 |
lemma [simp]: |
94 |
shows left_unique_conversep: "left_unique A\<inverse>\<inverse> \<longleftrightarrow> right_unique A" |
|
95 |
and right_unique_conversep: "right_unique A\<inverse>\<inverse> \<longleftrightarrow> left_unique A" |
|
96 |
by(auto simp add: left_unique_def right_unique_def) |
|
97 |
||
98 |
lemma [simp]: |
|
99 |
shows left_total_conversep: "left_total A\<inverse>\<inverse> \<longleftrightarrow> right_total A" |
|
100 |
and right_total_conversep: "right_total A\<inverse>\<inverse> \<longleftrightarrow> left_total A" |
|
101 |
by(simp_all add: left_total_def right_total_def) |
|
102 |
||
47308 | 103 |
subsection {* Quotient Predicate *} |
104 |
||
105 |
definition |
|
106 |
"Quotient R Abs Rep T \<longleftrightarrow> |
|
107 |
(\<forall>a. Abs (Rep a) = a) \<and> |
|
108 |
(\<forall>a. R (Rep a) (Rep a)) \<and> |
|
109 |
(\<forall>r s. R r s \<longleftrightarrow> R r r \<and> R s s \<and> Abs r = Abs s) \<and> |
|
110 |
T = (\<lambda>x y. R x x \<and> Abs x = y)" |
|
111 |
||
112 |
lemma QuotientI: |
|
113 |
assumes "\<And>a. Abs (Rep a) = a" |
|
114 |
and "\<And>a. R (Rep a) (Rep a)" |
|
115 |
and "\<And>r s. R r s \<longleftrightarrow> R r r \<and> R s s \<and> Abs r = Abs s" |
|
116 |
and "T = (\<lambda>x y. R x x \<and> Abs x = y)" |
|
117 |
shows "Quotient R Abs Rep T" |
|
118 |
using assms unfolding Quotient_def by blast |
|
119 |
||
47536 | 120 |
context |
121 |
fixes R Abs Rep T |
|
47308 | 122 |
assumes a: "Quotient R Abs Rep T" |
47536 | 123 |
begin |
124 |
||
125 |
lemma Quotient_abs_rep: "Abs (Rep a) = a" |
|
126 |
using a unfolding Quotient_def |
|
47308 | 127 |
by simp |
128 |
||
47536 | 129 |
lemma Quotient_rep_reflp: "R (Rep a) (Rep a)" |
130 |
using a unfolding Quotient_def |
|
47308 | 131 |
by blast |
132 |
||
133 |
lemma Quotient_rel: |
|
47536 | 134 |
"R r r \<and> R s s \<and> Abs r = Abs s \<longleftrightarrow> R r s" -- {* orientation does not loop on rewriting *} |
135 |
using a unfolding Quotient_def |
|
47308 | 136 |
by blast |
137 |
||
47536 | 138 |
lemma Quotient_cr_rel: "T = (\<lambda>x y. R x x \<and> Abs x = y)" |
47308 | 139 |
using a unfolding Quotient_def |
140 |
by blast |
|
141 |
||
47536 | 142 |
lemma Quotient_refl1: "R r s \<Longrightarrow> R r r" |
143 |
using a unfolding Quotient_def |
|
144 |
by fast |
|
145 |
||
146 |
lemma Quotient_refl2: "R r s \<Longrightarrow> R s s" |
|
147 |
using a unfolding Quotient_def |
|
148 |
by fast |
|
149 |
||
150 |
lemma Quotient_rel_rep: "R (Rep a) (Rep b) \<longleftrightarrow> a = b" |
|
151 |
using a unfolding Quotient_def |
|
152 |
by metis |
|
153 |
||
154 |
lemma Quotient_rep_abs: "R r r \<Longrightarrow> R (Rep (Abs r)) r" |
|
47308 | 155 |
using a unfolding Quotient_def |
156 |
by blast |
|
157 |
||
47937
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
158 |
lemma Quotient_rep_abs_fold_unmap: |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
159 |
assumes "x' \<equiv> Abs x" and "R x x" and "Rep x' \<equiv> Rep' x'" |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
160 |
shows "R (Rep' x') x" |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
161 |
proof - |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
162 |
have "R (Rep x') x" using assms(1-2) Quotient_rep_abs by auto |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
163 |
then show ?thesis using assms(3) by simp |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
164 |
qed |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
165 |
|
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
166 |
lemma Quotient_Rep_eq: |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
167 |
assumes "x' \<equiv> Abs x" |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
168 |
shows "Rep x' \<equiv> Rep x'" |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
169 |
by simp |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
170 |
|
47536 | 171 |
lemma Quotient_rel_abs: "R r s \<Longrightarrow> Abs r = Abs s" |
172 |
using a unfolding Quotient_def |
|
173 |
by blast |
|
174 |
||
47937
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
175 |
lemma Quotient_rel_abs2: |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
176 |
assumes "R (Rep x) y" |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
177 |
shows "x = Abs y" |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
178 |
proof - |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
179 |
from assms have "Abs (Rep x) = Abs y" by (auto intro: Quotient_rel_abs) |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
180 |
then show ?thesis using assms(1) by (simp add: Quotient_abs_rep) |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
181 |
qed |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
182 |
|
47536 | 183 |
lemma Quotient_symp: "symp R" |
47308 | 184 |
using a unfolding Quotient_def using sympI by (metis (full_types)) |
185 |
||
47536 | 186 |
lemma Quotient_transp: "transp R" |
47308 | 187 |
using a unfolding Quotient_def using transpI by (metis (full_types)) |
188 |
||
47536 | 189 |
lemma Quotient_part_equivp: "part_equivp R" |
190 |
by (metis Quotient_rep_reflp Quotient_symp Quotient_transp part_equivpI) |
|
191 |
||
192 |
end |
|
47308 | 193 |
|
194 |
lemma identity_quotient: "Quotient (op =) id id (op =)" |
|
195 |
unfolding Quotient_def by simp |
|
196 |
||
47652
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
197 |
text {* TODO: Use one of these alternatives as the real definition. *} |
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
198 |
|
47308 | 199 |
lemma Quotient_alt_def: |
200 |
"Quotient R Abs Rep T \<longleftrightarrow> |
|
201 |
(\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and> |
|
202 |
(\<forall>b. T (Rep b) b) \<and> |
|
203 |
(\<forall>x y. R x y \<longleftrightarrow> T x (Abs x) \<and> T y (Abs y) \<and> Abs x = Abs y)" |
|
204 |
apply safe |
|
205 |
apply (simp (no_asm_use) only: Quotient_def, fast) |
|
206 |
apply (simp (no_asm_use) only: Quotient_def, fast) |
|
207 |
apply (simp (no_asm_use) only: Quotient_def, fast) |
|
208 |
apply (simp (no_asm_use) only: Quotient_def, fast) |
|
209 |
apply (simp (no_asm_use) only: Quotient_def, fast) |
|
210 |
apply (simp (no_asm_use) only: Quotient_def, fast) |
|
211 |
apply (rule QuotientI) |
|
212 |
apply simp |
|
213 |
apply metis |
|
214 |
apply simp |
|
215 |
apply (rule ext, rule ext, metis) |
|
216 |
done |
|
217 |
||
218 |
lemma Quotient_alt_def2: |
|
219 |
"Quotient R Abs Rep T \<longleftrightarrow> |
|
220 |
(\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and> |
|
221 |
(\<forall>b. T (Rep b) b) \<and> |
|
222 |
(\<forall>x y. R x y \<longleftrightarrow> T x (Abs y) \<and> T y (Abs x))" |
|
223 |
unfolding Quotient_alt_def by (safe, metis+) |
|
224 |
||
47652
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
225 |
lemma Quotient_alt_def3: |
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
226 |
"Quotient R Abs Rep T \<longleftrightarrow> |
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
227 |
(\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and> (\<forall>b. T (Rep b) b) \<and> |
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
228 |
(\<forall>x y. R x y \<longleftrightarrow> (\<exists>z. T x z \<and> T y z))" |
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
229 |
unfolding Quotient_alt_def2 by (safe, metis+) |
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
230 |
|
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
231 |
lemma Quotient_alt_def4: |
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
232 |
"Quotient R Abs Rep T \<longleftrightarrow> |
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
233 |
(\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and> (\<forall>b. T (Rep b) b) \<and> R = T OO conversep T" |
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
234 |
unfolding Quotient_alt_def3 fun_eq_iff by auto |
1b722b100301
move alternative definition lemmas into Lifting.thy;
huffman
parents:
47651
diff
changeset
|
235 |
|
47308 | 236 |
lemma fun_quotient: |
237 |
assumes 1: "Quotient R1 abs1 rep1 T1" |
|
238 |
assumes 2: "Quotient R2 abs2 rep2 T2" |
|
239 |
shows "Quotient (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2) (T1 ===> T2)" |
|
240 |
using assms unfolding Quotient_alt_def2 |
|
241 |
unfolding fun_rel_def fun_eq_iff map_fun_apply |
|
242 |
by (safe, metis+) |
|
243 |
||
244 |
lemma apply_rsp: |
|
245 |
fixes f g::"'a \<Rightarrow> 'c" |
|
246 |
assumes q: "Quotient R1 Abs1 Rep1 T1" |
|
247 |
and a: "(R1 ===> R2) f g" "R1 x y" |
|
248 |
shows "R2 (f x) (g y)" |
|
249 |
using a by (auto elim: fun_relE) |
|
250 |
||
251 |
lemma apply_rsp': |
|
252 |
assumes a: "(R1 ===> R2) f g" "R1 x y" |
|
253 |
shows "R2 (f x) (g y)" |
|
254 |
using a by (auto elim: fun_relE) |
|
255 |
||
256 |
lemma apply_rsp'': |
|
257 |
assumes "Quotient R Abs Rep T" |
|
258 |
and "(R ===> S) f f" |
|
259 |
shows "S (f (Rep x)) (f (Rep x))" |
|
260 |
proof - |
|
261 |
from assms(1) have "R (Rep x) (Rep x)" by (rule Quotient_rep_reflp) |
|
262 |
then show ?thesis using assms(2) by (auto intro: apply_rsp') |
|
263 |
qed |
|
264 |
||
265 |
subsection {* Quotient composition *} |
|
266 |
||
267 |
lemma Quotient_compose: |
|
268 |
assumes 1: "Quotient R1 Abs1 Rep1 T1" |
|
269 |
assumes 2: "Quotient R2 Abs2 Rep2 T2" |
|
270 |
shows "Quotient (T1 OO R2 OO conversep T1) (Abs2 \<circ> Abs1) (Rep1 \<circ> Rep2) (T1 OO T2)" |
|
51994 | 271 |
using assms unfolding Quotient_alt_def4 by fastforce |
47308 | 272 |
|
47521
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47501
diff
changeset
|
273 |
lemma equivp_reflp2: |
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47501
diff
changeset
|
274 |
"equivp R \<Longrightarrow> reflp R" |
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47501
diff
changeset
|
275 |
by (erule equivpE) |
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47501
diff
changeset
|
276 |
|
47544 | 277 |
subsection {* Respects predicate *} |
278 |
||
279 |
definition Respects :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set" |
|
280 |
where "Respects R = {x. R x x}" |
|
281 |
||
282 |
lemma in_respects: "x \<in> Respects R \<longleftrightarrow> R x x" |
|
283 |
unfolding Respects_def by simp |
|
284 |
||
47308 | 285 |
subsection {* Invariant *} |
286 |
||
287 |
definition invariant :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" |
|
288 |
where "invariant R = (\<lambda>x y. R x \<and> x = y)" |
|
289 |
||
290 |
lemma invariant_to_eq: |
|
291 |
assumes "invariant P x y" |
|
292 |
shows "x = y" |
|
293 |
using assms by (simp add: invariant_def) |
|
294 |
||
295 |
lemma fun_rel_eq_invariant: |
|
296 |
shows "((invariant R) ===> S) = (\<lambda>f g. \<forall>x. R x \<longrightarrow> S (f x) (g x))" |
|
297 |
by (auto simp add: invariant_def fun_rel_def) |
|
298 |
||
299 |
lemma invariant_same_args: |
|
300 |
shows "invariant P x x \<equiv> P x" |
|
301 |
using assms by (auto simp add: invariant_def) |
|
302 |
||
53952 | 303 |
lemma invariant_transfer [transfer_rule]: |
304 |
assumes [transfer_rule]: "bi_unique A" |
|
305 |
shows "((A ===> op=) ===> A ===> A ===> op=) Lifting.invariant Lifting.invariant" |
|
306 |
unfolding invariant_def[abs_def] by transfer_prover |
|
307 |
||
47361
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
308 |
lemma UNIV_typedef_to_Quotient: |
47308 | 309 |
assumes "type_definition Rep Abs UNIV" |
47361
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
310 |
and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)" |
47308 | 311 |
shows "Quotient (op =) Abs Rep T" |
312 |
proof - |
|
313 |
interpret type_definition Rep Abs UNIV by fact |
|
47361
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
314 |
from Abs_inject Rep_inverse Abs_inverse T_def show ?thesis |
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
315 |
by (fastforce intro!: QuotientI fun_eq_iff) |
47308 | 316 |
qed |
317 |
||
47361
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
318 |
lemma UNIV_typedef_to_equivp: |
47308 | 319 |
fixes Abs :: "'a \<Rightarrow> 'b" |
320 |
and Rep :: "'b \<Rightarrow> 'a" |
|
321 |
assumes "type_definition Rep Abs (UNIV::'a set)" |
|
322 |
shows "equivp (op=::'a\<Rightarrow>'a\<Rightarrow>bool)" |
|
323 |
by (rule identity_equivp) |
|
324 |
||
47354 | 325 |
lemma typedef_to_Quotient: |
47361
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
326 |
assumes "type_definition Rep Abs S" |
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
327 |
and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)" |
47501 | 328 |
shows "Quotient (invariant (\<lambda>x. x \<in> S)) Abs Rep T" |
47361
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
329 |
proof - |
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
330 |
interpret type_definition Rep Abs S by fact |
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
331 |
from Rep Abs_inject Rep_inverse Abs_inverse T_def show ?thesis |
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
332 |
by (auto intro!: QuotientI simp: invariant_def fun_eq_iff) |
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
333 |
qed |
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
334 |
|
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
335 |
lemma typedef_to_part_equivp: |
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
336 |
assumes "type_definition Rep Abs S" |
47501 | 337 |
shows "part_equivp (invariant (\<lambda>x. x \<in> S))" |
47361
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
338 |
proof (intro part_equivpI) |
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
339 |
interpret type_definition Rep Abs S by fact |
47501 | 340 |
show "\<exists>x. invariant (\<lambda>x. x \<in> S) x x" using Rep by (auto simp: invariant_def) |
47361
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
341 |
next |
47501 | 342 |
show "symp (invariant (\<lambda>x. x \<in> S))" by (auto intro: sympI simp: invariant_def) |
47361
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
343 |
next |
47501 | 344 |
show "transp (invariant (\<lambda>x. x \<in> S))" by (auto intro: transpI simp: invariant_def) |
47361
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
345 |
qed |
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
346 |
|
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
347 |
lemma open_typedef_to_Quotient: |
47308 | 348 |
assumes "type_definition Rep Abs {x. P x}" |
47354 | 349 |
and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)" |
47308 | 350 |
shows "Quotient (invariant P) Abs Rep T" |
47651 | 351 |
using typedef_to_Quotient [OF assms] by simp |
47308 | 352 |
|
47361
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47354
diff
changeset
|
353 |
lemma open_typedef_to_part_equivp: |
47308 | 354 |
assumes "type_definition Rep Abs {x. P x}" |
355 |
shows "part_equivp (invariant P)" |
|
47651 | 356 |
using typedef_to_part_equivp [OF assms] by simp |
47308 | 357 |
|
47376 | 358 |
text {* Generating transfer rules for quotients. *} |
359 |
||
47537 | 360 |
context |
361 |
fixes R Abs Rep T |
|
362 |
assumes 1: "Quotient R Abs Rep T" |
|
363 |
begin |
|
47376 | 364 |
|
47537 | 365 |
lemma Quotient_right_unique: "right_unique T" |
366 |
using 1 unfolding Quotient_alt_def right_unique_def by metis |
|
367 |
||
368 |
lemma Quotient_right_total: "right_total T" |
|
369 |
using 1 unfolding Quotient_alt_def right_total_def by metis |
|
370 |
||
371 |
lemma Quotient_rel_eq_transfer: "(T ===> T ===> op =) R (op =)" |
|
372 |
using 1 unfolding Quotient_alt_def fun_rel_def by simp |
|
47376 | 373 |
|
47538 | 374 |
lemma Quotient_abs_induct: |
375 |
assumes "\<And>y. R y y \<Longrightarrow> P (Abs y)" shows "P x" |
|
376 |
using 1 assms unfolding Quotient_def by metis |
|
377 |
||
47537 | 378 |
end |
379 |
||
380 |
text {* Generating transfer rules for total quotients. *} |
|
47376 | 381 |
|
47537 | 382 |
context |
383 |
fixes R Abs Rep T |
|
384 |
assumes 1: "Quotient R Abs Rep T" and 2: "reflp R" |
|
385 |
begin |
|
47376 | 386 |
|
47537 | 387 |
lemma Quotient_bi_total: "bi_total T" |
388 |
using 1 2 unfolding Quotient_alt_def bi_total_def reflp_def by auto |
|
389 |
||
390 |
lemma Quotient_id_abs_transfer: "(op = ===> T) (\<lambda>x. x) Abs" |
|
391 |
using 1 2 unfolding Quotient_alt_def reflp_def fun_rel_def by simp |
|
392 |
||
47575 | 393 |
lemma Quotient_total_abs_induct: "(\<And>y. P (Abs y)) \<Longrightarrow> P x" |
394 |
using 1 2 assms unfolding Quotient_alt_def reflp_def by metis |
|
395 |
||
47889
29212a4bb866
lifting package produces abs_eq_iff rules for total quotients
huffman
parents:
47777
diff
changeset
|
396 |
lemma Quotient_total_abs_eq_iff: "Abs x = Abs y \<longleftrightarrow> R x y" |
29212a4bb866
lifting package produces abs_eq_iff rules for total quotients
huffman
parents:
47777
diff
changeset
|
397 |
using Quotient_rel [OF 1] 2 unfolding reflp_def by simp |
29212a4bb866
lifting package produces abs_eq_iff rules for total quotients
huffman
parents:
47777
diff
changeset
|
398 |
|
47537 | 399 |
end |
47376 | 400 |
|
47368
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
huffman
parents:
47354
diff
changeset
|
401 |
text {* Generating transfer rules for a type defined with @{text "typedef"}. *} |
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
huffman
parents:
47354
diff
changeset
|
402 |
|
47534
06cc372a80ed
use context block to organize typedef lifting theorems
huffman
parents:
47521
diff
changeset
|
403 |
context |
06cc372a80ed
use context block to organize typedef lifting theorems
huffman
parents:
47521
diff
changeset
|
404 |
fixes Rep Abs A T |
47368
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
huffman
parents:
47354
diff
changeset
|
405 |
assumes type: "type_definition Rep Abs A" |
47534
06cc372a80ed
use context block to organize typedef lifting theorems
huffman
parents:
47521
diff
changeset
|
406 |
assumes T_def: "T \<equiv> (\<lambda>(x::'a) (y::'b). x = Rep y)" |
06cc372a80ed
use context block to organize typedef lifting theorems
huffman
parents:
47521
diff
changeset
|
407 |
begin |
06cc372a80ed
use context block to organize typedef lifting theorems
huffman
parents:
47521
diff
changeset
|
408 |
|
51994 | 409 |
lemma typedef_left_unique: "left_unique T" |
410 |
unfolding left_unique_def T_def |
|
411 |
by (simp add: type_definition.Rep_inject [OF type]) |
|
412 |
||
47534
06cc372a80ed
use context block to organize typedef lifting theorems
huffman
parents:
47521
diff
changeset
|
413 |
lemma typedef_bi_unique: "bi_unique T" |
47368
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
huffman
parents:
47354
diff
changeset
|
414 |
unfolding bi_unique_def T_def |
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
huffman
parents:
47354
diff
changeset
|
415 |
by (simp add: type_definition.Rep_inject [OF type]) |
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
huffman
parents:
47354
diff
changeset
|
416 |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
417 |
(* the following two theorems are here only for convinience *) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
418 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
419 |
lemma typedef_right_unique: "right_unique T" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
420 |
using T_def type Quotient_right_unique typedef_to_Quotient |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
421 |
by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
422 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
423 |
lemma typedef_right_total: "right_total T" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
424 |
using T_def type Quotient_right_total typedef_to_Quotient |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
425 |
by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
426 |
|
47535
0f94b02fda1c
lifting_setup generates transfer rule for rep of typedefs
huffman
parents:
47534
diff
changeset
|
427 |
lemma typedef_rep_transfer: "(T ===> op =) (\<lambda>x. x) Rep" |
0f94b02fda1c
lifting_setup generates transfer rule for rep of typedefs
huffman
parents:
47534
diff
changeset
|
428 |
unfolding fun_rel_def T_def by simp |
0f94b02fda1c
lifting_setup generates transfer rule for rep of typedefs
huffman
parents:
47534
diff
changeset
|
429 |
|
47534
06cc372a80ed
use context block to organize typedef lifting theorems
huffman
parents:
47521
diff
changeset
|
430 |
end |
06cc372a80ed
use context block to organize typedef lifting theorems
huffman
parents:
47521
diff
changeset
|
431 |
|
47368
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
huffman
parents:
47354
diff
changeset
|
432 |
text {* Generating the correspondence rule for a constant defined with |
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
huffman
parents:
47354
diff
changeset
|
433 |
@{text "lift_definition"}. *} |
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
huffman
parents:
47354
diff
changeset
|
434 |
|
47351 | 435 |
lemma Quotient_to_transfer: |
436 |
assumes "Quotient R Abs Rep T" and "R c c" and "c' \<equiv> Abs c" |
|
437 |
shows "T c c'" |
|
438 |
using assms by (auto dest: Quotient_cr_rel) |
|
439 |
||
47982
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
440 |
text {* Proving reflexivity *} |
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
441 |
|
51994 | 442 |
definition reflp' :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where "reflp' R \<equiv> reflp R" |
47982
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
443 |
|
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
444 |
lemma Quotient_to_left_total: |
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
445 |
assumes q: "Quotient R Abs Rep T" |
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
446 |
and r_R: "reflp R" |
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
447 |
shows "left_total T" |
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
448 |
using r_R Quotient_cr_rel[OF q] unfolding left_total_def by (auto elim: reflpE) |
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
449 |
|
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
450 |
lemma reflp_Quotient_composition: |
51994 | 451 |
assumes "left_total R" |
452 |
assumes "reflp T" |
|
453 |
shows "reflp (R OO T OO R\<inverse>\<inverse>)" |
|
454 |
using assms unfolding reflp_def left_total_def by fast |
|
455 |
||
456 |
lemma reflp_fun1: |
|
457 |
assumes "is_equality R" |
|
458 |
assumes "reflp' S" |
|
459 |
shows "reflp (R ===> S)" |
|
460 |
using assms unfolding is_equality_def reflp'_def reflp_def fun_rel_def by blast |
|
461 |
||
462 |
lemma reflp_fun2: |
|
463 |
assumes "is_equality R" |
|
464 |
assumes "is_equality S" |
|
465 |
shows "reflp (R ===> S)" |
|
466 |
using assms unfolding is_equality_def reflp_def fun_rel_def by blast |
|
467 |
||
468 |
lemma is_equality_Quotient_composition: |
|
469 |
assumes "is_equality T" |
|
470 |
assumes "left_total R" |
|
471 |
assumes "left_unique R" |
|
472 |
shows "is_equality (R OO T OO R\<inverse>\<inverse>)" |
|
473 |
using assms unfolding is_equality_def left_total_def left_unique_def OO_def conversep_iff |
|
474 |
by fastforce |
|
47982
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
475 |
|
52307 | 476 |
lemma left_total_composition: "left_total R \<Longrightarrow> left_total S \<Longrightarrow> left_total (R OO S)" |
477 |
unfolding left_total_def OO_def by fast |
|
478 |
||
479 |
lemma left_unique_composition: "left_unique R \<Longrightarrow> left_unique S \<Longrightarrow> left_unique (R OO S)" |
|
480 |
unfolding left_unique_def OO_def by fast |
|
481 |
||
47982
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
482 |
lemma reflp_equality: "reflp (op =)" |
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
483 |
by (auto intro: reflpI) |
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47937
diff
changeset
|
484 |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
485 |
text {* Proving a parametrized correspondence relation *} |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
486 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
487 |
lemma eq_OO: "op= OO R = R" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
488 |
unfolding OO_def by metis |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
489 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
490 |
definition POS :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" where |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
491 |
"POS A B \<equiv> A \<le> B" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
492 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
493 |
definition NEG :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" where |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
494 |
"NEG A B \<equiv> B \<le> A" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
495 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
496 |
(* |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
497 |
The following two rules are here because we don't have any proper |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
498 |
left-unique ant left-total relations. Left-unique and left-total |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
499 |
assumptions show up in distributivity rules for the function type. |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
500 |
*) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
501 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
502 |
lemma bi_unique_left_unique[transfer_rule]: "bi_unique R \<Longrightarrow> left_unique R" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
503 |
unfolding bi_unique_def left_unique_def by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
504 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
505 |
lemma bi_total_left_total[transfer_rule]: "bi_total R \<Longrightarrow> left_total R" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
506 |
unfolding bi_total_def left_total_def by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
507 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
508 |
lemma pos_OO_eq: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
509 |
shows "POS (A OO op=) A" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
510 |
unfolding POS_def OO_def by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
511 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
512 |
lemma pos_eq_OO: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
513 |
shows "POS (op= OO A) A" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
514 |
unfolding POS_def OO_def by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
515 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
516 |
lemma neg_OO_eq: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
517 |
shows "NEG (A OO op=) A" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
518 |
unfolding NEG_def OO_def by auto |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
519 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
520 |
lemma neg_eq_OO: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
521 |
shows "NEG (op= OO A) A" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
522 |
unfolding NEG_def OO_def by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
523 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
524 |
lemma POS_trans: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
525 |
assumes "POS A B" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
526 |
assumes "POS B C" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
527 |
shows "POS A C" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
528 |
using assms unfolding POS_def by auto |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
529 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
530 |
lemma NEG_trans: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
531 |
assumes "NEG A B" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
532 |
assumes "NEG B C" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
533 |
shows "NEG A C" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
534 |
using assms unfolding NEG_def by auto |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
535 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
536 |
lemma POS_NEG: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
537 |
"POS A B \<equiv> NEG B A" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
538 |
unfolding POS_def NEG_def by auto |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
539 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
540 |
lemma NEG_POS: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
541 |
"NEG A B \<equiv> POS B A" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
542 |
unfolding POS_def NEG_def by auto |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
543 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
544 |
lemma POS_pcr_rule: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
545 |
assumes "POS (A OO B) C" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
546 |
shows "POS (A OO B OO X) (C OO X)" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
547 |
using assms unfolding POS_def OO_def by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
548 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
549 |
lemma NEG_pcr_rule: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
550 |
assumes "NEG (A OO B) C" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
551 |
shows "NEG (A OO B OO X) (C OO X)" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
552 |
using assms unfolding NEG_def OO_def by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
553 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
554 |
lemma POS_apply: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
555 |
assumes "POS R R'" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
556 |
assumes "R f g" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
557 |
shows "R' f g" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
558 |
using assms unfolding POS_def by auto |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
559 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
560 |
text {* Proving a parametrized correspondence relation *} |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
561 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
562 |
lemma fun_mono: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
563 |
assumes "A \<ge> C" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
564 |
assumes "B \<le> D" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
565 |
shows "(A ===> B) \<le> (C ===> D)" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
566 |
using assms unfolding fun_rel_def by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
567 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
568 |
lemma pos_fun_distr: "((R ===> S) OO (R' ===> S')) \<le> ((R OO R') ===> (S OO S'))" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
569 |
unfolding OO_def fun_rel_def by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
570 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
571 |
lemma functional_relation: "right_unique R \<Longrightarrow> left_total R \<Longrightarrow> \<forall>x. \<exists>!y. R x y" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
572 |
unfolding right_unique_def left_total_def by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
573 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
574 |
lemma functional_converse_relation: "left_unique R \<Longrightarrow> right_total R \<Longrightarrow> \<forall>y. \<exists>!x. R x y" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
575 |
unfolding left_unique_def right_total_def by blast |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
576 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
577 |
lemma neg_fun_distr1: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
578 |
assumes 1: "left_unique R" "right_total R" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
579 |
assumes 2: "right_unique R'" "left_total R'" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
580 |
shows "(R OO R' ===> S OO S') \<le> ((R ===> S) OO (R' ===> S')) " |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
581 |
using functional_relation[OF 2] functional_converse_relation[OF 1] |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
582 |
unfolding fun_rel_def OO_def |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
583 |
apply clarify |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
584 |
apply (subst all_comm) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
585 |
apply (subst all_conj_distrib[symmetric]) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
586 |
apply (intro choice) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
587 |
by metis |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
588 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
589 |
lemma neg_fun_distr2: |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
590 |
assumes 1: "right_unique R'" "left_total R'" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
591 |
assumes 2: "left_unique S'" "right_total S'" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
592 |
shows "(R OO R' ===> S OO S') \<le> ((R ===> S) OO (R' ===> S'))" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
593 |
using functional_converse_relation[OF 2] functional_relation[OF 1] |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
594 |
unfolding fun_rel_def OO_def |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
595 |
apply clarify |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
596 |
apply (subst all_comm) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
597 |
apply (subst all_conj_distrib[symmetric]) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
598 |
apply (intro choice) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
599 |
by metis |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
600 |
|
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
601 |
subsection {* Domains *} |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
602 |
|
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
603 |
lemma pcr_Domainp_par_left_total: |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
604 |
assumes "Domainp B = P" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
605 |
assumes "left_total A" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
606 |
assumes "(A ===> op=) P' P" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
607 |
shows "Domainp (A OO B) = P'" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
608 |
using assms |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
609 |
unfolding Domainp_iff[abs_def] OO_def bi_unique_def left_total_def fun_rel_def |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
610 |
by (fast intro: fun_eq_iff) |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
611 |
|
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
612 |
lemma pcr_Domainp_par: |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
613 |
assumes "Domainp B = P2" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
614 |
assumes "Domainp A = P1" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
615 |
assumes "(A ===> op=) P2' P2" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
616 |
shows "Domainp (A OO B) = (inf P1 P2')" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
617 |
using assms unfolding fun_rel_def Domainp_iff[abs_def] OO_def |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
618 |
by (fast intro: fun_eq_iff) |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
619 |
|
53151 | 620 |
definition rel_pred_comp :: "('a => 'b => bool) => ('b => bool) => 'a => bool" |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
621 |
where "rel_pred_comp R P \<equiv> \<lambda>x. \<exists>y. R x y \<and> P y" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
622 |
|
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
623 |
lemma pcr_Domainp: |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
624 |
assumes "Domainp B = P" |
53151 | 625 |
shows "Domainp (A OO B) = (\<lambda>x. \<exists>y. A x y \<and> P y)" |
626 |
using assms by blast |
|
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
627 |
|
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
628 |
lemma pcr_Domainp_total: |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
629 |
assumes "bi_total B" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
630 |
assumes "Domainp A = P" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
631 |
shows "Domainp (A OO B) = P" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
632 |
using assms unfolding bi_total_def |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
633 |
by fast |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
634 |
|
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
635 |
lemma Quotient_to_Domainp: |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
636 |
assumes "Quotient R Abs Rep T" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
637 |
shows "Domainp T = (\<lambda>x. R x x)" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
638 |
by (simp add: Domainp_iff[abs_def] Quotient_cr_rel[OF assms]) |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
639 |
|
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
640 |
lemma invariant_to_Domainp: |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
641 |
assumes "Quotient (Lifting.invariant P) Abs Rep T" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
642 |
shows "Domainp T = P" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
643 |
by (simp add: invariant_def Domainp_iff[abs_def] Quotient_cr_rel[OF assms]) |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
644 |
|
53011
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52307
diff
changeset
|
645 |
end |
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52307
diff
changeset
|
646 |
|
47308 | 647 |
subsection {* ML setup *} |
648 |
||
48891 | 649 |
ML_file "Tools/Lifting/lifting_util.ML" |
47308 | 650 |
|
48891 | 651 |
ML_file "Tools/Lifting/lifting_info.ML" |
47308 | 652 |
setup Lifting_Info.setup |
653 |
||
51994 | 654 |
lemmas [reflexivity_rule] = |
52036
1aa2e40df9ff
reflexivity rules for the function type and equality
kuncar
parents:
51994
diff
changeset
|
655 |
reflp_equality reflp_Quotient_composition is_equality_Quotient_composition |
52307 | 656 |
left_total_fun left_unique_fun left_total_eq left_unique_eq left_total_composition |
657 |
left_unique_composition |
|
51994 | 658 |
|
659 |
text {* add @{thm reflp_fun1} and @{thm reflp_fun2} manually through ML |
|
660 |
because we don't want to get reflp' variant of these theorems *} |
|
661 |
||
662 |
setup{* |
|
663 |
Context.theory_map |
|
664 |
(fold |
|
665 |
(snd oo (Thm.apply_attribute Lifting_Info.add_reflexivity_rule_raw_attribute)) |
|
666 |
[@{thm reflp_fun1}, @{thm reflp_fun2}]) |
|
667 |
*} |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
668 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
669 |
(* setup for the function type *) |
47777
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47698
diff
changeset
|
670 |
declare fun_quotient[quot_map] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
671 |
declare fun_mono[relator_mono] |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51112
diff
changeset
|
672 |
lemmas [relator_distr] = pos_fun_distr neg_fun_distr1 neg_fun_distr2 |
47308 | 673 |
|
48891 | 674 |
ML_file "Tools/Lifting/lifting_term.ML" |
47308 | 675 |
|
48891 | 676 |
ML_file "Tools/Lifting/lifting_def.ML" |
47308 | 677 |
|
48891 | 678 |
ML_file "Tools/Lifting/lifting_setup.ML" |
47308 | 679 |
|
51994 | 680 |
hide_const (open) invariant POS NEG reflp' |
47308 | 681 |
|
682 |
end |