| author | wenzelm | 
| Sat, 22 Mar 2014 18:16:54 +0100 | |
| changeset 56253 | 83b3c110f22d | 
| parent 55945 | e96383acecf9 | 
| child 56517 | 7e8a369eb10d | 
| 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"  | 
|
| 55945 | 55  | 
using assms unfolding left_unique_def[abs_def] right_total_def bi_unique_def rel_fun_def  | 
| 53952 | 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)"  | 
| 55945 | 72  | 
unfolding left_total_def rel_fun_def  | 
| 
52036
 
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)"  | 
| 55945 | 86  | 
unfolding left_total_def left_unique_def rel_fun_def  | 
| 
52036
 
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  | 
||
| 
55610
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55604 
diff
changeset
 | 
158  | 
lemma Quotient_rep_abs_eq: "R t t \<Longrightarrow> R \<le> op= \<Longrightarrow> Rep (Abs t) = t"  | 
| 
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55604 
diff
changeset
 | 
159  | 
using a unfolding Quotient_def  | 
| 
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55604 
diff
changeset
 | 
160  | 
by blast  | 
| 
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55604 
diff
changeset
 | 
161  | 
|
| 
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
 | 
162  | 
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
 | 
163  | 
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
 | 
164  | 
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
 | 
165  | 
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
 | 
166  | 
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
 | 
167  | 
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
 | 
168  | 
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
 | 
169  | 
|
| 
 
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  | 
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
 | 
171  | 
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
 | 
172  | 
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
 | 
173  | 
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
 | 
174  | 
|
| 47536 | 175  | 
lemma Quotient_rel_abs: "R r s \<Longrightarrow> Abs r = Abs s"  | 
176  | 
using a unfolding Quotient_def  | 
|
177  | 
by blast  | 
|
178  | 
||
| 
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
 | 
179  | 
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
 | 
180  | 
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
 | 
181  | 
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
 | 
182  | 
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
 | 
183  | 
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
 | 
184  | 
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
 | 
185  | 
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
 | 
186  | 
|
| 47536 | 187  | 
lemma Quotient_symp: "symp R"  | 
| 47308 | 188  | 
using a unfolding Quotient_def using sympI by (metis (full_types))  | 
189  | 
||
| 47536 | 190  | 
lemma Quotient_transp: "transp R"  | 
| 47308 | 191  | 
using a unfolding Quotient_def using transpI by (metis (full_types))  | 
192  | 
||
| 47536 | 193  | 
lemma Quotient_part_equivp: "part_equivp R"  | 
194  | 
by (metis Quotient_rep_reflp Quotient_symp Quotient_transp part_equivpI)  | 
|
195  | 
||
196  | 
end  | 
|
| 47308 | 197  | 
|
198  | 
lemma identity_quotient: "Quotient (op =) id id (op =)"  | 
|
199  | 
unfolding Quotient_def by simp  | 
|
200  | 
||
| 
47652
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
201  | 
text {* TODO: Use one of these alternatives as the real definition. *}
 | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
202  | 
|
| 47308 | 203  | 
lemma Quotient_alt_def:  | 
204  | 
"Quotient R Abs Rep T \<longleftrightarrow>  | 
|
205  | 
(\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and>  | 
|
206  | 
(\<forall>b. T (Rep b) b) \<and>  | 
|
207  | 
(\<forall>x y. R x y \<longleftrightarrow> T x (Abs x) \<and> T y (Abs y) \<and> Abs x = Abs y)"  | 
|
208  | 
apply safe  | 
|
209  | 
apply (simp (no_asm_use) only: Quotient_def, fast)  | 
|
210  | 
apply (simp (no_asm_use) only: Quotient_def, fast)  | 
|
211  | 
apply (simp (no_asm_use) only: Quotient_def, fast)  | 
|
212  | 
apply (simp (no_asm_use) only: Quotient_def, fast)  | 
|
213  | 
apply (simp (no_asm_use) only: Quotient_def, fast)  | 
|
214  | 
apply (simp (no_asm_use) only: Quotient_def, fast)  | 
|
215  | 
apply (rule QuotientI)  | 
|
216  | 
apply simp  | 
|
217  | 
apply metis  | 
|
218  | 
apply simp  | 
|
219  | 
apply (rule ext, rule ext, metis)  | 
|
220  | 
done  | 
|
221  | 
||
222  | 
lemma Quotient_alt_def2:  | 
|
223  | 
"Quotient R Abs Rep T \<longleftrightarrow>  | 
|
224  | 
(\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and>  | 
|
225  | 
(\<forall>b. T (Rep b) b) \<and>  | 
|
226  | 
(\<forall>x y. R x y \<longleftrightarrow> T x (Abs y) \<and> T y (Abs x))"  | 
|
227  | 
unfolding Quotient_alt_def by (safe, metis+)  | 
|
228  | 
||
| 
47652
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
229  | 
lemma Quotient_alt_def3:  | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
230  | 
"Quotient R Abs Rep T \<longleftrightarrow>  | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
231  | 
(\<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
 | 
232  | 
(\<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
 | 
233  | 
unfolding Quotient_alt_def2 by (safe, metis+)  | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
234  | 
|
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
235  | 
lemma Quotient_alt_def4:  | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
236  | 
"Quotient R Abs Rep T \<longleftrightarrow>  | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
237  | 
(\<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
 | 
238  | 
unfolding Quotient_alt_def3 fun_eq_iff by auto  | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
239  | 
|
| 47308 | 240  | 
lemma fun_quotient:  | 
241  | 
assumes 1: "Quotient R1 abs1 rep1 T1"  | 
|
242  | 
assumes 2: "Quotient R2 abs2 rep2 T2"  | 
|
243  | 
shows "Quotient (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2) (T1 ===> T2)"  | 
|
244  | 
using assms unfolding Quotient_alt_def2  | 
|
| 55945 | 245  | 
unfolding rel_fun_def fun_eq_iff map_fun_apply  | 
| 47308 | 246  | 
by (safe, metis+)  | 
247  | 
||
248  | 
lemma apply_rsp:  | 
|
249  | 
fixes f g::"'a \<Rightarrow> 'c"  | 
|
250  | 
assumes q: "Quotient R1 Abs1 Rep1 T1"  | 
|
251  | 
and a: "(R1 ===> R2) f g" "R1 x y"  | 
|
252  | 
shows "R2 (f x) (g y)"  | 
|
| 55945 | 253  | 
using a by (auto elim: rel_funE)  | 
| 47308 | 254  | 
|
255  | 
lemma apply_rsp':  | 
|
256  | 
assumes a: "(R1 ===> R2) f g" "R1 x y"  | 
|
257  | 
shows "R2 (f x) (g y)"  | 
|
| 55945 | 258  | 
using a by (auto elim: rel_funE)  | 
| 47308 | 259  | 
|
260  | 
lemma apply_rsp'':  | 
|
261  | 
assumes "Quotient R Abs Rep T"  | 
|
262  | 
and "(R ===> S) f f"  | 
|
263  | 
shows "S (f (Rep x)) (f (Rep x))"  | 
|
264  | 
proof -  | 
|
265  | 
from assms(1) have "R (Rep x) (Rep x)" by (rule Quotient_rep_reflp)  | 
|
266  | 
then show ?thesis using assms(2) by (auto intro: apply_rsp')  | 
|
267  | 
qed  | 
|
268  | 
||
269  | 
subsection {* Quotient composition *}
 | 
|
270  | 
||
271  | 
lemma Quotient_compose:  | 
|
272  | 
assumes 1: "Quotient R1 Abs1 Rep1 T1"  | 
|
273  | 
assumes 2: "Quotient R2 Abs2 Rep2 T2"  | 
|
274  | 
shows "Quotient (T1 OO R2 OO conversep T1) (Abs2 \<circ> Abs1) (Rep1 \<circ> Rep2) (T1 OO T2)"  | 
|
| 51994 | 275  | 
using assms unfolding Quotient_alt_def4 by fastforce  | 
| 47308 | 276  | 
|
| 
47521
 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 
kuncar 
parents: 
47501 
diff
changeset
 | 
277  | 
lemma equivp_reflp2:  | 
| 
 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 
kuncar 
parents: 
47501 
diff
changeset
 | 
278  | 
"equivp R \<Longrightarrow> reflp R"  | 
| 
 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 
kuncar 
parents: 
47501 
diff
changeset
 | 
279  | 
by (erule equivpE)  | 
| 
 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 
kuncar 
parents: 
47501 
diff
changeset
 | 
280  | 
|
| 47544 | 281  | 
subsection {* Respects predicate *}
 | 
282  | 
||
283  | 
definition Respects :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set"
 | 
|
284  | 
  where "Respects R = {x. R x x}"
 | 
|
285  | 
||
286  | 
lemma in_respects: "x \<in> Respects R \<longleftrightarrow> R x x"  | 
|
287  | 
unfolding Respects_def by simp  | 
|
288  | 
||
| 47308 | 289  | 
subsection {* Invariant *}
 | 
290  | 
||
291  | 
definition invariant :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" 
 | 
|
292  | 
where "invariant R = (\<lambda>x y. R x \<and> x = y)"  | 
|
293  | 
||
294  | 
lemma invariant_to_eq:  | 
|
295  | 
assumes "invariant P x y"  | 
|
296  | 
shows "x = y"  | 
|
297  | 
using assms by (simp add: invariant_def)  | 
|
298  | 
||
| 55945 | 299  | 
lemma rel_fun_eq_invariant: "(op= ===> Lifting.invariant P) = Lifting.invariant (\<lambda>f. \<forall>x. P(f x))"  | 
300  | 
unfolding invariant_def rel_fun_def by auto  | 
|
| 55737 | 301  | 
|
| 55945 | 302  | 
lemma rel_fun_invariant_rel:  | 
| 47308 | 303  | 
shows "((invariant R) ===> S) = (\<lambda>f g. \<forall>x. R x \<longrightarrow> S (f x) (g x))"  | 
| 55945 | 304  | 
by (auto simp add: invariant_def rel_fun_def)  | 
| 47308 | 305  | 
|
306  | 
lemma invariant_same_args:  | 
|
307  | 
shows "invariant P x x \<equiv> P x"  | 
|
308  | 
using assms by (auto simp add: invariant_def)  | 
|
309  | 
||
| 53952 | 310  | 
lemma invariant_transfer [transfer_rule]:  | 
311  | 
assumes [transfer_rule]: "bi_unique A"  | 
|
312  | 
shows "((A ===> op=) ===> A ===> A ===> op=) Lifting.invariant Lifting.invariant"  | 
|
313  | 
unfolding invariant_def[abs_def] by transfer_prover  | 
|
314  | 
||
| 
47361
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
315  | 
lemma UNIV_typedef_to_Quotient:  | 
| 47308 | 316  | 
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
 | 
317  | 
and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"  | 
| 47308 | 318  | 
shows "Quotient (op =) Abs Rep T"  | 
319  | 
proof -  | 
|
320  | 
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
 | 
321  | 
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
 | 
322  | 
by (fastforce intro!: QuotientI fun_eq_iff)  | 
| 47308 | 323  | 
qed  | 
324  | 
||
| 
47361
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
325  | 
lemma UNIV_typedef_to_equivp:  | 
| 47308 | 326  | 
fixes Abs :: "'a \<Rightarrow> 'b"  | 
327  | 
and Rep :: "'b \<Rightarrow> 'a"  | 
|
328  | 
assumes "type_definition Rep Abs (UNIV::'a set)"  | 
|
329  | 
shows "equivp (op=::'a\<Rightarrow>'a\<Rightarrow>bool)"  | 
|
330  | 
by (rule identity_equivp)  | 
|
331  | 
||
| 47354 | 332  | 
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
 | 
333  | 
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
 | 
334  | 
and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"  | 
| 47501 | 335  | 
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
 | 
336  | 
proof -  | 
| 
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
337  | 
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
 | 
338  | 
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
 | 
339  | 
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
 | 
340  | 
qed  | 
| 
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
341  | 
|
| 
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
342  | 
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
 | 
343  | 
assumes "type_definition Rep Abs S"  | 
| 47501 | 344  | 
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
 | 
345  | 
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
 | 
346  | 
interpret type_definition Rep Abs S by fact  | 
| 47501 | 347  | 
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
 | 
348  | 
next  | 
| 47501 | 349  | 
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
 | 
350  | 
next  | 
| 47501 | 351  | 
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
 | 
352  | 
qed  | 
| 
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
353  | 
|
| 
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
354  | 
lemma open_typedef_to_Quotient:  | 
| 47308 | 355  | 
  assumes "type_definition Rep Abs {x. P x}"
 | 
| 47354 | 356  | 
and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"  | 
| 47308 | 357  | 
shows "Quotient (invariant P) Abs Rep T"  | 
| 47651 | 358  | 
using typedef_to_Quotient [OF assms] by simp  | 
| 47308 | 359  | 
|
| 
47361
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
360  | 
lemma open_typedef_to_part_equivp:  | 
| 47308 | 361  | 
  assumes "type_definition Rep Abs {x. P x}"
 | 
362  | 
shows "part_equivp (invariant P)"  | 
|
| 47651 | 363  | 
using typedef_to_part_equivp [OF assms] by simp  | 
| 47308 | 364  | 
|
| 47376 | 365  | 
text {* Generating transfer rules for quotients. *}
 | 
366  | 
||
| 47537 | 367  | 
context  | 
368  | 
fixes R Abs Rep T  | 
|
369  | 
assumes 1: "Quotient R Abs Rep T"  | 
|
370  | 
begin  | 
|
| 47376 | 371  | 
|
| 47537 | 372  | 
lemma Quotient_right_unique: "right_unique T"  | 
373  | 
using 1 unfolding Quotient_alt_def right_unique_def by metis  | 
|
374  | 
||
375  | 
lemma Quotient_right_total: "right_total T"  | 
|
376  | 
using 1 unfolding Quotient_alt_def right_total_def by metis  | 
|
377  | 
||
378  | 
lemma Quotient_rel_eq_transfer: "(T ===> T ===> op =) R (op =)"  | 
|
| 55945 | 379  | 
using 1 unfolding Quotient_alt_def rel_fun_def by simp  | 
| 47376 | 380  | 
|
| 47538 | 381  | 
lemma Quotient_abs_induct:  | 
382  | 
assumes "\<And>y. R y y \<Longrightarrow> P (Abs y)" shows "P x"  | 
|
383  | 
using 1 assms unfolding Quotient_def by metis  | 
|
384  | 
||
| 47537 | 385  | 
end  | 
386  | 
||
387  | 
text {* Generating transfer rules for total quotients. *}
 | 
|
| 47376 | 388  | 
|
| 47537 | 389  | 
context  | 
390  | 
fixes R Abs Rep T  | 
|
391  | 
assumes 1: "Quotient R Abs Rep T" and 2: "reflp R"  | 
|
392  | 
begin  | 
|
| 47376 | 393  | 
|
| 47537 | 394  | 
lemma Quotient_bi_total: "bi_total T"  | 
395  | 
using 1 2 unfolding Quotient_alt_def bi_total_def reflp_def by auto  | 
|
396  | 
||
397  | 
lemma Quotient_id_abs_transfer: "(op = ===> T) (\<lambda>x. x) Abs"  | 
|
| 55945 | 398  | 
using 1 2 unfolding Quotient_alt_def reflp_def rel_fun_def by simp  | 
| 47537 | 399  | 
|
| 47575 | 400  | 
lemma Quotient_total_abs_induct: "(\<And>y. P (Abs y)) \<Longrightarrow> P x"  | 
401  | 
using 1 2 assms unfolding Quotient_alt_def reflp_def by metis  | 
|
402  | 
||
| 
47889
 
29212a4bb866
lifting package produces abs_eq_iff rules for total quotients
 
huffman 
parents: 
47777 
diff
changeset
 | 
403  | 
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
 | 
404  | 
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
 | 
405  | 
|
| 47537 | 406  | 
end  | 
| 47376 | 407  | 
|
| 
47368
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
408  | 
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
 | 
409  | 
|
| 
47534
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
410  | 
context  | 
| 
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
411  | 
fixes Rep Abs A T  | 
| 
47368
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
412  | 
assumes type: "type_definition Rep Abs A"  | 
| 
47534
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
413  | 
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
 | 
414  | 
begin  | 
| 
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
415  | 
|
| 51994 | 416  | 
lemma typedef_left_unique: "left_unique T"  | 
417  | 
unfolding left_unique_def T_def  | 
|
418  | 
by (simp add: type_definition.Rep_inject [OF type])  | 
|
419  | 
||
| 
47534
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
420  | 
lemma typedef_bi_unique: "bi_unique T"  | 
| 
47368
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
421  | 
unfolding bi_unique_def T_def  | 
| 
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
422  | 
by (simp add: type_definition.Rep_inject [OF type])  | 
| 
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
423  | 
|
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
424  | 
(* 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
 | 
425  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
426  | 
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
 | 
427  | 
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
 | 
428  | 
by blast  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
429  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
430  | 
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
 | 
431  | 
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
 | 
432  | 
by blast  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
433  | 
|
| 
47535
 
0f94b02fda1c
lifting_setup generates transfer rule for rep of typedefs
 
huffman 
parents: 
47534 
diff
changeset
 | 
434  | 
lemma typedef_rep_transfer: "(T ===> op =) (\<lambda>x. x) Rep"  | 
| 55945 | 435  | 
unfolding rel_fun_def T_def by simp  | 
| 
47535
 
0f94b02fda1c
lifting_setup generates transfer rule for rep of typedefs
 
huffman 
parents: 
47534 
diff
changeset
 | 
436  | 
|
| 
47534
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
437  | 
end  | 
| 
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
438  | 
|
| 
47368
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
439  | 
text {* Generating the correspondence rule for a constant defined with
 | 
| 
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
440  | 
  @{text "lift_definition"}. *}
 | 
| 
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
441  | 
|
| 47351 | 442  | 
lemma Quotient_to_transfer:  | 
443  | 
assumes "Quotient R Abs Rep T" and "R c c" and "c' \<equiv> Abs c"  | 
|
444  | 
shows "T c c'"  | 
|
445  | 
using assms by (auto dest: Quotient_cr_rel)  | 
|
446  | 
||
| 
47982
 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 
kuncar 
parents: 
47937 
diff
changeset
 | 
447  | 
text {* Proving reflexivity *}
 | 
| 
 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 
kuncar 
parents: 
47937 
diff
changeset
 | 
448  | 
|
| 
 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 
kuncar 
parents: 
47937 
diff
changeset
 | 
449  | 
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
 | 
450  | 
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
 | 
451  | 
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
 | 
452  | 
shows "left_total T"  | 
| 
 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 
kuncar 
parents: 
47937 
diff
changeset
 | 
453  | 
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
 | 
454  | 
|
| 
55563
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
455  | 
lemma Quotient_composition_ge_eq:  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
456  | 
assumes "left_total T"  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
457  | 
assumes "R \<ge> op="  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
458  | 
shows "(T OO R OO T\<inverse>\<inverse>) \<ge> op="  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
459  | 
using assms unfolding left_total_def by fast  | 
| 51994 | 460  | 
|
| 
55563
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
461  | 
lemma Quotient_composition_le_eq:  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
462  | 
assumes "left_unique T"  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
463  | 
assumes "R \<le> op="  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
464  | 
shows "(T OO R OO T\<inverse>\<inverse>) \<le> op="  | 
| 55604 | 465  | 
using assms unfolding left_unique_def by blast  | 
| 
47982
 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 
kuncar 
parents: 
47937 
diff
changeset
 | 
466  | 
|
| 52307 | 467  | 
lemma left_total_composition: "left_total R \<Longrightarrow> left_total S \<Longrightarrow> left_total (R OO S)"  | 
468  | 
unfolding left_total_def OO_def by fast  | 
|
469  | 
||
470  | 
lemma left_unique_composition: "left_unique R \<Longrightarrow> left_unique S \<Longrightarrow> left_unique (R OO S)"  | 
|
| 55604 | 471  | 
unfolding left_unique_def OO_def by blast  | 
| 52307 | 472  | 
|
| 
55563
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
473  | 
lemma invariant_le_eq:  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
474  | 
"invariant P \<le> op=" unfolding invariant_def by blast  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
475  | 
|
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
476  | 
lemma reflp_ge_eq:  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
477  | 
"reflp R \<Longrightarrow> R \<ge> op=" unfolding reflp_def by blast  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
478  | 
|
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
479  | 
lemma ge_eq_refl:  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
480  | 
"R \<ge> op= \<Longrightarrow> R x x" by blast  | 
| 
47982
 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 
kuncar 
parents: 
47937 
diff
changeset
 | 
481  | 
|
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
482  | 
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
 | 
483  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
484  | 
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
 | 
485  | 
"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
 | 
486  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
487  | 
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
 | 
488  | 
"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
 | 
489  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
490  | 
(*  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
491  | 
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
 | 
492  | 
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
 | 
493  | 
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
 | 
494  | 
*)  | 
| 
 
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  | 
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
 | 
497  | 
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
 | 
498  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
499  | 
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
 | 
500  | 
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
 | 
501  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
502  | 
lemma pos_OO_eq:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
503  | 
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
 | 
504  | 
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
 | 
505  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
506  | 
lemma pos_eq_OO:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
507  | 
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
 | 
508  | 
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
 | 
509  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
510  | 
lemma neg_OO_eq:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
511  | 
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
 | 
512  | 
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
 | 
513  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
514  | 
lemma neg_eq_OO:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
515  | 
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
 | 
516  | 
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
 | 
517  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
518  | 
lemma POS_trans:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
519  | 
assumes "POS A B"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
520  | 
assumes "POS B C"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
521  | 
shows "POS A C"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
522  | 
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
 | 
523  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
524  | 
lemma NEG_trans:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
525  | 
assumes "NEG A B"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
526  | 
assumes "NEG B C"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
527  | 
shows "NEG 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 NEG_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 POS_NEG:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
531  | 
"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
 | 
532  | 
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
 | 
533  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
534  | 
lemma NEG_POS:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
535  | 
"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
 | 
536  | 
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
 | 
537  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
538  | 
lemma POS_pcr_rule:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
539  | 
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
 | 
540  | 
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
 | 
541  | 
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
 | 
542  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
543  | 
lemma NEG_pcr_rule:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
544  | 
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
 | 
545  | 
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
 | 
546  | 
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
 | 
547  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
548  | 
lemma POS_apply:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
549  | 
assumes "POS R R'"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
550  | 
assumes "R f g"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
551  | 
shows "R' f g"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
552  | 
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
 | 
553  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
554  | 
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
 | 
555  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
556  | 
lemma fun_mono:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
557  | 
assumes "A \<ge> C"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
558  | 
assumes "B \<le> D"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
559  | 
shows "(A ===> B) \<le> (C ===> D)"  | 
| 55945 | 560  | 
using assms unfolding rel_fun_def by blast  | 
| 
51374
 
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 pos_fun_distr: "((R ===> S) OO (R' ===> S')) \<le> ((R OO R') ===> (S OO S'))"  | 
| 55945 | 563  | 
unfolding OO_def rel_fun_def by blast  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
564  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
565  | 
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
 | 
566  | 
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
 | 
567  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
568  | 
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
 | 
569  | 
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
 | 
570  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
571  | 
lemma neg_fun_distr1:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
572  | 
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
 | 
573  | 
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
 | 
574  | 
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
 | 
575  | 
using functional_relation[OF 2] functional_converse_relation[OF 1]  | 
| 55945 | 576  | 
unfolding rel_fun_def OO_def  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
577  | 
apply clarify  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
578  | 
apply (subst all_comm)  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
579  | 
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
 | 
580  | 
apply (intro choice)  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
581  | 
by metis  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
582  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
583  | 
lemma neg_fun_distr2:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
584  | 
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
 | 
585  | 
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
 | 
586  | 
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
 | 
587  | 
using functional_converse_relation[OF 2] functional_relation[OF 1]  | 
| 55945 | 588  | 
unfolding rel_fun_def OO_def  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
589  | 
apply clarify  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
590  | 
apply (subst all_comm)  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
591  | 
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
 | 
592  | 
apply (intro choice)  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
593  | 
by metis  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
594  | 
|
| 
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
 | 
595  | 
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
 | 
596  | 
|
| 
55731
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
597  | 
lemma composed_equiv_rel_invariant:  | 
| 
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
598  | 
assumes "left_unique R"  | 
| 
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
599  | 
assumes "(R ===> op=) P P'"  | 
| 
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
600  | 
assumes "Domainp R = P''"  | 
| 
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
601  | 
shows "(R OO Lifting.invariant P' OO R\<inverse>\<inverse>) = Lifting.invariant (inf P'' P)"  | 
| 55945 | 602  | 
using assms unfolding OO_def conversep_iff Domainp_iff[abs_def] left_unique_def rel_fun_def invariant_def  | 
| 
55731
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
603  | 
fun_eq_iff by blast  | 
| 
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
604  | 
|
| 
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
605  | 
lemma composed_equiv_rel_eq_invariant:  | 
| 
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
606  | 
assumes "left_unique R"  | 
| 
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
607  | 
assumes "Domainp R = P"  | 
| 
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
608  | 
shows "(R OO op= OO R\<inverse>\<inverse>) = Lifting.invariant P"  | 
| 
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
609  | 
using assms unfolding OO_def conversep_iff Domainp_iff[abs_def] left_unique_def invariant_def  | 
| 
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
610  | 
fun_eq_iff is_equality_def by metis  | 
| 
 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 
kuncar 
parents: 
55610 
diff
changeset
 | 
611  | 
|
| 
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
 | 
612  | 
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
 | 
613  | 
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
 | 
614  | 
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
 | 
615  | 
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
 | 
616  | 
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
 | 
617  | 
using assms  | 
| 55945 | 618  | 
unfolding Domainp_iff[abs_def] OO_def bi_unique_def left_total_def rel_fun_def  | 
| 
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
 | 
619  | 
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
 | 
620  | 
|
| 
 
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  | 
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
 | 
622  | 
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
 | 
623  | 
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
 | 
624  | 
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
 | 
625  | 
shows "Domainp (A OO B) = (inf P1 P2')"  | 
| 55945 | 626  | 
using assms unfolding rel_fun_def Domainp_iff[abs_def] OO_def  | 
| 
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  | 
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
 | 
628  | 
|
| 53151 | 629  | 
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
 | 
630  | 
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
 | 
631  | 
|
| 
 
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  | 
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
 | 
633  | 
assumes "Domainp B = P"  | 
| 53151 | 634  | 
shows "Domainp (A OO B) = (\<lambda>x. \<exists>y. A x y \<and> P y)"  | 
635  | 
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
 | 
636  | 
|
| 
 
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  | 
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
 | 
638  | 
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
 | 
639  | 
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
 | 
640  | 
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
 | 
641  | 
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
 | 
642  | 
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
 | 
643  | 
|
| 
 
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  | 
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
 | 
645  | 
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
 | 
646  | 
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
 | 
647  | 
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
 | 
648  | 
|
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51374 
diff
changeset
 | 
649  | 
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
 | 
650  | 
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
 | 
651  | 
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
 | 
652  | 
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
 | 
653  | 
|
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52307 
diff
changeset
 | 
654  | 
end  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52307 
diff
changeset
 | 
655  | 
|
| 47308 | 656  | 
subsection {* ML setup *}
 | 
657  | 
||
| 48891 | 658  | 
ML_file "Tools/Lifting/lifting_util.ML"  | 
| 47308 | 659  | 
|
| 48891 | 660  | 
ML_file "Tools/Lifting/lifting_info.ML"  | 
| 47308 | 661  | 
setup Lifting_Info.setup  | 
662  | 
||
| 51994 | 663  | 
lemmas [reflexivity_rule] =  | 
| 
55563
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
664  | 
order_refl[of "op="] invariant_le_eq Quotient_composition_le_eq  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
665  | 
Quotient_composition_ge_eq  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
666  | 
left_total_eq left_unique_eq left_total_composition left_unique_composition  | 
| 
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
667  | 
left_total_fun left_unique_fun  | 
| 
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  | 
|
| 
55563
 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 
kuncar 
parents: 
55083 
diff
changeset
 | 
680  | 
hide_const (open) invariant POS NEG  | 
| 47308 | 681  | 
|
682  | 
end  |