| author | wenzelm | 
| Mon, 17 Apr 2017 19:44:13 +0200 | |
| changeset 65495 | 60d4fbed2b1f | 
| parent 63343 | fb5d8a50c641 | 
| child 67229 | 4ecf0ef70efb | 
| 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  | 
||
| 60758 | 6  | 
section \<open>Lifting package\<close>  | 
| 47308 | 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  | 
||
| 60758 | 17  | 
subsection \<open>Function map\<close>  | 
| 47308 | 18  | 
|
| 63343 | 19  | 
context includes lifting_syntax  | 
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52307 
diff
changeset
 | 
20  | 
begin  | 
| 47308 | 21  | 
|
22  | 
lemma map_fun_id:  | 
|
23  | 
"(id ---> id) = id"  | 
|
24  | 
by (simp add: fun_eq_iff)  | 
|
25  | 
||
| 60758 | 26  | 
subsection \<open>Quotient Predicate\<close>  | 
| 47308 | 27  | 
|
28  | 
definition  | 
|
29  | 
"Quotient R Abs Rep T \<longleftrightarrow>  | 
|
| 58186 | 30  | 
(\<forall>a. Abs (Rep a) = a) \<and>  | 
| 47308 | 31  | 
(\<forall>a. R (Rep a) (Rep a)) \<and>  | 
32  | 
(\<forall>r s. R r s \<longleftrightarrow> R r r \<and> R s s \<and> Abs r = Abs s) \<and>  | 
|
33  | 
T = (\<lambda>x y. R x x \<and> Abs x = y)"  | 
|
34  | 
||
35  | 
lemma QuotientI:  | 
|
36  | 
assumes "\<And>a. Abs (Rep a) = a"  | 
|
37  | 
and "\<And>a. R (Rep a) (Rep a)"  | 
|
38  | 
and "\<And>r s. R r s \<longleftrightarrow> R r r \<and> R s s \<and> Abs r = Abs s"  | 
|
39  | 
and "T = (\<lambda>x y. R x x \<and> Abs x = y)"  | 
|
40  | 
shows "Quotient R Abs Rep T"  | 
|
41  | 
using assms unfolding Quotient_def by blast  | 
|
42  | 
||
| 47536 | 43  | 
context  | 
44  | 
fixes R Abs Rep T  | 
|
| 47308 | 45  | 
assumes a: "Quotient R Abs Rep T"  | 
| 47536 | 46  | 
begin  | 
47  | 
||
48  | 
lemma Quotient_abs_rep: "Abs (Rep a) = a"  | 
|
49  | 
using a unfolding Quotient_def  | 
|
| 47308 | 50  | 
by simp  | 
51  | 
||
| 47536 | 52  | 
lemma Quotient_rep_reflp: "R (Rep a) (Rep a)"  | 
53  | 
using a unfolding Quotient_def  | 
|
| 47308 | 54  | 
by blast  | 
55  | 
||
56  | 
lemma Quotient_rel:  | 
|
| 61799 | 57  | 
"R r r \<and> R s s \<and> Abs r = Abs s \<longleftrightarrow> R r s" \<comment> \<open>orientation does not loop on rewriting\<close>  | 
| 47536 | 58  | 
using a unfolding Quotient_def  | 
| 47308 | 59  | 
by blast  | 
60  | 
||
| 47536 | 61  | 
lemma Quotient_cr_rel: "T = (\<lambda>x y. R x x \<and> Abs x = y)"  | 
| 47308 | 62  | 
using a unfolding Quotient_def  | 
63  | 
by blast  | 
|
64  | 
||
| 47536 | 65  | 
lemma Quotient_refl1: "R r s \<Longrightarrow> R r r"  | 
66  | 
using a unfolding Quotient_def  | 
|
67  | 
by fast  | 
|
68  | 
||
69  | 
lemma Quotient_refl2: "R r s \<Longrightarrow> R s s"  | 
|
70  | 
using a unfolding Quotient_def  | 
|
71  | 
by fast  | 
|
72  | 
||
73  | 
lemma Quotient_rel_rep: "R (Rep a) (Rep b) \<longleftrightarrow> a = b"  | 
|
74  | 
using a unfolding Quotient_def  | 
|
75  | 
by metis  | 
|
76  | 
||
77  | 
lemma Quotient_rep_abs: "R r r \<Longrightarrow> R (Rep (Abs r)) r"  | 
|
| 47308 | 78  | 
using a unfolding Quotient_def  | 
79  | 
by blast  | 
|
80  | 
||
| 
55610
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55604 
diff
changeset
 | 
81  | 
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
 | 
82  | 
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
 | 
83  | 
by blast  | 
| 
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55604 
diff
changeset
 | 
84  | 
|
| 58186 | 85  | 
lemma Quotient_rep_abs_fold_unmap:  | 
86  | 
assumes "x' \<equiv> Abs x" and "R x x" and "Rep x' \<equiv> Rep' x'"  | 
|
| 
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
 | 
87  | 
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
 | 
88  | 
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
 | 
89  | 
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
 | 
90  | 
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
 | 
91  | 
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
 | 
92  | 
|
| 
 
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
 | 
93  | 
lemma Quotient_Rep_eq:  | 
| 58186 | 94  | 
assumes "x' \<equiv> Abs x"  | 
| 
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
 | 
95  | 
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
 | 
96  | 
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
 | 
97  | 
|
| 47536 | 98  | 
lemma Quotient_rel_abs: "R r s \<Longrightarrow> Abs r = Abs s"  | 
99  | 
using a unfolding Quotient_def  | 
|
100  | 
by blast  | 
|
101  | 
||
| 
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
 | 
102  | 
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
 | 
103  | 
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
 | 
104  | 
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
 | 
105  | 
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
 | 
106  | 
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
 | 
107  | 
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
 | 
108  | 
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
 | 
109  | 
|
| 47536 | 110  | 
lemma Quotient_symp: "symp R"  | 
| 47308 | 111  | 
using a unfolding Quotient_def using sympI by (metis (full_types))  | 
112  | 
||
| 47536 | 113  | 
lemma Quotient_transp: "transp R"  | 
| 47308 | 114  | 
using a unfolding Quotient_def using transpI by (metis (full_types))  | 
115  | 
||
| 47536 | 116  | 
lemma Quotient_part_equivp: "part_equivp R"  | 
117  | 
by (metis Quotient_rep_reflp Quotient_symp Quotient_transp part_equivpI)  | 
|
118  | 
||
119  | 
end  | 
|
| 47308 | 120  | 
|
121  | 
lemma identity_quotient: "Quotient (op =) id id (op =)"  | 
|
| 58186 | 122  | 
unfolding Quotient_def by simp  | 
| 47308 | 123  | 
|
| 60758 | 124  | 
text \<open>TODO: Use one of these alternatives as the real definition.\<close>  | 
| 
47652
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
125  | 
|
| 47308 | 126  | 
lemma Quotient_alt_def:  | 
127  | 
"Quotient R Abs Rep T \<longleftrightarrow>  | 
|
128  | 
(\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and>  | 
|
129  | 
(\<forall>b. T (Rep b) b) \<and>  | 
|
130  | 
(\<forall>x y. R x y \<longleftrightarrow> T x (Abs x) \<and> T y (Abs y) \<and> Abs x = Abs y)"  | 
|
131  | 
apply safe  | 
|
132  | 
apply (simp (no_asm_use) only: Quotient_def, fast)  | 
|
133  | 
apply (simp (no_asm_use) only: Quotient_def, fast)  | 
|
134  | 
apply (simp (no_asm_use) only: Quotient_def, fast)  | 
|
135  | 
apply (simp (no_asm_use) only: Quotient_def, fast)  | 
|
136  | 
apply (simp (no_asm_use) only: Quotient_def, fast)  | 
|
137  | 
apply (simp (no_asm_use) only: Quotient_def, fast)  | 
|
138  | 
apply (rule QuotientI)  | 
|
139  | 
apply simp  | 
|
140  | 
apply metis  | 
|
141  | 
apply simp  | 
|
142  | 
apply (rule ext, rule ext, metis)  | 
|
143  | 
done  | 
|
144  | 
||
145  | 
lemma Quotient_alt_def2:  | 
|
146  | 
"Quotient R Abs Rep T \<longleftrightarrow>  | 
|
147  | 
(\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and>  | 
|
148  | 
(\<forall>b. T (Rep b) b) \<and>  | 
|
149  | 
(\<forall>x y. R x y \<longleftrightarrow> T x (Abs y) \<and> T y (Abs x))"  | 
|
150  | 
unfolding Quotient_alt_def by (safe, metis+)  | 
|
151  | 
||
| 
47652
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
152  | 
lemma Quotient_alt_def3:  | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
153  | 
"Quotient R Abs Rep T \<longleftrightarrow>  | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
154  | 
(\<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
 | 
155  | 
(\<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
 | 
156  | 
unfolding Quotient_alt_def2 by (safe, metis+)  | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
157  | 
|
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
158  | 
lemma Quotient_alt_def4:  | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
159  | 
"Quotient R Abs Rep T \<longleftrightarrow>  | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
160  | 
(\<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
 | 
161  | 
unfolding Quotient_alt_def3 fun_eq_iff by auto  | 
| 
 
1b722b100301
move alternative definition lemmas into Lifting.thy;
 
huffman 
parents: 
47651 
diff
changeset
 | 
162  | 
|
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56519 
diff
changeset
 | 
163  | 
lemma Quotient_alt_def5:  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56519 
diff
changeset
 | 
164  | 
"Quotient R Abs Rep T \<longleftrightarrow>  | 
| 57398 | 165  | 
T \<le> BNF_Def.Grp UNIV Abs \<and> BNF_Def.Grp UNIV Rep \<le> T\<inverse>\<inverse> \<and> R = T OO T\<inverse>\<inverse>"  | 
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56519 
diff
changeset
 | 
166  | 
unfolding Quotient_alt_def4 Grp_def by blast  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56519 
diff
changeset
 | 
167  | 
|
| 47308 | 168  | 
lemma fun_quotient:  | 
169  | 
assumes 1: "Quotient R1 abs1 rep1 T1"  | 
|
170  | 
assumes 2: "Quotient R2 abs2 rep2 T2"  | 
|
171  | 
shows "Quotient (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2) (T1 ===> T2)"  | 
|
172  | 
using assms unfolding Quotient_alt_def2  | 
|
| 55945 | 173  | 
unfolding rel_fun_def fun_eq_iff map_fun_apply  | 
| 47308 | 174  | 
by (safe, metis+)  | 
175  | 
||
176  | 
lemma apply_rsp:  | 
|
177  | 
fixes f g::"'a \<Rightarrow> 'c"  | 
|
178  | 
assumes q: "Quotient R1 Abs1 Rep1 T1"  | 
|
179  | 
and a: "(R1 ===> R2) f g" "R1 x y"  | 
|
180  | 
shows "R2 (f x) (g y)"  | 
|
| 55945 | 181  | 
using a by (auto elim: rel_funE)  | 
| 47308 | 182  | 
|
183  | 
lemma apply_rsp':  | 
|
184  | 
assumes a: "(R1 ===> R2) f g" "R1 x y"  | 
|
185  | 
shows "R2 (f x) (g y)"  | 
|
| 55945 | 186  | 
using a by (auto elim: rel_funE)  | 
| 47308 | 187  | 
|
188  | 
lemma apply_rsp'':  | 
|
189  | 
assumes "Quotient R Abs Rep T"  | 
|
190  | 
and "(R ===> S) f f"  | 
|
191  | 
shows "S (f (Rep x)) (f (Rep x))"  | 
|
192  | 
proof -  | 
|
193  | 
from assms(1) have "R (Rep x) (Rep x)" by (rule Quotient_rep_reflp)  | 
|
194  | 
then show ?thesis using assms(2) by (auto intro: apply_rsp')  | 
|
195  | 
qed  | 
|
196  | 
||
| 60758 | 197  | 
subsection \<open>Quotient composition\<close>  | 
| 47308 | 198  | 
|
199  | 
lemma Quotient_compose:  | 
|
200  | 
assumes 1: "Quotient R1 Abs1 Rep1 T1"  | 
|
201  | 
assumes 2: "Quotient R2 Abs2 Rep2 T2"  | 
|
202  | 
shows "Quotient (T1 OO R2 OO conversep T1) (Abs2 \<circ> Abs1) (Rep1 \<circ> Rep2) (T1 OO T2)"  | 
|
| 51994 | 203  | 
using assms unfolding Quotient_alt_def4 by fastforce  | 
| 47308 | 204  | 
|
| 
47521
 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 
kuncar 
parents: 
47501 
diff
changeset
 | 
205  | 
lemma equivp_reflp2:  | 
| 
 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 
kuncar 
parents: 
47501 
diff
changeset
 | 
206  | 
"equivp R \<Longrightarrow> reflp R"  | 
| 
 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 
kuncar 
parents: 
47501 
diff
changeset
 | 
207  | 
by (erule equivpE)  | 
| 
 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 
kuncar 
parents: 
47501 
diff
changeset
 | 
208  | 
|
| 60758 | 209  | 
subsection \<open>Respects predicate\<close>  | 
| 47544 | 210  | 
|
211  | 
definition Respects :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set"
 | 
|
212  | 
  where "Respects R = {x. R x x}"
 | 
|
213  | 
||
214  | 
lemma in_respects: "x \<in> Respects R \<longleftrightarrow> R x x"  | 
|
215  | 
unfolding Respects_def by simp  | 
|
216  | 
||
| 
47361
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
217  | 
lemma UNIV_typedef_to_Quotient:  | 
| 47308 | 218  | 
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
 | 
219  | 
and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"  | 
| 47308 | 220  | 
shows "Quotient (op =) Abs Rep T"  | 
221  | 
proof -  | 
|
222  | 
interpret type_definition Rep Abs UNIV by fact  | 
|
| 58186 | 223  | 
from Abs_inject Rep_inverse Abs_inverse T_def show ?thesis  | 
| 
47361
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
224  | 
by (fastforce intro!: QuotientI fun_eq_iff)  | 
| 47308 | 225  | 
qed  | 
226  | 
||
| 
47361
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
227  | 
lemma UNIV_typedef_to_equivp:  | 
| 47308 | 228  | 
fixes Abs :: "'a \<Rightarrow> 'b"  | 
229  | 
and Rep :: "'b \<Rightarrow> 'a"  | 
|
230  | 
assumes "type_definition Rep Abs (UNIV::'a set)"  | 
|
231  | 
shows "equivp (op=::'a\<Rightarrow>'a\<Rightarrow>bool)"  | 
|
232  | 
by (rule identity_equivp)  | 
|
233  | 
||
| 47354 | 234  | 
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
 | 
235  | 
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
 | 
236  | 
and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"  | 
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
237  | 
shows "Quotient (eq_onp (\<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
 | 
238  | 
proof -  | 
| 
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
239  | 
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
 | 
240  | 
from Rep Abs_inject Rep_inverse Abs_inverse T_def show ?thesis  | 
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
241  | 
by (auto intro!: QuotientI simp: eq_onp_def fun_eq_iff)  | 
| 
47361
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
242  | 
qed  | 
| 
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
243  | 
|
| 
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
244  | 
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
 | 
245  | 
assumes "type_definition Rep Abs S"  | 
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
246  | 
shows "part_equivp (eq_onp (\<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
 | 
247  | 
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
 | 
248  | 
interpret type_definition Rep Abs S by fact  | 
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
249  | 
show "\<exists>x. eq_onp (\<lambda>x. x \<in> S) x x" using Rep by (auto simp: eq_onp_def)  | 
| 
47361
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
250  | 
next  | 
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
251  | 
show "symp (eq_onp (\<lambda>x. x \<in> S))" by (auto intro: sympI simp: eq_onp_def)  | 
| 
47361
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
252  | 
next  | 
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
253  | 
show "transp (eq_onp (\<lambda>x. x \<in> S))" by (auto intro: transpI simp: eq_onp_def)  | 
| 
47361
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
254  | 
qed  | 
| 
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
255  | 
|
| 
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
256  | 
lemma open_typedef_to_Quotient:  | 
| 47308 | 257  | 
  assumes "type_definition Rep Abs {x. P x}"
 | 
| 47354 | 258  | 
and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"  | 
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
259  | 
shows "Quotient (eq_onp P) Abs Rep T"  | 
| 47651 | 260  | 
using typedef_to_Quotient [OF assms] by simp  | 
| 47308 | 261  | 
|
| 
47361
 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 
kuncar 
parents: 
47354 
diff
changeset
 | 
262  | 
lemma open_typedef_to_part_equivp:  | 
| 47308 | 263  | 
  assumes "type_definition Rep Abs {x. P x}"
 | 
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
264  | 
shows "part_equivp (eq_onp P)"  | 
| 47651 | 265  | 
using typedef_to_part_equivp [OF assms] by simp  | 
| 47308 | 266  | 
|
| 
60229
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
267  | 
lemma type_definition_Quotient_not_empty: "Quotient (eq_onp P) Abs Rep T \<Longrightarrow> \<exists>x. P x"  | 
| 
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
268  | 
unfolding eq_onp_def by (drule Quotient_rep_reflp) blast  | 
| 
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
269  | 
|
| 
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
270  | 
lemma type_definition_Quotient_not_empty_witness: "Quotient (eq_onp P) Abs Rep T \<Longrightarrow> P (Rep undefined)"  | 
| 
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
271  | 
unfolding eq_onp_def by (drule Quotient_rep_reflp) blast  | 
| 
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
272  | 
|
| 
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
273  | 
|
| 60758 | 274  | 
text \<open>Generating transfer rules for quotients.\<close>  | 
| 47376 | 275  | 
|
| 47537 | 276  | 
context  | 
277  | 
fixes R Abs Rep T  | 
|
278  | 
assumes 1: "Quotient R Abs Rep T"  | 
|
279  | 
begin  | 
|
| 47376 | 280  | 
|
| 47537 | 281  | 
lemma Quotient_right_unique: "right_unique T"  | 
282  | 
using 1 unfolding Quotient_alt_def right_unique_def by metis  | 
|
283  | 
||
284  | 
lemma Quotient_right_total: "right_total T"  | 
|
285  | 
using 1 unfolding Quotient_alt_def right_total_def by metis  | 
|
286  | 
||
287  | 
lemma Quotient_rel_eq_transfer: "(T ===> T ===> op =) R (op =)"  | 
|
| 55945 | 288  | 
using 1 unfolding Quotient_alt_def rel_fun_def by simp  | 
| 47376 | 289  | 
|
| 47538 | 290  | 
lemma Quotient_abs_induct:  | 
291  | 
assumes "\<And>y. R y y \<Longrightarrow> P (Abs y)" shows "P x"  | 
|
292  | 
using 1 assms unfolding Quotient_def by metis  | 
|
293  | 
||
| 47537 | 294  | 
end  | 
295  | 
||
| 60758 | 296  | 
text \<open>Generating transfer rules for total quotients.\<close>  | 
| 47376 | 297  | 
|
| 47537 | 298  | 
context  | 
299  | 
fixes R Abs Rep T  | 
|
300  | 
assumes 1: "Quotient R Abs Rep T" and 2: "reflp R"  | 
|
301  | 
begin  | 
|
| 47376 | 302  | 
|
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56517 
diff
changeset
 | 
303  | 
lemma Quotient_left_total: "left_total T"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56517 
diff
changeset
 | 
304  | 
using 1 2 unfolding Quotient_alt_def left_total_def reflp_def by auto  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56517 
diff
changeset
 | 
305  | 
|
| 47537 | 306  | 
lemma Quotient_bi_total: "bi_total T"  | 
307  | 
using 1 2 unfolding Quotient_alt_def bi_total_def reflp_def by auto  | 
|
308  | 
||
309  | 
lemma Quotient_id_abs_transfer: "(op = ===> T) (\<lambda>x. x) Abs"  | 
|
| 55945 | 310  | 
using 1 2 unfolding Quotient_alt_def reflp_def rel_fun_def by simp  | 
| 47537 | 311  | 
|
| 47575 | 312  | 
lemma Quotient_total_abs_induct: "(\<And>y. P (Abs y)) \<Longrightarrow> P x"  | 
| 63092 | 313  | 
using 1 2 unfolding Quotient_alt_def reflp_def by metis  | 
| 47575 | 314  | 
|
| 
47889
 
29212a4bb866
lifting package produces abs_eq_iff rules for total quotients
 
huffman 
parents: 
47777 
diff
changeset
 | 
315  | 
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
 | 
316  | 
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
 | 
317  | 
|
| 47537 | 318  | 
end  | 
| 47376 | 319  | 
|
| 61799 | 320  | 
text \<open>Generating transfer rules for a type defined with \<open>typedef\<close>.\<close>  | 
| 
47368
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
321  | 
|
| 
47534
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
322  | 
context  | 
| 
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
323  | 
fixes Rep Abs A T  | 
| 
47368
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
324  | 
assumes type: "type_definition Rep Abs A"  | 
| 
47534
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
325  | 
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
 | 
326  | 
begin  | 
| 
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
327  | 
|
| 51994 | 328  | 
lemma typedef_left_unique: "left_unique T"  | 
329  | 
unfolding left_unique_def T_def  | 
|
330  | 
by (simp add: type_definition.Rep_inject [OF type])  | 
|
331  | 
||
| 
47534
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
332  | 
lemma typedef_bi_unique: "bi_unique T"  | 
| 
47368
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
333  | 
unfolding bi_unique_def T_def  | 
| 
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
334  | 
by (simp add: type_definition.Rep_inject [OF type])  | 
| 
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
335  | 
|
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
336  | 
(* 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
 | 
337  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
338  | 
lemma typedef_right_unique: "right_unique T"  | 
| 58186 | 339  | 
using T_def type Quotient_right_unique typedef_to_Quotient  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
340  | 
by blast  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
341  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
342  | 
lemma typedef_right_total: "right_total T"  | 
| 58186 | 343  | 
using T_def type Quotient_right_total typedef_to_Quotient  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
344  | 
by blast  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
345  | 
|
| 
47535
 
0f94b02fda1c
lifting_setup generates transfer rule for rep of typedefs
 
huffman 
parents: 
47534 
diff
changeset
 | 
346  | 
lemma typedef_rep_transfer: "(T ===> op =) (\<lambda>x. x) Rep"  | 
| 55945 | 347  | 
unfolding rel_fun_def T_def by simp  | 
| 
47535
 
0f94b02fda1c
lifting_setup generates transfer rule for rep of typedefs
 
huffman 
parents: 
47534 
diff
changeset
 | 
348  | 
|
| 
47534
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
349  | 
end  | 
| 
 
06cc372a80ed
use context block to organize typedef lifting theorems
 
huffman 
parents: 
47521 
diff
changeset
 | 
350  | 
|
| 60758 | 351  | 
text \<open>Generating the correspondence rule for a constant defined with  | 
| 61799 | 352  | 
\<open>lift_definition\<close>.\<close>  | 
| 
47368
 
4c522dff1f4d
add lemmas for generating transfer rules for typedefs
 
huffman 
parents: 
47354 
diff
changeset
 | 
353  | 
|
| 47351 | 354  | 
lemma Quotient_to_transfer:  | 
355  | 
assumes "Quotient R Abs Rep T" and "R c c" and "c' \<equiv> Abs c"  | 
|
356  | 
shows "T c c'"  | 
|
357  | 
using assms by (auto dest: Quotient_cr_rel)  | 
|
358  | 
||
| 60758 | 359  | 
text \<open>Proving reflexivity\<close>  | 
| 
47982
 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 
kuncar 
parents: 
47937 
diff
changeset
 | 
360  | 
|
| 
 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 
kuncar 
parents: 
47937 
diff
changeset
 | 
361  | 
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
 | 
362  | 
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
 | 
363  | 
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
 | 
364  | 
shows "left_total T"  | 
| 
 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 
kuncar 
parents: 
47937 
diff
changeset
 | 
365  | 
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
 | 
366  | 
|
| 
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
 | 
367  | 
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
 | 
368  | 
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
 | 
369  | 
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
 | 
370  | 
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
 | 
371  | 
using assms unfolding left_total_def by fast  | 
| 51994 | 372  | 
|
| 
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
 | 
373  | 
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
 | 
374  | 
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
 | 
375  | 
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
 | 
376  | 
shows "(T OO R OO T\<inverse>\<inverse>) \<le> op="  | 
| 55604 | 377  | 
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
 | 
378  | 
|
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
379  | 
lemma eq_onp_le_eq:  | 
| 
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
380  | 
"eq_onp P \<le> op=" unfolding eq_onp_def by blast  | 
| 
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
 | 
381  | 
|
| 
 
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
 | 
382  | 
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
 | 
383  | 
"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
 | 
384  | 
|
| 60758 | 385  | 
text \<open>Proving a parametrized correspondence relation\<close>  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
386  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
387  | 
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
 | 
388  | 
"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
 | 
389  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
390  | 
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
 | 
391  | 
"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
 | 
392  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
393  | 
lemma pos_OO_eq:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
394  | 
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
 | 
395  | 
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
 | 
396  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
397  | 
lemma pos_eq_OO:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
398  | 
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
 | 
399  | 
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
 | 
400  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
401  | 
lemma neg_OO_eq:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
402  | 
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
 | 
403  | 
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
 | 
404  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
405  | 
lemma neg_eq_OO:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
406  | 
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
 | 
407  | 
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
 | 
408  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
409  | 
lemma POS_trans:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
410  | 
assumes "POS A B"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
411  | 
assumes "POS B C"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
412  | 
shows "POS A C"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
413  | 
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
 | 
414  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
415  | 
lemma NEG_trans:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
416  | 
assumes "NEG A B"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
417  | 
assumes "NEG B C"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
418  | 
shows "NEG A C"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
419  | 
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
 | 
420  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
421  | 
lemma POS_NEG:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
422  | 
"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
 | 
423  | 
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
 | 
424  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
425  | 
lemma NEG_POS:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
426  | 
"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
 | 
427  | 
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
 | 
428  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
429  | 
lemma POS_pcr_rule:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
430  | 
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
 | 
431  | 
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
 | 
432  | 
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
 | 
433  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
434  | 
lemma NEG_pcr_rule:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
435  | 
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
 | 
436  | 
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
 | 
437  | 
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
 | 
438  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
439  | 
lemma POS_apply:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
440  | 
assumes "POS R R'"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
441  | 
assumes "R f g"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
442  | 
shows "R' f g"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
443  | 
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
 | 
444  | 
|
| 60758 | 445  | 
text \<open>Proving a parametrized correspondence relation\<close>  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
446  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
447  | 
lemma fun_mono:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
448  | 
assumes "A \<ge> C"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
449  | 
assumes "B \<le> D"  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
450  | 
shows "(A ===> B) \<le> (C ===> D)"  | 
| 55945 | 451  | 
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
 | 
452  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
453  | 
lemma pos_fun_distr: "((R ===> S) OO (R' ===> S')) \<le> ((R OO R') ===> (S OO S'))"  | 
| 55945 | 454  | 
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
 | 
455  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
456  | 
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
 | 
457  | 
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
 | 
458  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
459  | 
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
 | 
460  | 
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
 | 
461  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
462  | 
lemma neg_fun_distr1:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
463  | 
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
 | 
464  | 
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
 | 
465  | 
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
 | 
466  | 
using functional_relation[OF 2] functional_converse_relation[OF 1]  | 
| 55945 | 467  | 
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
 | 
468  | 
apply clarify  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
469  | 
apply (subst all_comm)  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
470  | 
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
 | 
471  | 
apply (intro choice)  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
472  | 
by metis  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
473  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
474  | 
lemma neg_fun_distr2:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
475  | 
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
 | 
476  | 
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
 | 
477  | 
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
 | 
478  | 
using functional_converse_relation[OF 2] functional_relation[OF 1]  | 
| 55945 | 479  | 
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
 | 
480  | 
apply clarify  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
481  | 
apply (subst all_comm)  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
482  | 
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
 | 
483  | 
apply (intro choice)  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
484  | 
by metis  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
485  | 
|
| 60758 | 486  | 
subsection \<open>Domains\<close>  | 
| 
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
 | 
487  | 
|
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
488  | 
lemma composed_equiv_rel_eq_onp:  | 
| 
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
 | 
489  | 
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
 | 
490  | 
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
 | 
491  | 
assumes "Domainp R = P''"  | 
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
492  | 
shows "(R OO eq_onp P' OO R\<inverse>\<inverse>) = eq_onp (inf P'' P)"  | 
| 
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
493  | 
using assms unfolding OO_def conversep_iff Domainp_iff[abs_def] left_unique_def rel_fun_def eq_onp_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
 | 
494  | 
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
 | 
495  | 
|
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
496  | 
lemma composed_equiv_rel_eq_eq_onp:  | 
| 
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
 | 
497  | 
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
 | 
498  | 
assumes "Domainp R = P"  | 
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
499  | 
shows "(R OO op= OO R\<inverse>\<inverse>) = eq_onp P"  | 
| 
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
500  | 
using assms unfolding OO_def conversep_iff Domainp_iff[abs_def] left_unique_def eq_onp_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
 | 
501  | 
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
 | 
502  | 
|
| 
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
 | 
503  | 
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
 | 
504  | 
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
 | 
505  | 
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
 | 
506  | 
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
 | 
507  | 
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
 | 
508  | 
using assms  | 
| 58186 | 509  | 
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
 | 
510  | 
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
 | 
511  | 
|
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51374 
diff
changeset
 | 
512  | 
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
 | 
513  | 
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
 | 
514  | 
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
 | 
515  | 
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
 | 
516  | 
shows "Domainp (A OO B) = (inf P1 P2')"  | 
| 55945 | 517  | 
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
 | 
518  | 
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
 | 
519  | 
|
| 53151 | 520  | 
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
 | 
521  | 
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
 | 
522  | 
|
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51374 
diff
changeset
 | 
523  | 
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
 | 
524  | 
assumes "Domainp B = P"  | 
| 53151 | 525  | 
shows "Domainp (A OO B) = (\<lambda>x. \<exists>y. A x y \<and> P y)"  | 
526  | 
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
 | 
527  | 
|
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51374 
diff
changeset
 | 
528  | 
lemma pcr_Domainp_total:  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56517 
diff
changeset
 | 
529  | 
assumes "left_total B"  | 
| 
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
 | 
530  | 
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
 | 
531  | 
shows "Domainp (A OO B) = P"  | 
| 58186 | 532  | 
using assms unfolding left_total_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
 | 
533  | 
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
 | 
534  | 
|
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51374 
diff
changeset
 | 
535  | 
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
 | 
536  | 
assumes "Quotient R Abs Rep T"  | 
| 58186 | 537  | 
shows "Domainp T = (\<lambda>x. R x x)"  | 
| 
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
 | 
538  | 
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
 | 
539  | 
|
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
540  | 
lemma eq_onp_to_Domainp:  | 
| 
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
541  | 
assumes "Quotient (eq_onp P) Abs Rep T"  | 
| 
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
 | 
542  | 
shows "Domainp T = P"  | 
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
543  | 
by (simp add: eq_onp_def Domainp_iff[abs_def] Quotient_cr_rel[OF assms])  | 
| 
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
 | 
544  | 
|
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52307 
diff
changeset
 | 
545  | 
end  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52307 
diff
changeset
 | 
546  | 
|
| 
60229
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
547  | 
(* needed for lifting_def_code_dt.ML (moved from Lifting_Set) *)  | 
| 
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
548  | 
lemma right_total_UNIV_transfer:  | 
| 
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
549  | 
assumes "right_total A"  | 
| 
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
550  | 
shows "(rel_set A) (Collect (Domainp A)) UNIV"  | 
| 
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
551  | 
using assms unfolding right_total_def rel_set_def Domainp_iff by blast  | 
| 
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
552  | 
|
| 60758 | 553  | 
subsection \<open>ML setup\<close>  | 
| 47308 | 554  | 
|
| 48891 | 555  | 
ML_file "Tools/Lifting/lifting_util.ML"  | 
| 47308 | 556  | 
|
| 57961 | 557  | 
named_theorems relator_eq_onp  | 
558  | 
"theorems that a relator of an eq_onp is an eq_onp of the corresponding predicate"  | 
|
| 48891 | 559  | 
ML_file "Tools/Lifting/lifting_info.ML"  | 
| 47308 | 560  | 
|
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51112 
diff
changeset
 | 
561  | 
(* setup for the function type *)  | 
| 
47777
 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 
kuncar 
parents: 
47698 
diff
changeset
 | 
562  | 
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
 | 
563  | 
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
 | 
564  | 
lemmas [relator_distr] = pos_fun_distr neg_fun_distr1 neg_fun_distr2  | 
| 47308 | 565  | 
|
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56519 
diff
changeset
 | 
566  | 
ML_file "Tools/Lifting/lifting_bnf.ML"  | 
| 48891 | 567  | 
ML_file "Tools/Lifting/lifting_term.ML"  | 
568  | 
ML_file "Tools/Lifting/lifting_def.ML"  | 
|
569  | 
ML_file "Tools/Lifting/lifting_setup.ML"  | 
|
| 
60229
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59726 
diff
changeset
 | 
570  | 
ML_file "Tools/Lifting/lifting_def_code_dt.ML"  | 
| 58186 | 571  | 
|
| 61630 | 572  | 
lemma pred_prod_beta: "pred_prod P Q xy \<longleftrightarrow> P (fst xy) \<and> Q (snd xy)"  | 
573  | 
by(cases xy) simp  | 
|
574  | 
||
575  | 
lemma pred_prod_split: "P (pred_prod Q R xy) \<longleftrightarrow> (\<forall>x y. xy = (x, y) \<longrightarrow> P (Q x \<and> R y))"  | 
|
576  | 
by(cases xy) simp  | 
|
577  | 
||
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
578  | 
hide_const (open) POS NEG  | 
| 47308 | 579  | 
|
580  | 
end  |