| author | Andreas Lochbihler | 
| Tue, 01 Dec 2015 12:35:11 +0100 | |
| changeset 61766 | 507b39df1a57 | 
| parent 61630 | 608520e0e8e2 | 
| child 62324 | ae44f16dcea5 | 
| 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: 
51955diff
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: 
59141diff
changeset | 9 | imports Basic_BNF_LFPs Hilbert_Choice Metis | 
| 47325 | 10 | begin | 
| 11 | ||
| 60758 | 12 | subsection \<open>Relator for function space\<close> | 
| 47325 | 13 | |
| 53011 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 kuncar parents: 
52358diff
changeset | 14 | locale lifting_syntax | 
| 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 kuncar parents: 
52358diff
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: 
52358diff
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: 
52358diff
changeset | 18 | end | 
| 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 kuncar parents: 
52358diff
changeset | 19 | |
| 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 kuncar parents: 
52358diff
changeset | 20 | context | 
| 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 kuncar parents: 
52358diff
changeset | 21 | begin | 
| 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 kuncar parents: 
52358diff
changeset | 22 | interpretation lifting_syntax . | 
| 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 kuncar parents: 
52358diff
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: 
47924diff
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: 
47924diff
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 | ||
| 60758 | 41 | subsection \<open>Transfer method\<close> | 
| 47325 | 42 | |
| 60758 | 43 | text \<open>Explicit tag for relation membership allows for | 
| 44 | backward proof methods.\<close> | |
| 47325 | 45 | |
| 46 | definition Rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
 | |
| 47 | where "Rel r \<equiv> r" | |
| 48 | ||
| 60758 | 49 | text \<open>Handling of equality relations\<close> | 
| 49975 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 huffman parents: 
48891diff
changeset | 50 | |
| 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 huffman parents: 
48891diff
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: 
48891diff
changeset | 52 | where "is_equality R \<longleftrightarrow> R = (op =)" | 
| 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 huffman parents: 
48891diff
changeset | 53 | |
| 51437 
8739f8abbecb
fixing transfer tactic - unfold fully identity relation by using relator_eq
 kuncar parents: 
51112diff
changeset | 54 | lemma is_equality_eq: "is_equality (op =)" | 
| 
8739f8abbecb
fixing transfer tactic - unfold fully identity relation by using relator_eq
 kuncar parents: 
51112diff
changeset | 55 | unfolding is_equality_def by simp | 
| 
8739f8abbecb
fixing transfer tactic - unfold fully identity relation by using relator_eq
 kuncar parents: 
51112diff
changeset | 56 | |
| 60758 | 57 | text \<open>Reverse implication for monotonicity rules\<close> | 
| 52354 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 58 | |
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 59 | definition rev_implies where | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 60 | "rev_implies x y \<longleftrightarrow> (y \<longrightarrow> x)" | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 61 | |
| 60758 | 62 | text \<open>Handling of meta-logic connectives\<close> | 
| 47325 | 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: 
47325diff
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: 
47325diff
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: 
47325diff
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: 
47325diff
changeset | 79 | lemma transfer_bforall_unfold: | 
| 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 huffman parents: 
47325diff
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: 
47325diff
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: 
47325diff
changeset | 82 | |
| 47658 
7631f6f7873d
enable variant of transfer method that proves an implication instead of an equivalence
 huffman parents: 
47637diff
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: 
47637diff
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: 
47627diff
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: 
47612diff
changeset | 90 | by simp | 
| 
1568dadd598a
make correspondence tactic more robust by replacing lhs with schematic variable before applying intro rules
 huffman parents: 
47612diff
changeset | 91 | |
| 52358 
f4c4bcb0d564
implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
 huffman parents: 
52354diff
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: 
52354diff
changeset | 93 | unfolding Rel_def by simp | 
| 
f4c4bcb0d564
implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
 huffman parents: 
52354diff
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: 
47684diff
changeset | 98 | lemma Rel_app: | 
| 47523 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 huffman parents: 
47503diff
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: 
47684diff
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: 
47503diff
changeset | 102 | |
| 47789 
71a526ee569a
implement transfer tactic with more scalable forward proof methods
 huffman parents: 
47684diff
changeset | 103 | lemma Rel_abs: | 
| 47523 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 huffman parents: 
47503diff
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: 
47684diff
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: 
47503diff
changeset | 107 | |
| 60758 | 108 | subsection \<open>Predicates on relations, i.e. ``class constraints''\<close> | 
| 47325 | 109 | |
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 110 | 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: 
56085diff
changeset | 111 | 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: 
56085diff
changeset | 112 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 113 | 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: 
56085diff
changeset | 114 | 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: 
56085diff
changeset | 115 | |
| 47325 | 116 | definition right_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 117 | where "right_total R \<longleftrightarrow> (\<forall>y. \<exists>x. R x y)" | |
| 118 | ||
| 119 | definition right_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | |
| 120 | where "right_unique R \<longleftrightarrow> (\<forall>x y z. R x y \<longrightarrow> R x z \<longrightarrow> y = z)" | |
| 121 | ||
| 122 | definition bi_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | |
| 123 | where "bi_total R \<longleftrightarrow> (\<forall>x. \<exists>y. R x y) \<and> (\<forall>y. \<exists>x. R x y)" | |
| 124 | ||
| 125 | definition bi_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | |
| 126 | where "bi_unique R \<longleftrightarrow> | |
| 127 | (\<forall>x y z. R x y \<longrightarrow> R x z \<longrightarrow> y = z) \<and> | |
| 128 | (\<forall>x y z. R x z \<longrightarrow> R y z \<longrightarrow> x = y)" | |
| 129 | ||
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 130 | 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: 
56085diff
changeset | 131 | 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: 
56085diff
changeset | 132 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 133 | 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: 
56085diff
changeset | 134 | 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: 
56085diff
changeset | 135 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 136 | lemma left_totalI: | 
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 137 | "(\<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: 
56085diff
changeset | 138 | 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: 
56085diff
changeset | 139 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 140 | lemma left_totalE: | 
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 141 | 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: 
56085diff
changeset | 142 | 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: 
56085diff
changeset | 143 | 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: 
56085diff
changeset | 144 | |
| 53927 | 145 | lemma bi_uniqueDr: "\<lbrakk> bi_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z" | 
| 146 | by(simp add: bi_unique_def) | |
| 147 | ||
| 148 | lemma bi_uniqueDl: "\<lbrakk> bi_unique A; A x y; A z y \<rbrakk> \<Longrightarrow> x = z" | |
| 149 | by(simp add: bi_unique_def) | |
| 150 | ||
| 151 | lemma right_uniqueI: "(\<And>x y z. \<lbrakk> A x y; A x z \<rbrakk> \<Longrightarrow> y = z) \<Longrightarrow> right_unique A" | |
| 56085 | 152 | unfolding right_unique_def by fast | 
| 53927 | 153 | |
| 154 | lemma right_uniqueD: "\<lbrakk> right_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z" | |
| 56085 | 155 | unfolding right_unique_def by fast | 
| 53927 | 156 | |
| 59514 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 157 | 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: 
59276diff
changeset | 158 | by(simp add: right_total_def) | 
| 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 159 | |
| 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 160 | lemma right_totalE: | 
| 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 161 | assumes "right_total A" | 
| 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 162 | obtains x where "A x y" | 
| 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 163 | using assms by(auto simp add: right_total_def) | 
| 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 164 | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 165 | lemma right_total_alt_def2: | 
| 47325 | 166 | "right_total R \<longleftrightarrow> ((R ===> op \<longrightarrow>) ===> op \<longrightarrow>) All All" | 
| 55945 | 167 | unfolding right_total_def rel_fun_def | 
| 47325 | 168 | apply (rule iffI, fast) | 
| 169 | apply (rule allI) | |
| 170 | apply (drule_tac x="\<lambda>x. True" in spec) | |
| 171 | apply (drule_tac x="\<lambda>y. \<exists>x. R x y" in spec) | |
| 172 | apply fast | |
| 173 | done | |
| 174 | ||
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 175 | lemma right_unique_alt_def2: | 
| 47325 | 176 | "right_unique R \<longleftrightarrow> (R ===> R ===> op \<longrightarrow>) (op =) (op =)" | 
| 55945 | 177 | unfolding right_unique_def rel_fun_def by auto | 
| 47325 | 178 | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 179 | lemma bi_total_alt_def2: | 
| 47325 | 180 | "bi_total R \<longleftrightarrow> ((R ===> op =) ===> op =) All All" | 
| 55945 | 181 | unfolding bi_total_def rel_fun_def | 
| 47325 | 182 | apply (rule iffI, fast) | 
| 183 | apply safe | |
| 184 | apply (drule_tac x="\<lambda>x. \<exists>y. R x y" in spec) | |
| 185 | apply (drule_tac x="\<lambda>y. True" in spec) | |
| 186 | apply fast | |
| 187 | apply (drule_tac x="\<lambda>x. True" in spec) | |
| 188 | apply (drule_tac x="\<lambda>y. \<exists>x. R x y" in spec) | |
| 189 | apply fast | |
| 190 | done | |
| 191 | ||
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 192 | lemma bi_unique_alt_def2: | 
| 47325 | 193 | "bi_unique R \<longleftrightarrow> (R ===> R ===> op =) (op =) (op =)" | 
| 55945 | 194 | unfolding bi_unique_def rel_fun_def by auto | 
| 47325 | 195 | |
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 196 | lemma [simp]: | 
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 197 | 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: 
56085diff
changeset | 198 | 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: 
56085diff
changeset | 199 | 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: 
56085diff
changeset | 200 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 201 | lemma [simp]: | 
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 202 | 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: 
56085diff
changeset | 203 | 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: 
56085diff
changeset | 204 | 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: 
56085diff
changeset | 205 | |
| 53944 | 206 | lemma bi_unique_conversep [simp]: "bi_unique R\<inverse>\<inverse> = bi_unique R" | 
| 207 | by(auto simp add: bi_unique_def) | |
| 208 | ||
| 209 | lemma bi_total_conversep [simp]: "bi_total R\<inverse>\<inverse> = bi_total R" | |
| 210 | by(auto simp add: bi_total_def) | |
| 211 | ||
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 212 | lemma right_unique_alt_def: "right_unique R = (conversep R OO R \<le> op=)" unfolding right_unique_def by blast | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 213 | lemma left_unique_alt_def: "left_unique R = (R OO (conversep R) \<le> op=)" unfolding left_unique_def by blast | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 214 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 215 | lemma right_total_alt_def: "right_total R = (conversep R OO R \<ge> op=)" unfolding right_total_def by blast | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 216 | lemma left_total_alt_def: "left_total R = (R OO conversep R \<ge> op=)" unfolding left_total_def by blast | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 217 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 218 | 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: 
56085diff
changeset | 219 | 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: 
56085diff
changeset | 220 | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 221 | 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: 
56085diff
changeset | 222 | 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: 
56085diff
changeset | 223 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 224 | 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: 
56520diff
changeset | 225 | 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: 
56085diff
changeset | 226 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 227 | 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: 
56520diff
changeset | 228 | unfolding bi_unique_alt_def .. | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 229 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 230 | end | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 231 | |
| 60758 | 232 | subsection \<open>Equality restricted by a predicate\<close> | 
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 233 | |
| 58182 | 234 | definition eq_onp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 235 | where "eq_onp R = (\<lambda>x y. R x \<and> x = y)" | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 236 | |
| 58182 | 237 | lemma eq_onp_Grp: "eq_onp P = BNF_Def.Grp (Collect P) id" | 
| 238 | unfolding eq_onp_def Grp_def by auto | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 239 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 240 | lemma eq_onp_to_eq: | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 241 | assumes "eq_onp P x y" | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 242 | shows "x = y" | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 243 | using assms by (simp add: eq_onp_def) | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 244 | |
| 58182 | 245 | lemma eq_onp_top_eq_eq: "eq_onp top = op=" | 
| 56677 
660ffb526069
predicator simplification rules: support also partially specialized types e.g. 'a * nat
 kuncar parents: 
56654diff
changeset | 246 | by (simp add: eq_onp_def) | 
| 
660ffb526069
predicator simplification rules: support also partially specialized types e.g. 'a * nat
 kuncar parents: 
56654diff
changeset | 247 | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 248 | lemma eq_onp_same_args: | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 249 | shows "eq_onp P x x = P x" | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 250 | using assms by (auto simp add: eq_onp_def) | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 251 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 252 | lemma Ball_Collect: "Ball A P = (A \<subseteq> (Collect P))" | 
| 57260 
8747af0d1012
fixed postprocessing of 'coinduct' formula to obtain right property format (without needless hypotheses)
 blanchet parents: 
56677diff
changeset | 253 | by auto | 
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 254 | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 255 | ML_file "Tools/Transfer/transfer.ML" | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 256 | declare refl [transfer_rule] | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 257 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 258 | hide_const (open) Rel | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 259 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 260 | context | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 261 | begin | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 262 | interpretation lifting_syntax . | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 263 | |
| 60758 | 264 | text \<open>Handling of domains\<close> | 
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 265 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 266 | 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: 
56520diff
changeset | 267 | by auto | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 268 | |
| 58386 | 269 | lemma Domainp_refl[transfer_domain_rule]: | 
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 270 | "Domainp T = Domainp T" .. | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 271 | |
| 60229 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
59523diff
changeset | 272 | lemma Domain_eq_top: "Domainp op= = top" by auto | 
| 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
59523diff
changeset | 273 | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 274 | lemma Domainp_prod_fun_eq[relator_domain]: | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 275 | "Domainp (op= ===> T) = (\<lambda>f. \<forall>x. (Domainp T) (f x))" | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 276 | by (auto intro: choice simp: Domainp_iff rel_fun_def fun_eq_iff) | 
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
changeset | 299 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
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: 
56085diff
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: 
56085diff
changeset | 302 | |
| 47325 | 303 | |
| 60758 | 304 | subsection \<open>Properties of relators\<close> | 
| 47325 | 305 | |
| 58182 | 306 | lemma left_total_eq[transfer_rule]: "left_total op=" | 
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
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: 
56085diff
changeset | 308 | |
| 58182 | 309 | lemma left_unique_eq[transfer_rule]: "left_unique op=" | 
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
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: 
56085diff
changeset | 311 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 312 | lemma right_total_eq [transfer_rule]: "right_total op=" | 
| 47325 | 313 | unfolding right_total_def by simp | 
| 314 | ||
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 315 | lemma right_unique_eq [transfer_rule]: "right_unique op=" | 
| 47325 | 316 | unfolding right_unique_def by simp | 
| 317 | ||
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 318 | lemma bi_total_eq[transfer_rule]: "bi_total (op =)" | 
| 47325 | 319 | unfolding bi_total_def by simp | 
| 320 | ||
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 321 | lemma bi_unique_eq[transfer_rule]: "bi_unique (op =)" | 
| 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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
changeset | 336 | done | 
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 337 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
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: 
56520diff
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: 
56085diff
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: 
56085diff
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: 
56520diff
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: 
56085diff
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: 
56524diff
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: 
56524diff
changeset | 373 | |
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59141diff
changeset | 374 | lemma if_conn: | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59141diff
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: 
59141diff
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: 
59141diff
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: 
59141diff
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: 
59141diff
changeset | 379 | by auto | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59141diff
changeset | 380 | |
| 58182 | 381 | ML_file "Tools/Transfer/transfer_bnf.ML" | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59141diff
changeset | 382 | ML_file "Tools/BNF/bnf_fp_rec_sugar_transfer.ML" | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59141diff
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: 
56524diff
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: 
56524diff
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: 
56524diff
changeset | 386 | |
| 60758 | 387 | subsection \<open>Transfer rules\<close> | 
| 47325 | 388 | |
| 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: 
56524diff
changeset | 389 | context | 
| 
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
 kuncar parents: 
56524diff
changeset | 390 | 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: 
56524diff
changeset | 391 | interpretation lifting_syntax . | 
| 
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
 kuncar parents: 
56524diff
changeset | 392 | |
| 53952 | 393 | lemma Domainp_forall_transfer [transfer_rule]: | 
| 394 | assumes "right_total A" | |
| 395 | shows "((A ===> op =) ===> op =) | |
| 396 | (transfer_bforall (Domainp A)) transfer_forall" | |
| 397 | using assms unfolding right_total_def | |
| 55945 | 398 | unfolding transfer_forall_def transfer_bforall_def rel_fun_def Domainp_iff | 
| 56085 | 399 | by fast | 
| 53952 | 400 | |
| 60758 | 401 | text \<open>Transfer rules using implication instead of equality on booleans.\<close> | 
| 47684 | 402 | |
| 52354 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 403 | lemma transfer_forall_transfer [transfer_rule]: | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 404 | "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: 
51956diff
changeset | 405 | "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: 
51956diff
changeset | 406 | "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: 
51956diff
changeset | 407 | "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: 
51956diff
changeset | 408 | "bi_total A \<Longrightarrow> ((A ===> rev_implies) ===> rev_implies) transfer_forall transfer_forall" | 
| 55945 | 409 | unfolding transfer_forall_def rev_implies_def rel_fun_def right_total_def bi_total_def | 
| 56085 | 410 | by fast+ | 
| 52354 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 411 | |
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 412 | lemma transfer_implies_transfer [transfer_rule]: | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 413 | "(op = ===> op = ===> op = ) transfer_implies transfer_implies" | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 414 | "(rev_implies ===> implies ===> implies ) transfer_implies transfer_implies" | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 415 | "(rev_implies ===> op = ===> implies ) transfer_implies transfer_implies" | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 416 | "(op = ===> implies ===> implies ) transfer_implies transfer_implies" | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 417 | "(op = ===> op = ===> implies ) transfer_implies transfer_implies" | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 418 | "(implies ===> rev_implies ===> rev_implies) transfer_implies transfer_implies" | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 419 | "(implies ===> op = ===> rev_implies) transfer_implies transfer_implies" | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 420 | "(op = ===> rev_implies ===> rev_implies) transfer_implies transfer_implies" | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 421 | "(op = ===> op = ===> rev_implies) transfer_implies transfer_implies" | 
| 55945 | 422 | 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: 
51956diff
changeset | 423 | |
| 47684 | 424 | lemma eq_imp_transfer [transfer_rule]: | 
| 425 | "right_unique A \<Longrightarrow> (A ===> A ===> op \<longrightarrow>) (op =) (op =)" | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 426 | unfolding right_unique_alt_def2 . | 
| 47684 | 427 | |
| 60758 | 428 | 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: 
56085diff
changeset | 429 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 430 | 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: 
56085diff
changeset | 431 | 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: 
56085diff
changeset | 432 | 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: 
56085diff
changeset | 433 | assumes "bi_unique A" | 
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 434 | shows "((A ===> B ===> op=) ===> implies) left_unique left_unique" | 
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 435 | 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: 
56085diff
changeset | 436 | by metis | 
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 437 | |
| 47636 | 438 | lemma eq_transfer [transfer_rule]: | 
| 47325 | 439 | assumes "bi_unique A" | 
| 440 | shows "(A ===> A ===> op =) (op =) (op =)" | |
| 55945 | 441 | using assms unfolding bi_unique_def rel_fun_def by auto | 
| 47325 | 442 | |
| 51956 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 kuncar parents: 
51955diff
changeset | 443 | 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: 
51955diff
changeset | 444 | 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: 
51955diff
changeset | 445 | shows "((A ===> op=) ===> op=) (Bex (Collect (Domainp A))) Ex" | 
| 55945 | 446 | using assms unfolding right_total_def Bex_def rel_fun_def Domainp_iff[abs_def] | 
| 56085 | 447 | by fast | 
| 51956 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 kuncar parents: 
51955diff
changeset | 448 | |
| 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 kuncar parents: 
51955diff
changeset | 449 | 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: 
51955diff
changeset | 450 | 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: 
51955diff
changeset | 451 | shows "((A ===> op =) ===> op =) (Ball (Collect (Domainp A))) All" | 
| 55945 | 452 | using assms unfolding right_total_def Ball_def rel_fun_def Domainp_iff[abs_def] | 
| 56085 | 453 | by fast | 
| 51956 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 kuncar parents: 
51955diff
changeset | 454 | |
| 47636 | 455 | lemma All_transfer [transfer_rule]: | 
| 47325 | 456 | assumes "bi_total A" | 
| 457 | shows "((A ===> op =) ===> op =) All All" | |
| 55945 | 458 | using assms unfolding bi_total_def rel_fun_def by fast | 
| 47325 | 459 | |
| 47636 | 460 | lemma Ex_transfer [transfer_rule]: | 
| 47325 | 461 | assumes "bi_total A" | 
| 462 | shows "((A ===> op =) ===> op =) Ex Ex" | |
| 55945 | 463 | using assms unfolding bi_total_def rel_fun_def by fast | 
| 47325 | 464 | |
| 59515 | 465 | lemma Ex1_parametric [transfer_rule]: | 
| 466 | assumes [transfer_rule]: "bi_unique A" "bi_total A" | |
| 467 | shows "((A ===> op =) ===> op =) Ex1 Ex1" | |
| 468 | unfolding Ex1_def[abs_def] by transfer_prover | |
| 469 | ||
| 58448 | 470 | declare If_transfer [transfer_rule] | 
| 47325 | 471 | |
| 47636 | 472 | lemma Let_transfer [transfer_rule]: "(A ===> (A ===> B) ===> B) Let Let" | 
| 55945 | 473 | unfolding rel_fun_def by simp | 
| 47612 | 474 | |
| 58916 | 475 | declare id_transfer [transfer_rule] | 
| 47625 | 476 | |
| 58444 | 477 | declare comp_transfer [transfer_rule] | 
| 47325 | 478 | |
| 58916 | 479 | lemma curry_transfer [transfer_rule]: | 
| 480 | "((rel_prod A B ===> C) ===> A ===> B ===> C) curry curry" | |
| 481 | unfolding curry_def by transfer_prover | |
| 482 | ||
| 47636 | 483 | lemma fun_upd_transfer [transfer_rule]: | 
| 47325 | 484 | assumes [transfer_rule]: "bi_unique A" | 
| 485 | shows "((A ===> B) ===> A ===> B ===> A ===> B) fun_upd fun_upd" | |
| 47635 
ebb79474262c
rename 'correspondence' method to 'transfer_prover'
 huffman parents: 
47627diff
changeset | 486 | unfolding fun_upd_def [abs_def] by transfer_prover | 
| 47325 | 487 | |
| 55415 | 488 | lemma case_nat_transfer [transfer_rule]: | 
| 489 | "(A ===> (op = ===> A) ===> op = ===> A) case_nat case_nat" | |
| 55945 | 490 | 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: 
47625diff
changeset | 491 | |
| 55415 | 492 | lemma rec_nat_transfer [transfer_rule]: | 
| 493 | "(A ===> (op = ===> A ===> A) ===> op = ===> A) rec_nat rec_nat" | |
| 55945 | 494 | unfolding rel_fun_def by (clarsimp, rename_tac n, induct_tac n, simp_all) | 
| 47924 | 495 | |
| 496 | lemma funpow_transfer [transfer_rule]: | |
| 497 | "(op = ===> (A ===> A) ===> (A ===> A)) compow compow" | |
| 498 | unfolding funpow_def by transfer_prover | |
| 499 | ||
| 53952 | 500 | lemma mono_transfer[transfer_rule]: | 
| 501 | assumes [transfer_rule]: "bi_total A" | |
| 502 | assumes [transfer_rule]: "(A ===> A ===> op=) op\<le> op\<le>" | |
| 503 | assumes [transfer_rule]: "(B ===> B ===> op=) op\<le> op\<le>" | |
| 504 | shows "((A ===> B) ===> op=) mono mono" | |
| 505 | unfolding mono_def[abs_def] by transfer_prover | |
| 506 | ||
| 58182 | 507 | lemma right_total_relcompp_transfer[transfer_rule]: | 
| 53952 | 508 | assumes [transfer_rule]: "right_total B" | 
| 58182 | 509 | shows "((A ===> B ===> op=) ===> (B ===> C ===> op=) ===> A ===> C ===> op=) | 
| 53952 | 510 | (\<lambda>R S x z. \<exists>y\<in>Collect (Domainp B). R x y \<and> S y z) op OO" | 
| 511 | unfolding OO_def[abs_def] by transfer_prover | |
| 512 | ||
| 58182 | 513 | lemma relcompp_transfer[transfer_rule]: | 
| 53952 | 514 | assumes [transfer_rule]: "bi_total B" | 
| 515 | shows "((A ===> B ===> op=) ===> (B ===> C ===> op=) ===> A ===> C ===> op=) op OO op OO" | |
| 516 | unfolding OO_def[abs_def] by transfer_prover | |
| 47627 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 huffman parents: 
47625diff
changeset | 517 | |
| 53952 | 518 | lemma right_total_Domainp_transfer[transfer_rule]: | 
| 519 | assumes [transfer_rule]: "right_total B" | |
| 520 | shows "((A ===> B ===> op=) ===> A ===> op=) (\<lambda>T x. \<exists>y\<in>Collect(Domainp B). T x y) Domainp" | |
| 521 | apply(subst(2) Domainp_iff[abs_def]) by transfer_prover | |
| 522 | ||
| 523 | lemma Domainp_transfer[transfer_rule]: | |
| 524 | assumes [transfer_rule]: "bi_total B" | |
| 525 | shows "((A ===> B ===> op=) ===> A ===> op=) Domainp Domainp" | |
| 526 | unfolding Domainp_iff[abs_def] by transfer_prover | |
| 527 | ||
| 58182 | 528 | lemma reflp_transfer[transfer_rule]: | 
| 53952 | 529 | "bi_total A \<Longrightarrow> ((A ===> A ===> op=) ===> op=) reflp reflp" | 
| 530 | "right_total A \<Longrightarrow> ((A ===> A ===> implies) ===> implies) reflp reflp" | |
| 531 | "right_total A \<Longrightarrow> ((A ===> A ===> op=) ===> implies) reflp reflp" | |
| 532 | "bi_total A \<Longrightarrow> ((A ===> A ===> rev_implies) ===> rev_implies) reflp reflp" | |
| 533 | "bi_total A \<Longrightarrow> ((A ===> A ===> op=) ===> rev_implies) reflp reflp" | |
| 58182 | 534 | using assms unfolding reflp_def[abs_def] rev_implies_def bi_total_def right_total_def rel_fun_def | 
| 53952 | 535 | by fast+ | 
| 536 | ||
| 537 | lemma right_unique_transfer [transfer_rule]: | |
| 59523 | 538 | "\<lbrakk> right_total A; right_total B; bi_unique B \<rbrakk> | 
| 539 | \<Longrightarrow> ((A ===> B ===> op=) ===> implies) right_unique right_unique" | |
| 540 | unfolding right_unique_def[abs_def] right_total_def bi_unique_def rel_fun_def | |
| 53952 | 541 | by metis | 
| 47325 | 542 | |
| 59523 | 543 | lemma left_total_parametric [transfer_rule]: | 
| 544 | assumes [transfer_rule]: "bi_total A" "bi_total B" | |
| 545 | shows "((A ===> B ===> op =) ===> op =) left_total left_total" | |
| 546 | unfolding left_total_def[abs_def] by transfer_prover | |
| 547 | ||
| 548 | lemma right_total_parametric [transfer_rule]: | |
| 549 | assumes [transfer_rule]: "bi_total A" "bi_total B" | |
| 550 | shows "((A ===> B ===> op =) ===> op =) right_total right_total" | |
| 551 | unfolding right_total_def[abs_def] by transfer_prover | |
| 552 | ||
| 553 | lemma left_unique_parametric [transfer_rule]: | |
| 554 | assumes [transfer_rule]: "bi_unique A" "bi_total A" "bi_total B" | |
| 555 | shows "((A ===> B ===> op =) ===> op =) left_unique left_unique" | |
| 556 | unfolding left_unique_def[abs_def] by transfer_prover | |
| 557 | ||
| 558 | lemma prod_pred_parametric [transfer_rule]: | |
| 559 | "((A ===> op =) ===> (B ===> op =) ===> rel_prod A B ===> op =) pred_prod pred_prod" | |
| 560 | unfolding pred_prod_def[abs_def] Basic_BNFs.fsts_def Basic_BNFs.snds_def fstsp.simps sndsp.simps | |
| 561 | by simp transfer_prover | |
| 562 | ||
| 563 | lemma apfst_parametric [transfer_rule]: | |
| 564 | "((A ===> B) ===> rel_prod A C ===> rel_prod B C) apfst apfst" | |
| 565 | unfolding apfst_def[abs_def] by transfer_prover | |
| 566 | ||
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 567 | lemma rel_fun_eq_eq_onp: "(op= ===> eq_onp P) = eq_onp (\<lambda>f. \<forall>x. P(f x))" | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 568 | unfolding eq_onp_def rel_fun_def by auto | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 569 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 570 | lemma rel_fun_eq_onp_rel: | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 571 | 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: 
56520diff
changeset | 572 | by (auto simp add: eq_onp_def rel_fun_def) | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 573 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 574 | lemma eq_onp_transfer [transfer_rule]: | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 575 | assumes [transfer_rule]: "bi_unique A" | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 576 | shows "((A ===> op=) ===> A ===> A ===> op=) eq_onp eq_onp" | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 577 | unfolding eq_onp_def[abs_def] by transfer_prover | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 578 | |
| 57599 | 579 | lemma rtranclp_parametric [transfer_rule]: | 
| 580 | assumes "bi_unique A" "bi_total A" | |
| 581 | shows "((A ===> A ===> op =) ===> A ===> A ===> op =) rtranclp rtranclp" | |
| 582 | proof(rule rel_funI iffI)+ | |
| 583 | fix R :: "'a \<Rightarrow> 'a \<Rightarrow> bool" and R' x y x' y' | |
| 584 | assume R: "(A ===> A ===> op =) R R'" and "A x x'" | |
| 585 |   {
 | |
| 586 | assume "R\<^sup>*\<^sup>* x y" "A y y'" | |
| 587 | thus "R'\<^sup>*\<^sup>* x' y'" | |
| 588 | proof(induction arbitrary: y') | |
| 589 | case base | |
| 60758 | 590 | with \<open>bi_unique A\<close> \<open>A x x'\<close> have "x' = y'" by(rule bi_uniqueDr) | 
| 57599 | 591 | thus ?case by simp | 
| 592 | next | |
| 593 | case (step y z z') | |
| 60758 | 594 | from \<open>bi_total A\<close> obtain y' where "A y y'" unfolding bi_total_def by blast | 
| 57599 | 595 | hence "R'\<^sup>*\<^sup>* x' y'" by(rule step.IH) | 
| 60758 | 596 | moreover from R \<open>A y y'\<close> \<open>A z z'\<close> \<open>R y z\<close> | 
| 57599 | 597 | have "R' y' z'" by(auto dest: rel_funD) | 
| 598 | ultimately show ?case .. | |
| 599 | qed | |
| 600 | next | |
| 601 | assume "R'\<^sup>*\<^sup>* x' y'" "A y y'" | |
| 602 | thus "R\<^sup>*\<^sup>* x y" | |
| 603 | proof(induction arbitrary: y) | |
| 604 | case base | |
| 60758 | 605 | with \<open>bi_unique A\<close> \<open>A x x'\<close> have "x = y" by(rule bi_uniqueDl) | 
| 57599 | 606 | thus ?case by simp | 
| 607 | next | |
| 608 | case (step y' z' z) | |
| 60758 | 609 | from \<open>bi_total A\<close> obtain y where "A y y'" unfolding bi_total_def by blast | 
| 57599 | 610 | hence "R\<^sup>*\<^sup>* x y" by(rule step.IH) | 
| 60758 | 611 | moreover from R \<open>A y y'\<close> \<open>A z z'\<close> \<open>R' y' z'\<close> | 
| 57599 | 612 | have "R y z" by(auto dest: rel_funD) | 
| 613 | ultimately show ?case .. | |
| 614 | qed | |
| 615 | } | |
| 616 | qed | |
| 617 | ||
| 59523 | 618 | lemma right_unique_parametric [transfer_rule]: | 
| 619 | assumes [transfer_rule]: "bi_total A" "bi_unique B" "bi_total B" | |
| 620 | shows "((A ===> B ===> op =) ===> op =) right_unique right_unique" | |
| 621 | unfolding right_unique_def[abs_def] by transfer_prover | |
| 622 | ||
| 61630 | 623 | lemma map_fun_parametric [transfer_rule]: | 
| 624 | "((A ===> B) ===> (C ===> D) ===> (B ===> C) ===> A ===> D) map_fun map_fun" | |
| 625 | unfolding map_fun_def[abs_def] by transfer_prover | |
| 626 | ||
| 47325 | 627 | end | 
| 53011 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 kuncar parents: 
52358diff
changeset | 628 | |
| 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 kuncar parents: 
52358diff
changeset | 629 | end |