| author | wenzelm | 
| Sat, 22 Mar 2014 18:16:54 +0100 | |
| changeset 56253 | 83b3c110f22d | 
| parent 56085 | 3d11892ea537 | 
| child 56518 | beb3b6851665 | 
| permissions | -rw-r--r-- | 
| 47325 | 1  | 
(* Title: HOL/Transfer.thy  | 
2  | 
Author: Brian Huffman, TU Muenchen  | 
|
| 
51956
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
3  | 
Author: Ondrej Kuncar, TU Muenchen  | 
| 47325 | 4  | 
*)  | 
5  | 
||
6  | 
header {* Generic theorem transfer using relations *}
 | 
|
7  | 
||
8  | 
theory Transfer  | 
|
| 55811 | 9  | 
imports Hilbert_Choice Basic_BNFs Metis  | 
| 47325 | 10  | 
begin  | 
11  | 
||
12  | 
subsection {* Relator for function space *}
 | 
|
13  | 
||
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
14  | 
locale lifting_syntax  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
15  | 
begin  | 
| 55945 | 16  | 
notation rel_fun (infixr "===>" 55)  | 
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
17  | 
notation map_fun (infixr "--->" 55)  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
18  | 
end  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
19  | 
|
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
20  | 
context  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
21  | 
begin  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
22  | 
interpretation lifting_syntax .  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
23  | 
|
| 55945 | 24  | 
lemma rel_funD2:  | 
25  | 
assumes "rel_fun A B f g" and "A x 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: 
47924 
diff
changeset
 | 
26  | 
shows "B (f x) (g x)"  | 
| 55945 | 27  | 
using assms by (rule rel_funD)  | 
| 
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: 
47924 
diff
changeset
 | 
28  | 
|
| 55945 | 29  | 
lemma rel_funE:  | 
30  | 
assumes "rel_fun A B f g" and "A x y"  | 
|
| 47325 | 31  | 
obtains "B (f x) (g y)"  | 
| 55945 | 32  | 
using assms by (simp add: rel_fun_def)  | 
| 47325 | 33  | 
|
| 55945 | 34  | 
lemmas rel_fun_eq = fun.rel_eq  | 
| 47325 | 35  | 
|
| 55945 | 36  | 
lemma rel_fun_eq_rel:  | 
37  | 
shows "rel_fun (op =) R = (\<lambda>f g. \<forall>x. R (f x) (g x))"  | 
|
38  | 
by (simp add: rel_fun_def)  | 
|
| 47325 | 39  | 
|
40  | 
||
41  | 
subsection {* Transfer method *}
 | 
|
42  | 
||
| 
47789
 
71a526ee569a
implement transfer tactic with more scalable forward proof methods
 
huffman 
parents: 
47684 
diff
changeset
 | 
43  | 
text {* Explicit tag for relation membership allows for
 | 
| 
 
71a526ee569a
implement transfer tactic with more scalable forward proof methods
 
huffman 
parents: 
47684 
diff
changeset
 | 
44  | 
backward proof methods. *}  | 
| 47325 | 45  | 
|
46  | 
definition Rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
 | 
|
47  | 
where "Rel r \<equiv> r"  | 
|
48  | 
||
| 
49975
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
48891 
diff
changeset
 | 
49  | 
text {* Handling of equality relations *}
 | 
| 
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
48891 
diff
changeset
 | 
50  | 
|
| 
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
48891 
diff
changeset
 | 
51  | 
definition is_equality :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
48891 
diff
changeset
 | 
52  | 
where "is_equality R \<longleftrightarrow> R = (op =)"  | 
| 
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
48891 
diff
changeset
 | 
53  | 
|
| 
51437
 
8739f8abbecb
fixing transfer tactic - unfold fully identity relation by using relator_eq
 
kuncar 
parents: 
51112 
diff
changeset
 | 
54  | 
lemma is_equality_eq: "is_equality (op =)"  | 
| 
 
8739f8abbecb
fixing transfer tactic - unfold fully identity relation by using relator_eq
 
kuncar 
parents: 
51112 
diff
changeset
 | 
55  | 
unfolding is_equality_def by simp  | 
| 
 
8739f8abbecb
fixing transfer tactic - unfold fully identity relation by using relator_eq
 
kuncar 
parents: 
51112 
diff
changeset
 | 
56  | 
|
| 
52354
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
57  | 
text {* Reverse implication for monotonicity rules *}
 | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
58  | 
|
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
59  | 
definition rev_implies where  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
60  | 
"rev_implies x y \<longleftrightarrow> (y \<longrightarrow> x)"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
61  | 
|
| 47325 | 62  | 
text {* Handling of meta-logic connectives *}
 | 
63  | 
||
64  | 
definition transfer_forall where  | 
|
65  | 
"transfer_forall \<equiv> All"  | 
|
66  | 
||
67  | 
definition transfer_implies where  | 
|
68  | 
"transfer_implies \<equiv> op \<longrightarrow>"  | 
|
69  | 
||
| 
47355
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
70  | 
definition transfer_bforall :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
71  | 
where "transfer_bforall \<equiv> (\<lambda>P Q. \<forall>x. P x \<longrightarrow> Q x)"  | 
| 
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
72  | 
|
| 47325 | 73  | 
lemma transfer_forall_eq: "(\<And>x. P x) \<equiv> Trueprop (transfer_forall (\<lambda>x. P x))"  | 
74  | 
unfolding atomize_all transfer_forall_def ..  | 
|
75  | 
||
76  | 
lemma transfer_implies_eq: "(A \<Longrightarrow> B) \<equiv> Trueprop (transfer_implies A B)"  | 
|
77  | 
unfolding atomize_imp transfer_implies_def ..  | 
|
78  | 
||
| 
47355
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
79  | 
lemma transfer_bforall_unfold:  | 
| 
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
80  | 
"Trueprop (transfer_bforall P (\<lambda>x. Q x)) \<equiv> (\<And>x. P x \<Longrightarrow> Q x)"  | 
| 
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
81  | 
unfolding transfer_bforall_def atomize_imp atomize_all ..  | 
| 
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
82  | 
|
| 
47658
 
7631f6f7873d
enable variant of transfer method that proves an implication instead of an equivalence
 
huffman 
parents: 
47637 
diff
changeset
 | 
83  | 
lemma transfer_start: "\<lbrakk>P; Rel (op =) P Q\<rbrakk> \<Longrightarrow> Q"  | 
| 47325 | 84  | 
unfolding Rel_def by simp  | 
85  | 
||
| 
47658
 
7631f6f7873d
enable variant of transfer method that proves an implication instead of an equivalence
 
huffman 
parents: 
47637 
diff
changeset
 | 
86  | 
lemma transfer_start': "\<lbrakk>P; Rel (op \<longrightarrow>) P Q\<rbrakk> \<Longrightarrow> Q"  | 
| 47325 | 87  | 
unfolding Rel_def by simp  | 
88  | 
||
| 
47635
 
ebb79474262c
rename 'correspondence' method to 'transfer_prover'
 
huffman 
parents: 
47627 
diff
changeset
 | 
89  | 
lemma transfer_prover_start: "\<lbrakk>x = x'; Rel R x' y\<rbrakk> \<Longrightarrow> Rel R x y"  | 
| 
47618
 
1568dadd598a
make correspondence tactic more robust by replacing lhs with schematic variable before applying intro rules
 
huffman 
parents: 
47612 
diff
changeset
 | 
90  | 
by simp  | 
| 
 
1568dadd598a
make correspondence tactic more robust by replacing lhs with schematic variable before applying intro rules
 
huffman 
parents: 
47612 
diff
changeset
 | 
91  | 
|
| 
52358
 
f4c4bcb0d564
implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
 
huffman 
parents: 
52354 
diff
changeset
 | 
92  | 
lemma untransfer_start: "\<lbrakk>Q; Rel (op =) P Q\<rbrakk> \<Longrightarrow> P"  | 
| 
 
f4c4bcb0d564
implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
 
huffman 
parents: 
52354 
diff
changeset
 | 
93  | 
unfolding Rel_def by simp  | 
| 
 
f4c4bcb0d564
implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
 
huffman 
parents: 
52354 
diff
changeset
 | 
94  | 
|
| 47325 | 95  | 
lemma Rel_eq_refl: "Rel (op =) x x"  | 
96  | 
unfolding Rel_def ..  | 
|
97  | 
||
| 
47789
 
71a526ee569a
implement transfer tactic with more scalable forward proof methods
 
huffman 
parents: 
47684 
diff
changeset
 | 
98  | 
lemma Rel_app:  | 
| 
47523
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
99  | 
assumes "Rel (A ===> B) f g" and "Rel A x y"  | 
| 
47789
 
71a526ee569a
implement transfer tactic with more scalable forward proof methods
 
huffman 
parents: 
47684 
diff
changeset
 | 
100  | 
shows "Rel B (f x) (g y)"  | 
| 55945 | 101  | 
using assms unfolding Rel_def rel_fun_def by fast  | 
| 
47523
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
102  | 
|
| 
47789
 
71a526ee569a
implement transfer tactic with more scalable forward proof methods
 
huffman 
parents: 
47684 
diff
changeset
 | 
103  | 
lemma Rel_abs:  | 
| 
47523
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
104  | 
assumes "\<And>x y. Rel A x y \<Longrightarrow> Rel B (f x) (g y)"  | 
| 
47789
 
71a526ee569a
implement transfer tactic with more scalable forward proof methods
 
huffman 
parents: 
47684 
diff
changeset
 | 
105  | 
shows "Rel (A ===> B) (\<lambda>x. f x) (\<lambda>y. g y)"  | 
| 55945 | 106  | 
using assms unfolding Rel_def rel_fun_def by fast  | 
| 
47523
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
107  | 
|
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
108  | 
end  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
109  | 
|
| 48891 | 110  | 
ML_file "Tools/transfer.ML"  | 
| 47325 | 111  | 
setup Transfer.setup  | 
112  | 
||
| 
49975
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
48891 
diff
changeset
 | 
113  | 
declare refl [transfer_rule]  | 
| 
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
48891 
diff
changeset
 | 
114  | 
|
| 55945 | 115  | 
declare rel_fun_eq [relator_eq]  | 
| 47503 | 116  | 
|
| 
47789
 
71a526ee569a
implement transfer tactic with more scalable forward proof methods
 
huffman 
parents: 
47684 
diff
changeset
 | 
117  | 
hide_const (open) Rel  | 
| 47325 | 118  | 
|
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
119  | 
context  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
120  | 
begin  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
121  | 
interpretation lifting_syntax .  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
122  | 
|
| 
51956
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
123  | 
text {* Handling of domains *}
 | 
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
124  | 
|
| 
55760
 
aaaccc8e015f
transfer domain rule for special case of functions - was missing
 
kuncar 
parents: 
55415 
diff
changeset
 | 
125  | 
lemma Domainp_iff: "Domainp T x \<longleftrightarrow> (\<exists>y. T x y)"  | 
| 
 
aaaccc8e015f
transfer domain rule for special case of functions - was missing
 
kuncar 
parents: 
55415 
diff
changeset
 | 
126  | 
by auto  | 
| 
 
aaaccc8e015f
transfer domain rule for special case of functions - was missing
 
kuncar 
parents: 
55415 
diff
changeset
 | 
127  | 
|
| 
51956
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
128  | 
lemma Domaimp_refl[transfer_domain_rule]:  | 
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
129  | 
"Domainp T = Domainp T" ..  | 
| 47325 | 130  | 
|
| 
55760
 
aaaccc8e015f
transfer domain rule for special case of functions - was missing
 
kuncar 
parents: 
55415 
diff
changeset
 | 
131  | 
lemma Domainp_prod_fun_eq[transfer_domain_rule]:  | 
| 
 
aaaccc8e015f
transfer domain rule for special case of functions - was missing
 
kuncar 
parents: 
55415 
diff
changeset
 | 
132  | 
assumes "Domainp T = P"  | 
| 
 
aaaccc8e015f
transfer domain rule for special case of functions - was missing
 
kuncar 
parents: 
55415 
diff
changeset
 | 
133  | 
shows "Domainp (op= ===> T) = (\<lambda>f. \<forall>x. P (f x))"  | 
| 55945 | 134  | 
by (auto intro: choice simp: assms[symmetric] Domainp_iff rel_fun_def fun_eq_iff)  | 
| 
55760
 
aaaccc8e015f
transfer domain rule for special case of functions - was missing
 
kuncar 
parents: 
55415 
diff
changeset
 | 
135  | 
|
| 47325 | 136  | 
subsection {* Predicates on relations, i.e. ``class constraints'' *}
 | 
137  | 
||
138  | 
definition right_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
|
139  | 
where "right_total R \<longleftrightarrow> (\<forall>y. \<exists>x. R x y)"  | 
|
140  | 
||
141  | 
definition right_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
|
142  | 
where "right_unique R \<longleftrightarrow> (\<forall>x y z. R x y \<longrightarrow> R x z \<longrightarrow> y = z)"  | 
|
143  | 
||
144  | 
definition bi_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
|
145  | 
where "bi_total R \<longleftrightarrow> (\<forall>x. \<exists>y. R x y) \<and> (\<forall>y. \<exists>x. R x y)"  | 
|
146  | 
||
147  | 
definition bi_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
|
148  | 
where "bi_unique R \<longleftrightarrow>  | 
|
149  | 
(\<forall>x y z. R x y \<longrightarrow> R x z \<longrightarrow> y = z) \<and>  | 
|
150  | 
(\<forall>x y z. R x z \<longrightarrow> R y z \<longrightarrow> x = y)"  | 
|
151  | 
||
| 53927 | 152  | 
lemma bi_uniqueDr: "\<lbrakk> bi_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z"  | 
153  | 
by(simp add: bi_unique_def)  | 
|
154  | 
||
155  | 
lemma bi_uniqueDl: "\<lbrakk> bi_unique A; A x y; A z y \<rbrakk> \<Longrightarrow> x = z"  | 
|
156  | 
by(simp add: bi_unique_def)  | 
|
157  | 
||
158  | 
lemma right_uniqueI: "(\<And>x y z. \<lbrakk> A x y; A x z \<rbrakk> \<Longrightarrow> y = z) \<Longrightarrow> right_unique A"  | 
|
| 56085 | 159  | 
unfolding right_unique_def by fast  | 
| 53927 | 160  | 
|
161  | 
lemma right_uniqueD: "\<lbrakk> right_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z"  | 
|
| 56085 | 162  | 
unfolding right_unique_def by fast  | 
| 53927 | 163  | 
|
| 47325 | 164  | 
lemma right_total_alt_def:  | 
165  | 
"right_total R \<longleftrightarrow> ((R ===> op \<longrightarrow>) ===> op \<longrightarrow>) All All"  | 
|
| 55945 | 166  | 
unfolding right_total_def rel_fun_def  | 
| 47325 | 167  | 
apply (rule iffI, fast)  | 
168  | 
apply (rule allI)  | 
|
169  | 
apply (drule_tac x="\<lambda>x. True" in spec)  | 
|
170  | 
apply (drule_tac x="\<lambda>y. \<exists>x. R x y" in spec)  | 
|
171  | 
apply fast  | 
|
172  | 
done  | 
|
173  | 
||
174  | 
lemma right_unique_alt_def:  | 
|
175  | 
"right_unique R \<longleftrightarrow> (R ===> R ===> op \<longrightarrow>) (op =) (op =)"  | 
|
| 55945 | 176  | 
unfolding right_unique_def rel_fun_def by auto  | 
| 47325 | 177  | 
|
178  | 
lemma bi_total_alt_def:  | 
|
179  | 
"bi_total R \<longleftrightarrow> ((R ===> op =) ===> op =) All All"  | 
|
| 55945 | 180  | 
unfolding bi_total_def rel_fun_def  | 
| 47325 | 181  | 
apply (rule iffI, fast)  | 
182  | 
apply safe  | 
|
183  | 
apply (drule_tac x="\<lambda>x. \<exists>y. R x y" in spec)  | 
|
184  | 
apply (drule_tac x="\<lambda>y. True" in spec)  | 
|
185  | 
apply fast  | 
|
186  | 
apply (drule_tac x="\<lambda>x. True" in spec)  | 
|
187  | 
apply (drule_tac x="\<lambda>y. \<exists>x. R x y" in spec)  | 
|
188  | 
apply fast  | 
|
189  | 
done  | 
|
190  | 
||
191  | 
lemma bi_unique_alt_def:  | 
|
192  | 
"bi_unique R \<longleftrightarrow> (R ===> R ===> op =) (op =) (op =)"  | 
|
| 55945 | 193  | 
unfolding bi_unique_def rel_fun_def by auto  | 
| 47325 | 194  | 
|
| 53944 | 195  | 
lemma bi_unique_conversep [simp]: "bi_unique R\<inverse>\<inverse> = bi_unique R"  | 
196  | 
by(auto simp add: bi_unique_def)  | 
|
197  | 
||
198  | 
lemma bi_total_conversep [simp]: "bi_total R\<inverse>\<inverse> = bi_total R"  | 
|
199  | 
by(auto simp add: bi_total_def)  | 
|
200  | 
||
| 47660 | 201  | 
text {* Properties are preserved by relation composition. *}
 | 
202  | 
||
203  | 
lemma OO_def: "R OO S = (\<lambda>x z. \<exists>y. R x y \<and> S y z)"  | 
|
204  | 
by auto  | 
|
205  | 
||
206  | 
lemma bi_total_OO: "\<lbrakk>bi_total A; bi_total B\<rbrakk> \<Longrightarrow> bi_total (A OO B)"  | 
|
| 56085 | 207  | 
unfolding bi_total_def OO_def by fast  | 
| 47660 | 208  | 
|
209  | 
lemma bi_unique_OO: "\<lbrakk>bi_unique A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A OO B)"  | 
|
| 56085 | 210  | 
unfolding bi_unique_def OO_def by blast  | 
| 47660 | 211  | 
|
212  | 
lemma right_total_OO:  | 
|
213  | 
"\<lbrakk>right_total A; right_total B\<rbrakk> \<Longrightarrow> right_total (A OO B)"  | 
|
| 56085 | 214  | 
unfolding right_total_def OO_def by fast  | 
| 47660 | 215  | 
|
216  | 
lemma right_unique_OO:  | 
|
217  | 
"\<lbrakk>right_unique A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A OO B)"  | 
|
| 56085 | 218  | 
unfolding right_unique_def OO_def by fast  | 
| 47660 | 219  | 
|
| 47325 | 220  | 
|
221  | 
subsection {* Properties of relators *}
 | 
|
222  | 
||
223  | 
lemma right_total_eq [transfer_rule]: "right_total (op =)"  | 
|
224  | 
unfolding right_total_def by simp  | 
|
225  | 
||
226  | 
lemma right_unique_eq [transfer_rule]: "right_unique (op =)"  | 
|
227  | 
unfolding right_unique_def by simp  | 
|
228  | 
||
229  | 
lemma bi_total_eq [transfer_rule]: "bi_total (op =)"  | 
|
230  | 
unfolding bi_total_def by simp  | 
|
231  | 
||
232  | 
lemma bi_unique_eq [transfer_rule]: "bi_unique (op =)"  | 
|
233  | 
unfolding bi_unique_def by simp  | 
|
234  | 
||
235  | 
lemma right_total_fun [transfer_rule]:  | 
|
236  | 
"\<lbrakk>right_unique A; right_total B\<rbrakk> \<Longrightarrow> right_total (A ===> B)"  | 
|
| 55945 | 237  | 
unfolding right_total_def rel_fun_def  | 
| 47325 | 238  | 
apply (rule allI, rename_tac g)  | 
239  | 
apply (rule_tac x="\<lambda>x. SOME z. B z (g (THE y. A x y))" in exI)  | 
|
240  | 
apply clarify  | 
|
241  | 
apply (subgoal_tac "(THE y. A x y) = y", simp)  | 
|
242  | 
apply (rule someI_ex)  | 
|
243  | 
apply (simp)  | 
|
244  | 
apply (rule the_equality)  | 
|
245  | 
apply assumption  | 
|
246  | 
apply (simp add: right_unique_def)  | 
|
247  | 
done  | 
|
248  | 
||
249  | 
lemma right_unique_fun [transfer_rule]:  | 
|
250  | 
"\<lbrakk>right_total A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A ===> B)"  | 
|
| 55945 | 251  | 
unfolding right_total_def right_unique_def rel_fun_def  | 
| 47325 | 252  | 
by (clarify, rule ext, fast)  | 
253  | 
||
254  | 
lemma bi_total_fun [transfer_rule]:  | 
|
255  | 
"\<lbrakk>bi_unique A; bi_total B\<rbrakk> \<Longrightarrow> bi_total (A ===> B)"  | 
|
| 55945 | 256  | 
unfolding bi_total_def rel_fun_def  | 
| 47325 | 257  | 
apply safe  | 
258  | 
apply (rename_tac f)  | 
|
259  | 
apply (rule_tac x="\<lambda>y. SOME z. B (f (THE x. A x y)) z" in exI)  | 
|
260  | 
apply clarify  | 
|
261  | 
apply (subgoal_tac "(THE x. A x y) = x", simp)  | 
|
262  | 
apply (rule someI_ex)  | 
|
263  | 
apply (simp)  | 
|
264  | 
apply (rule the_equality)  | 
|
265  | 
apply assumption  | 
|
266  | 
apply (simp add: bi_unique_def)  | 
|
267  | 
apply (rename_tac g)  | 
|
268  | 
apply (rule_tac x="\<lambda>x. SOME z. B z (g (THE y. A x y))" in exI)  | 
|
269  | 
apply clarify  | 
|
270  | 
apply (subgoal_tac "(THE y. A x y) = y", simp)  | 
|
271  | 
apply (rule someI_ex)  | 
|
272  | 
apply (simp)  | 
|
273  | 
apply (rule the_equality)  | 
|
274  | 
apply assumption  | 
|
275  | 
apply (simp add: bi_unique_def)  | 
|
276  | 
done  | 
|
277  | 
||
278  | 
lemma bi_unique_fun [transfer_rule]:  | 
|
279  | 
"\<lbrakk>bi_total A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A ===> B)"  | 
|
| 55945 | 280  | 
unfolding bi_total_def bi_unique_def rel_fun_def fun_eq_iff  | 
| 56085 | 281  | 
by fast+  | 
| 47325 | 282  | 
|
283  | 
||
| 
47635
 
ebb79474262c
rename 'correspondence' method to 'transfer_prover'
 
huffman 
parents: 
47627 
diff
changeset
 | 
284  | 
subsection {* Transfer rules *}
 | 
| 47325 | 285  | 
|
| 53952 | 286  | 
lemma Domainp_forall_transfer [transfer_rule]:  | 
287  | 
assumes "right_total A"  | 
|
288  | 
shows "((A ===> op =) ===> op =)  | 
|
289  | 
(transfer_bforall (Domainp A)) transfer_forall"  | 
|
290  | 
using assms unfolding right_total_def  | 
|
| 55945 | 291  | 
unfolding transfer_forall_def transfer_bforall_def rel_fun_def Domainp_iff  | 
| 56085 | 292  | 
by fast  | 
| 53952 | 293  | 
|
| 47684 | 294  | 
text {* Transfer rules using implication instead of equality on booleans. *}
 | 
295  | 
||
| 
52354
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
296  | 
lemma transfer_forall_transfer [transfer_rule]:  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
297  | 
"bi_total A \<Longrightarrow> ((A ===> op =) ===> op =) transfer_forall transfer_forall"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
298  | 
"right_total A \<Longrightarrow> ((A ===> op =) ===> implies) transfer_forall transfer_forall"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
299  | 
"right_total A \<Longrightarrow> ((A ===> implies) ===> implies) transfer_forall transfer_forall"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
300  | 
"bi_total A \<Longrightarrow> ((A ===> op =) ===> rev_implies) transfer_forall transfer_forall"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
301  | 
"bi_total A \<Longrightarrow> ((A ===> rev_implies) ===> rev_implies) transfer_forall transfer_forall"  | 
| 55945 | 302  | 
unfolding transfer_forall_def rev_implies_def rel_fun_def right_total_def bi_total_def  | 
| 56085 | 303  | 
by fast+  | 
| 
52354
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
304  | 
|
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
305  | 
lemma transfer_implies_transfer [transfer_rule]:  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
306  | 
"(op = ===> op = ===> op = ) transfer_implies transfer_implies"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
307  | 
"(rev_implies ===> implies ===> implies ) transfer_implies transfer_implies"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
308  | 
"(rev_implies ===> op = ===> implies ) transfer_implies transfer_implies"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
309  | 
"(op = ===> implies ===> implies ) transfer_implies transfer_implies"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
310  | 
"(op = ===> op = ===> implies ) transfer_implies transfer_implies"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
311  | 
"(implies ===> rev_implies ===> rev_implies) transfer_implies transfer_implies"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
312  | 
"(implies ===> op = ===> rev_implies) transfer_implies transfer_implies"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
313  | 
"(op = ===> rev_implies ===> rev_implies) transfer_implies transfer_implies"  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
314  | 
"(op = ===> op = ===> rev_implies) transfer_implies transfer_implies"  | 
| 55945 | 315  | 
unfolding transfer_implies_def rev_implies_def rel_fun_def by auto  | 
| 
52354
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
316  | 
|
| 47684 | 317  | 
lemma eq_imp_transfer [transfer_rule]:  | 
318  | 
"right_unique A \<Longrightarrow> (A ===> A ===> op \<longrightarrow>) (op =) (op =)"  | 
|
319  | 
unfolding right_unique_alt_def .  | 
|
320  | 
||
| 47636 | 321  | 
lemma eq_transfer [transfer_rule]:  | 
| 47325 | 322  | 
assumes "bi_unique A"  | 
323  | 
shows "(A ===> A ===> op =) (op =) (op =)"  | 
|
| 55945 | 324  | 
using assms unfolding bi_unique_def rel_fun_def by auto  | 
| 47325 | 325  | 
|
| 
51956
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
326  | 
lemma right_total_Ex_transfer[transfer_rule]:  | 
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
327  | 
assumes "right_total A"  | 
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
328  | 
shows "((A ===> op=) ===> op=) (Bex (Collect (Domainp A))) Ex"  | 
| 55945 | 329  | 
using assms unfolding right_total_def Bex_def rel_fun_def Domainp_iff[abs_def]  | 
| 56085 | 330  | 
by fast  | 
| 
51956
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
331  | 
|
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
332  | 
lemma right_total_All_transfer[transfer_rule]:  | 
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
333  | 
assumes "right_total A"  | 
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
334  | 
shows "((A ===> op =) ===> op =) (Ball (Collect (Domainp A))) All"  | 
| 55945 | 335  | 
using assms unfolding right_total_def Ball_def rel_fun_def Domainp_iff[abs_def]  | 
| 56085 | 336  | 
by fast  | 
| 
51956
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
337  | 
|
| 47636 | 338  | 
lemma All_transfer [transfer_rule]:  | 
| 47325 | 339  | 
assumes "bi_total A"  | 
340  | 
shows "((A ===> op =) ===> op =) All All"  | 
|
| 55945 | 341  | 
using assms unfolding bi_total_def rel_fun_def by fast  | 
| 47325 | 342  | 
|
| 47636 | 343  | 
lemma Ex_transfer [transfer_rule]:  | 
| 47325 | 344  | 
assumes "bi_total A"  | 
345  | 
shows "((A ===> op =) ===> op =) Ex Ex"  | 
|
| 55945 | 346  | 
using assms unfolding bi_total_def rel_fun_def by fast  | 
| 47325 | 347  | 
|
| 47636 | 348  | 
lemma If_transfer [transfer_rule]: "(op = ===> A ===> A ===> A) If If"  | 
| 55945 | 349  | 
unfolding rel_fun_def by simp  | 
| 47325 | 350  | 
|
| 47636 | 351  | 
lemma Let_transfer [transfer_rule]: "(A ===> (A ===> B) ===> B) Let Let"  | 
| 55945 | 352  | 
unfolding rel_fun_def by simp  | 
| 47612 | 353  | 
|
| 47636 | 354  | 
lemma id_transfer [transfer_rule]: "(A ===> A) id id"  | 
| 55945 | 355  | 
unfolding rel_fun_def by simp  | 
| 47625 | 356  | 
|
| 47636 | 357  | 
lemma comp_transfer [transfer_rule]:  | 
| 47325 | 358  | 
"((B ===> C) ===> (A ===> B) ===> (A ===> C)) (op \<circ>) (op \<circ>)"  | 
| 55945 | 359  | 
unfolding rel_fun_def by simp  | 
| 47325 | 360  | 
|
| 47636 | 361  | 
lemma fun_upd_transfer [transfer_rule]:  | 
| 47325 | 362  | 
assumes [transfer_rule]: "bi_unique A"  | 
363  | 
shows "((A ===> B) ===> A ===> B ===> A ===> B) fun_upd fun_upd"  | 
|
| 
47635
 
ebb79474262c
rename 'correspondence' method to 'transfer_prover'
 
huffman 
parents: 
47627 
diff
changeset
 | 
364  | 
unfolding fun_upd_def [abs_def] by transfer_prover  | 
| 47325 | 365  | 
|
| 55415 | 366  | 
lemma case_nat_transfer [transfer_rule]:  | 
367  | 
"(A ===> (op = ===> A) ===> op = ===> A) case_nat case_nat"  | 
|
| 55945 | 368  | 
unfolding rel_fun_def by (simp split: nat.split)  | 
| 
47627
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
369  | 
|
| 55415 | 370  | 
lemma rec_nat_transfer [transfer_rule]:  | 
371  | 
"(A ===> (op = ===> A ===> A) ===> op = ===> A) rec_nat rec_nat"  | 
|
| 55945 | 372  | 
unfolding rel_fun_def by (clarsimp, rename_tac n, induct_tac n, simp_all)  | 
| 47924 | 373  | 
|
374  | 
lemma funpow_transfer [transfer_rule]:  | 
|
375  | 
"(op = ===> (A ===> A) ===> (A ===> A)) compow compow"  | 
|
376  | 
unfolding funpow_def by transfer_prover  | 
|
377  | 
||
| 53952 | 378  | 
lemma mono_transfer[transfer_rule]:  | 
379  | 
assumes [transfer_rule]: "bi_total A"  | 
|
380  | 
assumes [transfer_rule]: "(A ===> A ===> op=) op\<le> op\<le>"  | 
|
381  | 
assumes [transfer_rule]: "(B ===> B ===> op=) op\<le> op\<le>"  | 
|
382  | 
shows "((A ===> B) ===> op=) mono mono"  | 
|
383  | 
unfolding mono_def[abs_def] by transfer_prover  | 
|
384  | 
||
385  | 
lemma right_total_relcompp_transfer[transfer_rule]:  | 
|
386  | 
assumes [transfer_rule]: "right_total B"  | 
|
387  | 
shows "((A ===> B ===> op=) ===> (B ===> C ===> op=) ===> A ===> C ===> op=)  | 
|
388  | 
(\<lambda>R S x z. \<exists>y\<in>Collect (Domainp B). R x y \<and> S y z) op OO"  | 
|
389  | 
unfolding OO_def[abs_def] by transfer_prover  | 
|
390  | 
||
391  | 
lemma relcompp_transfer[transfer_rule]:  | 
|
392  | 
assumes [transfer_rule]: "bi_total B"  | 
|
393  | 
shows "((A ===> B ===> op=) ===> (B ===> C ===> op=) ===> A ===> C ===> op=) op OO op OO"  | 
|
394  | 
unfolding OO_def[abs_def] by transfer_prover  | 
|
| 
47627
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
395  | 
|
| 53952 | 396  | 
lemma right_total_Domainp_transfer[transfer_rule]:  | 
397  | 
assumes [transfer_rule]: "right_total B"  | 
|
398  | 
shows "((A ===> B ===> op=) ===> A ===> op=) (\<lambda>T x. \<exists>y\<in>Collect(Domainp B). T x y) Domainp"  | 
|
399  | 
apply(subst(2) Domainp_iff[abs_def]) by transfer_prover  | 
|
400  | 
||
401  | 
lemma Domainp_transfer[transfer_rule]:  | 
|
402  | 
assumes [transfer_rule]: "bi_total B"  | 
|
403  | 
shows "((A ===> B ===> op=) ===> A ===> op=) Domainp Domainp"  | 
|
404  | 
unfolding Domainp_iff[abs_def] by transfer_prover  | 
|
405  | 
||
406  | 
lemma reflp_transfer[transfer_rule]:  | 
|
407  | 
"bi_total A \<Longrightarrow> ((A ===> A ===> op=) ===> op=) reflp reflp"  | 
|
408  | 
"right_total A \<Longrightarrow> ((A ===> A ===> implies) ===> implies) reflp reflp"  | 
|
409  | 
"right_total A \<Longrightarrow> ((A ===> A ===> op=) ===> implies) reflp reflp"  | 
|
410  | 
"bi_total A \<Longrightarrow> ((A ===> A ===> rev_implies) ===> rev_implies) reflp reflp"  | 
|
411  | 
"bi_total A \<Longrightarrow> ((A ===> A ===> op=) ===> rev_implies) reflp reflp"  | 
|
| 55945 | 412  | 
using assms unfolding reflp_def[abs_def] rev_implies_def bi_total_def right_total_def rel_fun_def  | 
| 53952 | 413  | 
by fast+  | 
414  | 
||
415  | 
lemma right_unique_transfer [transfer_rule]:  | 
|
416  | 
assumes [transfer_rule]: "right_total A"  | 
|
417  | 
assumes [transfer_rule]: "right_total B"  | 
|
418  | 
assumes [transfer_rule]: "bi_unique B"  | 
|
419  | 
shows "((A ===> B ===> op=) ===> implies) right_unique right_unique"  | 
|
| 55945 | 420  | 
using assms unfolding right_unique_def[abs_def] right_total_def bi_unique_def rel_fun_def  | 
| 53952 | 421  | 
by metis  | 
| 47325 | 422  | 
|
423  | 
end  | 
|
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
424  | 
|
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
425  | 
end  |