| author | haftmann | 
| Sat, 19 Oct 2019 09:15:41 +0000 | |
| changeset 70903 | c550368a4e29 | 
| parent 70491 | 8cac53925407 | 
| child 70927 | cc204e10385c | 
| 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  | 
||
| 60758 | 6  | 
section \<open>Generic theorem transfer using relations\<close>  | 
| 47325 | 7  | 
|
8  | 
theory Transfer  | 
|
| 
59275
 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 
desharna 
parents: 
59141 
diff
changeset
 | 
9  | 
imports Basic_BNF_LFPs Hilbert_Choice Metis  | 
| 47325 | 10  | 
begin  | 
11  | 
||
| 60758 | 12  | 
subsection \<open>Relator for function space\<close>  | 
| 47325 | 13  | 
|
| 63343 | 14  | 
bundle lifting_syntax  | 
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
15  | 
begin  | 
| 63343 | 16  | 
notation rel_fun (infixr "===>" 55)  | 
17  | 
notation map_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
 | 
18  | 
end  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
19  | 
|
| 63343 | 20  | 
context includes lifting_syntax  | 
| 
53011
 
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  | 
|
| 55945 | 23  | 
lemma rel_funD2:  | 
24  | 
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
 | 
25  | 
shows "B (f x) (g x)"  | 
| 55945 | 26  | 
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
 | 
27  | 
|
| 55945 | 28  | 
lemma rel_funE:  | 
29  | 
assumes "rel_fun A B f g" and "A x y"  | 
|
| 47325 | 30  | 
obtains "B (f x) (g y)"  | 
| 55945 | 31  | 
using assms by (simp add: rel_fun_def)  | 
| 47325 | 32  | 
|
| 55945 | 33  | 
lemmas rel_fun_eq = fun.rel_eq  | 
| 47325 | 34  | 
|
| 55945 | 35  | 
lemma rel_fun_eq_rel:  | 
| 67399 | 36  | 
shows "rel_fun (=) R = (\<lambda>f g. \<forall>x. R (f x) (g x))"  | 
| 55945 | 37  | 
by (simp add: rel_fun_def)  | 
| 47325 | 38  | 
|
39  | 
||
| 60758 | 40  | 
subsection \<open>Transfer method\<close>  | 
| 47325 | 41  | 
|
| 60758 | 42  | 
text \<open>Explicit tag for relation membership allows for  | 
43  | 
backward proof methods.\<close>  | 
|
| 47325 | 44  | 
|
45  | 
definition Rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
 | 
|
46  | 
where "Rel r \<equiv> r"  | 
|
47  | 
||
| 60758 | 48  | 
text \<open>Handling of equality relations\<close>  | 
| 
49975
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
48891 
diff
changeset
 | 
49  | 
|
| 
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
48891 
diff
changeset
 | 
50  | 
definition is_equality :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 67399 | 51  | 
where "is_equality R \<longleftrightarrow> R = (=)"  | 
| 
49975
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
48891 
diff
changeset
 | 
52  | 
|
| 67399 | 53  | 
lemma is_equality_eq: "is_equality (=)"  | 
| 
51437
 
8739f8abbecb
fixing transfer tactic - unfold fully identity relation by using relator_eq
 
kuncar 
parents: 
51112 
diff
changeset
 | 
54  | 
unfolding is_equality_def by simp  | 
| 
 
8739f8abbecb
fixing transfer tactic - unfold fully identity relation by using relator_eq
 
kuncar 
parents: 
51112 
diff
changeset
 | 
55  | 
|
| 60758 | 56  | 
text \<open>Reverse implication for monotonicity rules\<close>  | 
| 
52354
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
57  | 
|
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
58  | 
definition rev_implies where  | 
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
59  | 
"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
 | 
60  | 
|
| 60758 | 61  | 
text \<open>Handling of meta-logic connectives\<close>  | 
| 47325 | 62  | 
|
63  | 
definition transfer_forall where  | 
|
64  | 
"transfer_forall \<equiv> All"  | 
|
65  | 
||
66  | 
definition transfer_implies where  | 
|
| 67399 | 67  | 
"transfer_implies \<equiv> (\<longrightarrow>)"  | 
| 47325 | 68  | 
|
| 
47355
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
69  | 
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
 | 
70  | 
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
 | 
71  | 
|
| 47325 | 72  | 
lemma transfer_forall_eq: "(\<And>x. P x) \<equiv> Trueprop (transfer_forall (\<lambda>x. P x))"  | 
73  | 
unfolding atomize_all transfer_forall_def ..  | 
|
74  | 
||
75  | 
lemma transfer_implies_eq: "(A \<Longrightarrow> B) \<equiv> Trueprop (transfer_implies A B)"  | 
|
76  | 
unfolding atomize_imp transfer_implies_def ..  | 
|
77  | 
||
| 
47355
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
78  | 
lemma transfer_bforall_unfold:  | 
| 
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
79  | 
"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
 | 
80  | 
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
 | 
81  | 
|
| 67399 | 82  | 
lemma transfer_start: "\<lbrakk>P; Rel (=) P Q\<rbrakk> \<Longrightarrow> Q"  | 
| 47325 | 83  | 
unfolding Rel_def by simp  | 
84  | 
||
| 67399 | 85  | 
lemma transfer_start': "\<lbrakk>P; Rel (\<longrightarrow>) P Q\<rbrakk> \<Longrightarrow> Q"  | 
| 47325 | 86  | 
unfolding Rel_def by simp  | 
87  | 
||
| 
47635
 
ebb79474262c
rename 'correspondence' method to 'transfer_prover'
 
huffman 
parents: 
47627 
diff
changeset
 | 
88  | 
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
 | 
89  | 
by simp  | 
| 
 
1568dadd598a
make correspondence tactic more robust by replacing lhs with schematic variable before applying intro rules
 
huffman 
parents: 
47612 
diff
changeset
 | 
90  | 
|
| 67399 | 91  | 
lemma untransfer_start: "\<lbrakk>Q; Rel (=) P Q\<rbrakk> \<Longrightarrow> P"  | 
| 
52358
 
f4c4bcb0d564
implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
 
huffman 
parents: 
52354 
diff
changeset
 | 
92  | 
unfolding Rel_def by simp  | 
| 
 
f4c4bcb0d564
implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
 
huffman 
parents: 
52354 
diff
changeset
 | 
93  | 
|
| 67399 | 94  | 
lemma Rel_eq_refl: "Rel (=) x x"  | 
| 47325 | 95  | 
unfolding Rel_def ..  | 
96  | 
||
| 
47789
 
71a526ee569a
implement transfer tactic with more scalable forward proof methods
 
huffman 
parents: 
47684 
diff
changeset
 | 
97  | 
lemma Rel_app:  | 
| 
47523
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
98  | 
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
 | 
99  | 
shows "Rel B (f x) (g y)"  | 
| 55945 | 100  | 
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
 | 
101  | 
|
| 
47789
 
71a526ee569a
implement transfer tactic with more scalable forward proof methods
 
huffman 
parents: 
47684 
diff
changeset
 | 
102  | 
lemma Rel_abs:  | 
| 
47523
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
103  | 
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
 | 
104  | 
shows "Rel (A ===> B) (\<lambda>x. f x) (\<lambda>y. g y)"  | 
| 55945 | 105  | 
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
 | 
106  | 
|
| 60758 | 107  | 
subsection \<open>Predicates on relations, i.e. ``class constraints''\<close>  | 
| 47325 | 108  | 
|
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
109  | 
definition left_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
110  | 
where "left_total R \<longleftrightarrow> (\<forall>x. \<exists>y. R x y)"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
111  | 
|
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
112  | 
definition left_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
113  | 
where "left_unique R \<longleftrightarrow> (\<forall>x y z. R x z \<longrightarrow> R y z \<longrightarrow> x = y)"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
114  | 
|
| 47325 | 115  | 
definition right_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
116  | 
where "right_total R \<longleftrightarrow> (\<forall>y. \<exists>x. R x y)"  | 
|
117  | 
||
118  | 
definition right_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
|
119  | 
where "right_unique R \<longleftrightarrow> (\<forall>x y z. R x y \<longrightarrow> R x z \<longrightarrow> y = z)"  | 
|
120  | 
||
121  | 
definition bi_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
|
122  | 
where "bi_total R \<longleftrightarrow> (\<forall>x. \<exists>y. R x y) \<and> (\<forall>y. \<exists>x. R x y)"  | 
|
123  | 
||
124  | 
definition bi_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
|
125  | 
where "bi_unique R \<longleftrightarrow>  | 
|
126  | 
(\<forall>x y z. R x y \<longrightarrow> R x z \<longrightarrow> y = z) \<and>  | 
|
127  | 
(\<forall>x y z. R x z \<longrightarrow> R y z \<longrightarrow> x = y)"  | 
|
128  | 
||
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
129  | 
lemma left_uniqueI: "(\<And>x y z. \<lbrakk> A x z; A y z \<rbrakk> \<Longrightarrow> x = y) \<Longrightarrow> left_unique A"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
130  | 
unfolding left_unique_def by blast  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
131  | 
|
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
132  | 
lemma left_uniqueD: "\<lbrakk> left_unique A; A x z; A y z \<rbrakk> \<Longrightarrow> x = y"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
133  | 
unfolding left_unique_def by blast  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
134  | 
|
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
135  | 
lemma left_totalI:  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
136  | 
"(\<And>x. \<exists>y. R x y) \<Longrightarrow> left_total R"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
137  | 
unfolding left_total_def by blast  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
138  | 
|
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
139  | 
lemma left_totalE:  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
140  | 
assumes "left_total R"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
141  | 
obtains "(\<And>x. \<exists>y. R x y)"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
142  | 
using assms unfolding left_total_def by blast  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
143  | 
|
| 53927 | 144  | 
lemma bi_uniqueDr: "\<lbrakk> bi_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z"  | 
145  | 
by(simp add: bi_unique_def)  | 
|
146  | 
||
147  | 
lemma bi_uniqueDl: "\<lbrakk> bi_unique A; A x y; A z y \<rbrakk> \<Longrightarrow> x = z"  | 
|
148  | 
by(simp add: bi_unique_def)  | 
|
149  | 
||
150  | 
lemma right_uniqueI: "(\<And>x y z. \<lbrakk> A x y; A x z \<rbrakk> \<Longrightarrow> y = z) \<Longrightarrow> right_unique A"  | 
|
| 56085 | 151  | 
unfolding right_unique_def by fast  | 
| 53927 | 152  | 
|
153  | 
lemma right_uniqueD: "\<lbrakk> right_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z"  | 
|
| 56085 | 154  | 
unfolding right_unique_def by fast  | 
| 53927 | 155  | 
|
| 
59514
 
509caf5edfa6
add intro and elim rules for right_total
 
Andreas Lochbihler 
parents: 
59276 
diff
changeset
 | 
156  | 
lemma right_totalI: "(\<And>y. \<exists>x. A x y) \<Longrightarrow> right_total A"  | 
| 
 
509caf5edfa6
add intro and elim rules for right_total
 
Andreas Lochbihler 
parents: 
59276 
diff
changeset
 | 
157  | 
by(simp add: right_total_def)  | 
| 
 
509caf5edfa6
add intro and elim rules for right_total
 
Andreas Lochbihler 
parents: 
59276 
diff
changeset
 | 
158  | 
|
| 
 
509caf5edfa6
add intro and elim rules for right_total
 
Andreas Lochbihler 
parents: 
59276 
diff
changeset
 | 
159  | 
lemma right_totalE:  | 
| 
 
509caf5edfa6
add intro and elim rules for right_total
 
Andreas Lochbihler 
parents: 
59276 
diff
changeset
 | 
160  | 
assumes "right_total A"  | 
| 
 
509caf5edfa6
add intro and elim rules for right_total
 
Andreas Lochbihler 
parents: 
59276 
diff
changeset
 | 
161  | 
obtains x where "A x y"  | 
| 
 
509caf5edfa6
add intro and elim rules for right_total
 
Andreas Lochbihler 
parents: 
59276 
diff
changeset
 | 
162  | 
using assms by(auto simp add: right_total_def)  | 
| 
 
509caf5edfa6
add intro and elim rules for right_total
 
Andreas Lochbihler 
parents: 
59276 
diff
changeset
 | 
163  | 
|
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
164  | 
lemma right_total_alt_def2:  | 
| 67399 | 165  | 
"right_total R \<longleftrightarrow> ((R ===> (\<longrightarrow>)) ===> (\<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  | 
||
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
174  | 
lemma right_unique_alt_def2:  | 
| 67399 | 175  | 
"right_unique R \<longleftrightarrow> (R ===> R ===> (\<longrightarrow>)) (=) (=)"  | 
| 55945 | 176  | 
unfolding right_unique_def rel_fun_def by auto  | 
| 47325 | 177  | 
|
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
178  | 
lemma bi_total_alt_def2:  | 
| 67399 | 179  | 
"bi_total R \<longleftrightarrow> ((R ===> (=)) ===> (=)) 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  | 
||
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
191  | 
lemma bi_unique_alt_def2:  | 
| 67399 | 192  | 
"bi_unique R \<longleftrightarrow> (R ===> R ===> (=)) (=) (=)"  | 
| 55945 | 193  | 
unfolding bi_unique_def rel_fun_def by auto  | 
| 47325 | 194  | 
|
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
195  | 
lemma [simp]:  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
196  | 
shows left_unique_conversep: "left_unique A\<inverse>\<inverse> \<longleftrightarrow> right_unique A"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
197  | 
and right_unique_conversep: "right_unique A\<inverse>\<inverse> \<longleftrightarrow> left_unique A"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
198  | 
by(auto simp add: left_unique_def right_unique_def)  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
199  | 
|
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
200  | 
lemma [simp]:  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
201  | 
shows left_total_conversep: "left_total A\<inverse>\<inverse> \<longleftrightarrow> right_total A"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
202  | 
and right_total_conversep: "right_total A\<inverse>\<inverse> \<longleftrightarrow> left_total A"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
203  | 
by(simp_all add: left_total_def right_total_def)  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
204  | 
|
| 53944 | 205  | 
lemma bi_unique_conversep [simp]: "bi_unique R\<inverse>\<inverse> = bi_unique R"  | 
206  | 
by(auto simp add: bi_unique_def)  | 
|
207  | 
||
208  | 
lemma bi_total_conversep [simp]: "bi_total R\<inverse>\<inverse> = bi_total R"  | 
|
209  | 
by(auto simp add: bi_total_def)  | 
|
210  | 
||
| 67399 | 211  | 
lemma right_unique_alt_def: "right_unique R = (conversep R OO R \<le> (=))" unfolding right_unique_def by blast  | 
212  | 
lemma left_unique_alt_def: "left_unique R = (R OO (conversep R) \<le> (=))" unfolding left_unique_def by blast  | 
|
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
213  | 
|
| 67399 | 214  | 
lemma right_total_alt_def: "right_total R = (conversep R OO R \<ge> (=))" unfolding right_total_def by blast  | 
215  | 
lemma left_total_alt_def: "left_total R = (R OO conversep R \<ge> (=))" unfolding left_total_def by blast  | 
|
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
216  | 
|
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
217  | 
lemma bi_total_alt_def: "bi_total A = (left_total A \<and> right_total A)"  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
218  | 
unfolding left_total_def right_total_def bi_total_def by blast  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
219  | 
|
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
220  | 
lemma bi_unique_alt_def: "bi_unique A = (left_unique A \<and> right_unique A)"  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
221  | 
unfolding left_unique_def right_unique_def bi_unique_def by blast  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
222  | 
|
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
223  | 
lemma bi_totalI: "left_total R \<Longrightarrow> right_total R \<Longrightarrow> bi_total R"  | 
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
224  | 
unfolding bi_total_alt_def ..  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
225  | 
|
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
226  | 
lemma bi_uniqueI: "left_unique R \<Longrightarrow> right_unique R \<Longrightarrow> bi_unique R"  | 
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
227  | 
unfolding bi_unique_alt_def ..  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
228  | 
|
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
229  | 
end  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
230  | 
|
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
231  | 
|
| 70491 | 232  | 
lemma is_equality_lemma: "(\<And>R. is_equality R \<Longrightarrow> PROP (P R)) \<equiv> PROP (P (=))"  | 
233  | 
apply (unfold is_equality_def)  | 
|
234  | 
apply (rule equal_intr_rule)  | 
|
235  | 
apply (drule meta_spec)  | 
|
236  | 
apply (erule meta_mp)  | 
|
237  | 
apply (rule refl)  | 
|
238  | 
apply simp  | 
|
239  | 
done  | 
|
240  | 
||
241  | 
lemma Domainp_lemma: "(\<And>R. Domainp T = R \<Longrightarrow> PROP (P R)) \<equiv> PROP (P (Domainp T))"  | 
|
242  | 
apply (rule equal_intr_rule)  | 
|
243  | 
apply (drule meta_spec)  | 
|
244  | 
apply (erule meta_mp)  | 
|
245  | 
apply (rule refl)  | 
|
246  | 
apply simp  | 
|
247  | 
done  | 
|
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
248  | 
|
| 69605 | 249  | 
ML_file \<open>Tools/Transfer/transfer.ML\<close>  | 
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
250  | 
declare refl [transfer_rule]  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
251  | 
|
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
252  | 
hide_const (open) Rel  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
253  | 
|
| 63343 | 254  | 
context includes lifting_syntax  | 
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
255  | 
begin  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
256  | 
|
| 60758 | 257  | 
text \<open>Handling of domains\<close>  | 
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
258  | 
|
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
259  | 
lemma Domainp_iff: "Domainp T x \<longleftrightarrow> (\<exists>y. T x y)"  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
260  | 
by auto  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
261  | 
|
| 58386 | 262  | 
lemma Domainp_refl[transfer_domain_rule]:  | 
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
263  | 
"Domainp T = Domainp T" ..  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
264  | 
|
| 67399 | 265  | 
lemma Domain_eq_top[transfer_domain_rule]: "Domainp (=) = top" by auto  | 
| 
60229
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
59523 
diff
changeset
 | 
266  | 
|
| 
64425
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
267  | 
lemma Domainp_pred_fun_eq[relator_domain]:  | 
| 
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
268  | 
assumes "left_unique T"  | 
| 
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
269  | 
shows "Domainp (T ===> S) = pred_fun (Domainp T) (Domainp S)"  | 
| 
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
270  | 
using assms unfolding rel_fun_def Domainp_iff[abs_def] left_unique_def fun_eq_iff pred_fun_def  | 
| 
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
271  | 
apply safe  | 
| 
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
272  | 
apply blast  | 
| 
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
273  | 
apply (subst all_comm)  | 
| 
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
274  | 
apply (rule choice)  | 
| 
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
275  | 
apply blast  | 
| 
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
276  | 
done  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
277  | 
|
| 60758 | 278  | 
text \<open>Properties are preserved by relation composition.\<close>  | 
| 47660 | 279  | 
|
280  | 
lemma OO_def: "R OO S = (\<lambda>x z. \<exists>y. R x y \<and> S y z)"  | 
|
281  | 
by auto  | 
|
282  | 
||
283  | 
lemma bi_total_OO: "\<lbrakk>bi_total A; bi_total B\<rbrakk> \<Longrightarrow> bi_total (A OO B)"  | 
|
| 56085 | 284  | 
unfolding bi_total_def OO_def by fast  | 
| 47660 | 285  | 
|
286  | 
lemma bi_unique_OO: "\<lbrakk>bi_unique A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A OO B)"  | 
|
| 56085 | 287  | 
unfolding bi_unique_def OO_def by blast  | 
| 47660 | 288  | 
|
289  | 
lemma right_total_OO:  | 
|
290  | 
"\<lbrakk>right_total A; right_total B\<rbrakk> \<Longrightarrow> right_total (A OO B)"  | 
|
| 56085 | 291  | 
unfolding right_total_def OO_def by fast  | 
| 47660 | 292  | 
|
293  | 
lemma right_unique_OO:  | 
|
294  | 
"\<lbrakk>right_unique A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A OO B)"  | 
|
| 56085 | 295  | 
unfolding right_unique_def OO_def by fast  | 
| 47660 | 296  | 
|
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
297  | 
lemma left_total_OO: "left_total R \<Longrightarrow> left_total S \<Longrightarrow> left_total (R OO S)"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
298  | 
unfolding left_total_def OO_def by fast  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
299  | 
|
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
300  | 
lemma left_unique_OO: "left_unique R \<Longrightarrow> left_unique S \<Longrightarrow> left_unique (R OO S)"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
301  | 
unfolding left_unique_def OO_def by blast  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
302  | 
|
| 47325 | 303  | 
|
| 60758 | 304  | 
subsection \<open>Properties of relators\<close>  | 
| 47325 | 305  | 
|
| 67399 | 306  | 
lemma left_total_eq[transfer_rule]: "left_total (=)"  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
307  | 
unfolding left_total_def by blast  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
308  | 
|
| 67399 | 309  | 
lemma left_unique_eq[transfer_rule]: "left_unique (=)"  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
310  | 
unfolding left_unique_def by blast  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
311  | 
|
| 67399 | 312  | 
lemma right_total_eq [transfer_rule]: "right_total (=)"  | 
| 47325 | 313  | 
unfolding right_total_def by simp  | 
314  | 
||
| 67399 | 315  | 
lemma right_unique_eq [transfer_rule]: "right_unique (=)"  | 
| 47325 | 316  | 
unfolding right_unique_def by simp  | 
317  | 
||
| 67399 | 318  | 
lemma bi_total_eq[transfer_rule]: "bi_total (=)"  | 
| 47325 | 319  | 
unfolding bi_total_def by simp  | 
320  | 
||
| 67399 | 321  | 
lemma bi_unique_eq[transfer_rule]: "bi_unique (=)"  | 
| 47325 | 322  | 
unfolding bi_unique_def by simp  | 
323  | 
||
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
324  | 
lemma left_total_fun[transfer_rule]:  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
325  | 
"\<lbrakk>left_unique A; left_total B\<rbrakk> \<Longrightarrow> left_total (A ===> B)"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
326  | 
unfolding left_total_def rel_fun_def  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
327  | 
apply (rule allI, rename_tac f)  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
328  | 
apply (rule_tac x="\<lambda>y. SOME z. B (f (THE x. A x y)) z" in exI)  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
329  | 
apply clarify  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
330  | 
apply (subgoal_tac "(THE x. A x y) = x", simp)  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
331  | 
apply (rule someI_ex)  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
332  | 
apply (simp)  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
333  | 
apply (rule the_equality)  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
334  | 
apply assumption  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
335  | 
apply (simp add: left_unique_def)  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
336  | 
done  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
337  | 
|
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
338  | 
lemma left_unique_fun[transfer_rule]:  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
339  | 
"\<lbrakk>left_total A; left_unique B\<rbrakk> \<Longrightarrow> left_unique (A ===> B)"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
340  | 
unfolding left_total_def left_unique_def rel_fun_def  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
341  | 
by (clarify, rule ext, fast)  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
342  | 
|
| 47325 | 343  | 
lemma right_total_fun [transfer_rule]:  | 
344  | 
"\<lbrakk>right_unique A; right_total B\<rbrakk> \<Longrightarrow> right_total (A ===> B)"  | 
|
| 55945 | 345  | 
unfolding right_total_def rel_fun_def  | 
| 47325 | 346  | 
apply (rule allI, rename_tac g)  | 
347  | 
apply (rule_tac x="\<lambda>x. SOME z. B z (g (THE y. A x y))" in exI)  | 
|
348  | 
apply clarify  | 
|
349  | 
apply (subgoal_tac "(THE y. A x y) = y", simp)  | 
|
350  | 
apply (rule someI_ex)  | 
|
351  | 
apply (simp)  | 
|
352  | 
apply (rule the_equality)  | 
|
353  | 
apply assumption  | 
|
354  | 
apply (simp add: right_unique_def)  | 
|
355  | 
done  | 
|
356  | 
||
357  | 
lemma right_unique_fun [transfer_rule]:  | 
|
358  | 
"\<lbrakk>right_total A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A ===> B)"  | 
|
| 55945 | 359  | 
unfolding right_total_def right_unique_def rel_fun_def  | 
| 47325 | 360  | 
by (clarify, rule ext, fast)  | 
361  | 
||
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
362  | 
lemma bi_total_fun[transfer_rule]:  | 
| 47325 | 363  | 
"\<lbrakk>bi_unique A; bi_total B\<rbrakk> \<Longrightarrow> bi_total (A ===> B)"  | 
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
364  | 
unfolding bi_unique_alt_def bi_total_alt_def  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
365  | 
by (blast intro: right_total_fun left_total_fun)  | 
| 47325 | 366  | 
|
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
367  | 
lemma bi_unique_fun[transfer_rule]:  | 
| 47325 | 368  | 
"\<lbrakk>bi_total A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A ===> B)"  | 
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
369  | 
unfolding bi_unique_alt_def bi_total_alt_def  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
370  | 
by (blast intro: right_unique_fun left_unique_fun)  | 
| 47325 | 371  | 
|
| 
56543
 
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
 
kuncar 
parents: 
56524 
diff
changeset
 | 
372  | 
end  | 
| 
 
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
 
kuncar 
parents: 
56524 
diff
changeset
 | 
373  | 
|
| 
59275
 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 
desharna 
parents: 
59141 
diff
changeset
 | 
374  | 
lemma if_conn:  | 
| 
 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 
desharna 
parents: 
59141 
diff
changeset
 | 
375  | 
"(if P \<and> Q then t else e) = (if P then if Q then t else e else e)"  | 
| 
 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 
desharna 
parents: 
59141 
diff
changeset
 | 
376  | 
"(if P \<or> Q then t else e) = (if P then t else if Q then t else e)"  | 
| 
 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 
desharna 
parents: 
59141 
diff
changeset
 | 
377  | 
"(if P \<longrightarrow> Q then t else e) = (if P then if Q then t else e else t)"  | 
| 
 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 
desharna 
parents: 
59141 
diff
changeset
 | 
378  | 
"(if \<not> P then t else e) = (if P then e else t)"  | 
| 
 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 
desharna 
parents: 
59141 
diff
changeset
 | 
379  | 
by auto  | 
| 
 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 
desharna 
parents: 
59141 
diff
changeset
 | 
380  | 
|
| 69605 | 381  | 
ML_file \<open>Tools/Transfer/transfer_bnf.ML\<close>  | 
382  | 
ML_file \<open>Tools/BNF/bnf_fp_rec_sugar_transfer.ML\<close>  | 
|
| 
59275
 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 
desharna 
parents: 
59141 
diff
changeset
 | 
383  | 
|
| 
56543
 
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
 
kuncar 
parents: 
56524 
diff
changeset
 | 
384  | 
declare pred_fun_def [simp]  | 
| 
 
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
 
kuncar 
parents: 
56524 
diff
changeset
 | 
385  | 
declare rel_fun_eq [relator_eq]  | 
| 
 
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
 
kuncar 
parents: 
56524 
diff
changeset
 | 
386  | 
|
| 
64425
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
387  | 
(* Delete the automated generated rule from the bnf command;  | 
| 
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
388  | 
we have a more general rule (Domainp_pred_fun_eq) that subsumes it. *)  | 
| 
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
389  | 
declare fun.Domainp_rel[relator_domain del]  | 
| 
 
b17acc1834e3
a more general relator domain rule for the function type
 
kuncar 
parents: 
64014 
diff
changeset
 | 
390  | 
|
| 60758 | 391  | 
subsection \<open>Transfer rules\<close>  | 
| 47325 | 392  | 
|
| 63343 | 393  | 
context includes lifting_syntax  | 
| 
56543
 
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
 
kuncar 
parents: 
56524 
diff
changeset
 | 
394  | 
begin  | 
| 
 
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
 
kuncar 
parents: 
56524 
diff
changeset
 | 
395  | 
|
| 53952 | 396  | 
lemma Domainp_forall_transfer [transfer_rule]:  | 
397  | 
assumes "right_total A"  | 
|
| 67399 | 398  | 
shows "((A ===> (=)) ===> (=))  | 
| 53952 | 399  | 
(transfer_bforall (Domainp A)) transfer_forall"  | 
400  | 
using assms unfolding right_total_def  | 
|
| 55945 | 401  | 
unfolding transfer_forall_def transfer_bforall_def rel_fun_def Domainp_iff  | 
| 56085 | 402  | 
by fast  | 
| 53952 | 403  | 
|
| 60758 | 404  | 
text \<open>Transfer rules using implication instead of equality on booleans.\<close>  | 
| 47684 | 405  | 
|
| 
52354
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
406  | 
lemma transfer_forall_transfer [transfer_rule]:  | 
| 67399 | 407  | 
"bi_total A \<Longrightarrow> ((A ===> (=)) ===> (=)) transfer_forall transfer_forall"  | 
408  | 
"right_total A \<Longrightarrow> ((A ===> (=)) ===> implies) transfer_forall transfer_forall"  | 
|
| 
52354
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
409  | 
"right_total A \<Longrightarrow> ((A ===> implies) ===> implies) transfer_forall transfer_forall"  | 
| 67399 | 410  | 
"bi_total A \<Longrightarrow> ((A ===> (=)) ===> rev_implies) transfer_forall transfer_forall"  | 
| 
52354
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
411  | 
"bi_total A \<Longrightarrow> ((A ===> rev_implies) ===> rev_implies) transfer_forall transfer_forall"  | 
| 55945 | 412  | 
unfolding transfer_forall_def rev_implies_def rel_fun_def right_total_def bi_total_def  | 
| 56085 | 413  | 
by fast+  | 
| 
52354
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
414  | 
|
| 
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
415  | 
lemma transfer_implies_transfer [transfer_rule]:  | 
| 67399 | 416  | 
"((=) ===> (=) ===> (=) ) transfer_implies transfer_implies"  | 
| 
52354
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
417  | 
"(rev_implies ===> implies ===> implies ) transfer_implies transfer_implies"  | 
| 67399 | 418  | 
"(rev_implies ===> (=) ===> implies ) transfer_implies transfer_implies"  | 
419  | 
"((=) ===> implies ===> implies ) transfer_implies transfer_implies"  | 
|
420  | 
"((=) ===> (=) ===> implies ) transfer_implies transfer_implies"  | 
|
| 
52354
 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 
huffman 
parents: 
51956 
diff
changeset
 | 
421  | 
"(implies ===> rev_implies ===> rev_implies) transfer_implies transfer_implies"  | 
| 67399 | 422  | 
"(implies ===> (=) ===> rev_implies) transfer_implies transfer_implies"  | 
423  | 
"((=) ===> rev_implies ===> rev_implies) transfer_implies transfer_implies"  | 
|
424  | 
"((=) ===> (=) ===> rev_implies) transfer_implies transfer_implies"  | 
|
| 55945 | 425  | 
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
 | 
426  | 
|
| 47684 | 427  | 
lemma eq_imp_transfer [transfer_rule]:  | 
| 67399 | 428  | 
"right_unique A \<Longrightarrow> (A ===> A ===> (\<longrightarrow>)) (=) (=)"  | 
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
429  | 
unfolding right_unique_alt_def2 .  | 
| 47684 | 430  | 
|
| 60758 | 431  | 
text \<open>Transfer rules using equality.\<close>  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
432  | 
|
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
433  | 
lemma left_unique_transfer [transfer_rule]:  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
434  | 
assumes "right_total A"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
435  | 
assumes "right_total B"  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
436  | 
assumes "bi_unique A"  | 
| 67399 | 437  | 
shows "((A ===> B ===> (=)) ===> implies) left_unique left_unique"  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
438  | 
using assms unfolding left_unique_def[abs_def] right_total_def bi_unique_def rel_fun_def  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
439  | 
by metis  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56085 
diff
changeset
 | 
440  | 
|
| 47636 | 441  | 
lemma eq_transfer [transfer_rule]:  | 
| 47325 | 442  | 
assumes "bi_unique A"  | 
| 67399 | 443  | 
shows "(A ===> A ===> (=)) (=) (=)"  | 
| 55945 | 444  | 
using assms unfolding bi_unique_def rel_fun_def by auto  | 
| 47325 | 445  | 
|
| 
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
 | 
446  | 
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
 | 
447  | 
assumes "right_total A"  | 
| 67399 | 448  | 
shows "((A ===> (=)) ===> (=)) (Bex (Collect (Domainp A))) Ex"  | 
| 55945 | 449  | 
using assms unfolding right_total_def Bex_def rel_fun_def Domainp_iff[abs_def]  | 
| 56085 | 450  | 
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
 | 
451  | 
|
| 
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51955 
diff
changeset
 | 
452  | 
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
 | 
453  | 
assumes "right_total A"  | 
| 67399 | 454  | 
shows "((A ===> (=)) ===> (=)) (Ball (Collect (Domainp A))) All"  | 
| 55945 | 455  | 
using assms unfolding right_total_def Ball_def rel_fun_def Domainp_iff[abs_def]  | 
| 56085 | 456  | 
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
 | 
457  | 
|
| 68521 | 458  | 
lemma right_total_fun_eq_transfer:  | 
459  | 
includes lifting_syntax  | 
|
460  | 
assumes [transfer_rule]: "right_total A" "bi_unique B"  | 
|
461  | 
shows "((A ===> B) ===> (A ===> B) ===> (=)) (\<lambda>f g. \<forall>x\<in>Collect(Domainp A). f x = g x) (=)"  | 
|
462  | 
unfolding fun_eq_iff  | 
|
463  | 
by transfer_prover  | 
|
464  | 
||
| 47636 | 465  | 
lemma All_transfer [transfer_rule]:  | 
| 47325 | 466  | 
assumes "bi_total A"  | 
| 67399 | 467  | 
shows "((A ===> (=)) ===> (=)) All All"  | 
| 55945 | 468  | 
using assms unfolding bi_total_def rel_fun_def by fast  | 
| 47325 | 469  | 
|
| 47636 | 470  | 
lemma Ex_transfer [transfer_rule]:  | 
| 47325 | 471  | 
assumes "bi_total A"  | 
| 67399 | 472  | 
shows "((A ===> (=)) ===> (=)) Ex Ex"  | 
| 55945 | 473  | 
using assms unfolding bi_total_def rel_fun_def by fast  | 
| 47325 | 474  | 
|
| 59515 | 475  | 
lemma Ex1_parametric [transfer_rule]:  | 
476  | 
assumes [transfer_rule]: "bi_unique A" "bi_total A"  | 
|
| 67399 | 477  | 
shows "((A ===> (=)) ===> (=)) Ex1 Ex1"  | 
| 59515 | 478  | 
unfolding Ex1_def[abs_def] by transfer_prover  | 
479  | 
||
| 58448 | 480  | 
declare If_transfer [transfer_rule]  | 
| 47325 | 481  | 
|
| 47636 | 482  | 
lemma Let_transfer [transfer_rule]: "(A ===> (A ===> B) ===> B) Let Let"  | 
| 55945 | 483  | 
unfolding rel_fun_def by simp  | 
| 47612 | 484  | 
|
| 58916 | 485  | 
declare id_transfer [transfer_rule]  | 
| 47625 | 486  | 
|
| 58444 | 487  | 
declare comp_transfer [transfer_rule]  | 
| 47325 | 488  | 
|
| 58916 | 489  | 
lemma curry_transfer [transfer_rule]:  | 
490  | 
"((rel_prod A B ===> C) ===> A ===> B ===> C) curry curry"  | 
|
491  | 
unfolding curry_def by transfer_prover  | 
|
492  | 
||
| 47636 | 493  | 
lemma fun_upd_transfer [transfer_rule]:  | 
| 47325 | 494  | 
assumes [transfer_rule]: "bi_unique A"  | 
495  | 
shows "((A ===> B) ===> A ===> B ===> A ===> B) fun_upd fun_upd"  | 
|
| 
47635
 
ebb79474262c
rename 'correspondence' method to 'transfer_prover'
 
huffman 
parents: 
47627 
diff
changeset
 | 
496  | 
unfolding fun_upd_def [abs_def] by transfer_prover  | 
| 47325 | 497  | 
|
| 55415 | 498  | 
lemma case_nat_transfer [transfer_rule]:  | 
| 67399 | 499  | 
"(A ===> ((=) ===> A) ===> (=) ===> A) case_nat case_nat"  | 
| 55945 | 500  | 
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
 | 
501  | 
|
| 55415 | 502  | 
lemma rec_nat_transfer [transfer_rule]:  | 
| 67399 | 503  | 
"(A ===> ((=) ===> A ===> A) ===> (=) ===> A) rec_nat rec_nat"  | 
| 55945 | 504  | 
unfolding rel_fun_def by (clarsimp, rename_tac n, induct_tac n, simp_all)  | 
| 47924 | 505  | 
|
506  | 
lemma funpow_transfer [transfer_rule]:  | 
|
| 67399 | 507  | 
"((=) ===> (A ===> A) ===> (A ===> A)) compow compow"  | 
| 47924 | 508  | 
unfolding funpow_def by transfer_prover  | 
509  | 
||
| 53952 | 510  | 
lemma mono_transfer[transfer_rule]:  | 
511  | 
assumes [transfer_rule]: "bi_total A"  | 
|
| 67399 | 512  | 
assumes [transfer_rule]: "(A ===> A ===> (=)) (\<le>) (\<le>)"  | 
513  | 
assumes [transfer_rule]: "(B ===> B ===> (=)) (\<le>) (\<le>)"  | 
|
514  | 
shows "((A ===> B) ===> (=)) mono mono"  | 
|
| 53952 | 515  | 
unfolding mono_def[abs_def] by transfer_prover  | 
516  | 
||
| 58182 | 517  | 
lemma right_total_relcompp_transfer[transfer_rule]:  | 
| 53952 | 518  | 
assumes [transfer_rule]: "right_total B"  | 
| 67399 | 519  | 
shows "((A ===> B ===> (=)) ===> (B ===> C ===> (=)) ===> A ===> C ===> (=))  | 
520  | 
(\<lambda>R S x z. \<exists>y\<in>Collect (Domainp B). R x y \<and> S y z) (OO)"  | 
|
| 53952 | 521  | 
unfolding OO_def[abs_def] by transfer_prover  | 
522  | 
||
| 58182 | 523  | 
lemma relcompp_transfer[transfer_rule]:  | 
| 53952 | 524  | 
assumes [transfer_rule]: "bi_total B"  | 
| 67399 | 525  | 
shows "((A ===> B ===> (=)) ===> (B ===> C ===> (=)) ===> A ===> C ===> (=)) (OO) (OO)"  | 
| 53952 | 526  | 
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
 | 
527  | 
|
| 53952 | 528  | 
lemma right_total_Domainp_transfer[transfer_rule]:  | 
529  | 
assumes [transfer_rule]: "right_total B"  | 
|
| 67399 | 530  | 
shows "((A ===> B ===> (=)) ===> A ===> (=)) (\<lambda>T x. \<exists>y\<in>Collect(Domainp B). T x y) Domainp"  | 
| 53952 | 531  | 
apply(subst(2) Domainp_iff[abs_def]) by transfer_prover  | 
532  | 
||
533  | 
lemma Domainp_transfer[transfer_rule]:  | 
|
534  | 
assumes [transfer_rule]: "bi_total B"  | 
|
| 67399 | 535  | 
shows "((A ===> B ===> (=)) ===> A ===> (=)) Domainp Domainp"  | 
| 53952 | 536  | 
unfolding Domainp_iff[abs_def] by transfer_prover  | 
537  | 
||
| 58182 | 538  | 
lemma reflp_transfer[transfer_rule]:  | 
| 67399 | 539  | 
"bi_total A \<Longrightarrow> ((A ===> A ===> (=)) ===> (=)) reflp reflp"  | 
| 53952 | 540  | 
"right_total A \<Longrightarrow> ((A ===> A ===> implies) ===> implies) reflp reflp"  | 
| 67399 | 541  | 
"right_total A \<Longrightarrow> ((A ===> A ===> (=)) ===> implies) reflp reflp"  | 
| 53952 | 542  | 
"bi_total A \<Longrightarrow> ((A ===> A ===> rev_implies) ===> rev_implies) reflp reflp"  | 
| 67399 | 543  | 
"bi_total A \<Longrightarrow> ((A ===> A ===> (=)) ===> rev_implies) reflp reflp"  | 
| 63092 | 544  | 
unfolding reflp_def[abs_def] rev_implies_def bi_total_def right_total_def rel_fun_def  | 
| 53952 | 545  | 
by fast+  | 
546  | 
||
547  | 
lemma right_unique_transfer [transfer_rule]:  | 
|
| 59523 | 548  | 
"\<lbrakk> right_total A; right_total B; bi_unique B \<rbrakk>  | 
| 67399 | 549  | 
\<Longrightarrow> ((A ===> B ===> (=)) ===> implies) right_unique right_unique"  | 
| 59523 | 550  | 
unfolding right_unique_def[abs_def] right_total_def bi_unique_def rel_fun_def  | 
| 53952 | 551  | 
by metis  | 
| 47325 | 552  | 
|
| 59523 | 553  | 
lemma left_total_parametric [transfer_rule]:  | 
554  | 
assumes [transfer_rule]: "bi_total A" "bi_total B"  | 
|
| 67399 | 555  | 
shows "((A ===> B ===> (=)) ===> (=)) left_total left_total"  | 
| 59523 | 556  | 
unfolding left_total_def[abs_def] by transfer_prover  | 
557  | 
||
558  | 
lemma right_total_parametric [transfer_rule]:  | 
|
559  | 
assumes [transfer_rule]: "bi_total A" "bi_total B"  | 
|
| 67399 | 560  | 
shows "((A ===> B ===> (=)) ===> (=)) right_total right_total"  | 
| 59523 | 561  | 
unfolding right_total_def[abs_def] by transfer_prover  | 
562  | 
||
563  | 
lemma left_unique_parametric [transfer_rule]:  | 
|
564  | 
assumes [transfer_rule]: "bi_unique A" "bi_total A" "bi_total B"  | 
|
| 67399 | 565  | 
shows "((A ===> B ===> (=)) ===> (=)) left_unique left_unique"  | 
| 59523 | 566  | 
unfolding left_unique_def[abs_def] by transfer_prover  | 
567  | 
||
568  | 
lemma prod_pred_parametric [transfer_rule]:  | 
|
| 67399 | 569  | 
"((A ===> (=)) ===> (B ===> (=)) ===> rel_prod A B ===> (=)) pred_prod pred_prod"  | 
| 62324 | 570  | 
unfolding prod.pred_set[abs_def] Basic_BNFs.fsts_def Basic_BNFs.snds_def fstsp.simps sndsp.simps  | 
| 59523 | 571  | 
by simp transfer_prover  | 
572  | 
||
573  | 
lemma apfst_parametric [transfer_rule]:  | 
|
574  | 
"((A ===> B) ===> rel_prod A C ===> rel_prod B C) apfst apfst"  | 
|
575  | 
unfolding apfst_def[abs_def] by transfer_prover  | 
|
576  | 
||
| 67399 | 577  | 
lemma rel_fun_eq_eq_onp: "((=) ===> eq_onp P) = eq_onp (\<lambda>f. \<forall>x. P(f x))"  | 
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
578  | 
unfolding eq_onp_def rel_fun_def by auto  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
579  | 
|
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
580  | 
lemma rel_fun_eq_onp_rel:  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
581  | 
shows "((eq_onp R) ===> S) = (\<lambda>f g. \<forall>x. R x \<longrightarrow> S (f x) (g x))"  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
582  | 
by (auto simp add: eq_onp_def rel_fun_def)  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
583  | 
|
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
584  | 
lemma eq_onp_transfer [transfer_rule]:  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
585  | 
assumes [transfer_rule]: "bi_unique A"  | 
| 67399 | 586  | 
shows "((A ===> (=)) ===> A ===> A ===> (=)) eq_onp eq_onp"  | 
| 
56524
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
587  | 
unfolding eq_onp_def[abs_def] by transfer_prover  | 
| 
 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 
kuncar 
parents: 
56520 
diff
changeset
 | 
588  | 
|
| 57599 | 589  | 
lemma rtranclp_parametric [transfer_rule]:  | 
590  | 
assumes "bi_unique A" "bi_total A"  | 
|
| 67399 | 591  | 
shows "((A ===> A ===> (=)) ===> A ===> A ===> (=)) rtranclp rtranclp"  | 
| 57599 | 592  | 
proof(rule rel_funI iffI)+  | 
593  | 
fix R :: "'a \<Rightarrow> 'a \<Rightarrow> bool" and R' x y x' y'  | 
|
| 67399 | 594  | 
assume R: "(A ===> A ===> (=)) R R'" and "A x x'"  | 
| 57599 | 595  | 
  {
 | 
596  | 
assume "R\<^sup>*\<^sup>* x y" "A y y'"  | 
|
597  | 
thus "R'\<^sup>*\<^sup>* x' y'"  | 
|
598  | 
proof(induction arbitrary: y')  | 
|
599  | 
case base  | 
|
| 60758 | 600  | 
with \<open>bi_unique A\<close> \<open>A x x'\<close> have "x' = y'" by(rule bi_uniqueDr)  | 
| 57599 | 601  | 
thus ?case by simp  | 
602  | 
next  | 
|
603  | 
case (step y z z')  | 
|
| 60758 | 604  | 
from \<open>bi_total A\<close> obtain y' where "A y y'" unfolding bi_total_def by blast  | 
| 57599 | 605  | 
hence "R'\<^sup>*\<^sup>* x' y'" by(rule step.IH)  | 
| 60758 | 606  | 
moreover from R \<open>A y y'\<close> \<open>A z z'\<close> \<open>R y z\<close>  | 
| 57599 | 607  | 
have "R' y' z'" by(auto dest: rel_funD)  | 
608  | 
ultimately show ?case ..  | 
|
609  | 
qed  | 
|
610  | 
next  | 
|
611  | 
assume "R'\<^sup>*\<^sup>* x' y'" "A y y'"  | 
|
612  | 
thus "R\<^sup>*\<^sup>* x y"  | 
|
613  | 
proof(induction arbitrary: y)  | 
|
614  | 
case base  | 
|
| 60758 | 615  | 
with \<open>bi_unique A\<close> \<open>A x x'\<close> have "x = y" by(rule bi_uniqueDl)  | 
| 57599 | 616  | 
thus ?case by simp  | 
617  | 
next  | 
|
618  | 
case (step y' z' z)  | 
|
| 60758 | 619  | 
from \<open>bi_total A\<close> obtain y where "A y y'" unfolding bi_total_def by blast  | 
| 57599 | 620  | 
hence "R\<^sup>*\<^sup>* x y" by(rule step.IH)  | 
| 60758 | 621  | 
moreover from R \<open>A y y'\<close> \<open>A z z'\<close> \<open>R' y' z'\<close>  | 
| 57599 | 622  | 
have "R y z" by(auto dest: rel_funD)  | 
623  | 
ultimately show ?case ..  | 
|
624  | 
qed  | 
|
625  | 
}  | 
|
626  | 
qed  | 
|
627  | 
||
| 59523 | 628  | 
lemma right_unique_parametric [transfer_rule]:  | 
629  | 
assumes [transfer_rule]: "bi_total A" "bi_unique B" "bi_total B"  | 
|
| 67399 | 630  | 
shows "((A ===> B ===> (=)) ===> (=)) right_unique right_unique"  | 
| 59523 | 631  | 
unfolding right_unique_def[abs_def] by transfer_prover  | 
632  | 
||
| 61630 | 633  | 
lemma map_fun_parametric [transfer_rule]:  | 
634  | 
"((A ===> B) ===> (C ===> D) ===> (B ===> C) ===> A ===> D) map_fun map_fun"  | 
|
635  | 
unfolding map_fun_def[abs_def] by transfer_prover  | 
|
636  | 
||
| 47325 | 637  | 
end  | 
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
638  | 
|
| 64014 | 639  | 
|
| 69593 | 640  | 
subsection \<open>\<^const>\<open>of_nat\<close>\<close>  | 
| 64014 | 641  | 
|
642  | 
lemma transfer_rule_of_nat:  | 
|
643  | 
fixes R :: "'a::semiring_1 \<Rightarrow> 'b::semiring_1 \<Rightarrow> bool"  | 
|
644  | 
assumes [transfer_rule]: "R 0 0" "R 1 1"  | 
|
645  | 
"rel_fun R (rel_fun R R) plus plus"  | 
|
646  | 
shows "rel_fun HOL.eq R of_nat of_nat"  | 
|
647  | 
by (unfold of_nat_def [abs_def]) transfer_prover  | 
|
648  | 
||
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
52358 
diff
changeset
 | 
649  | 
end  |