| author | wenzelm | 
| Fri, 08 Dec 2023 15:37:46 +0100 | |
| changeset 79207 | f991d3003ec8 | 
| parent 75669 | 43f5dfb7fa35 | 
| child 80932 | 261cd8722677 | 
| 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 | |
| 63343 | 14 | bundle lifting_syntax | 
| 53011 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 kuncar parents: 
52358diff
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: 
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 | |
| 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: 
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 | |
| 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: 
47924diff
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: 
47924diff
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: 
48891diff
changeset | 49 | |
| 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 huffman parents: 
48891diff
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: 
48891diff
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: 
51112diff
changeset | 54 | unfolding is_equality_def by simp | 
| 
8739f8abbecb
fixing transfer tactic - unfold fully identity relation by using relator_eq
 kuncar parents: 
51112diff
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: 
51956diff
changeset | 57 | |
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 58 | definition rev_implies where | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 59 | "rev_implies x y \<longleftrightarrow> (y \<longrightarrow> x)" | 
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
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: 
47325diff
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: 
47325diff
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: 
47325diff
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: 
47325diff
changeset | 78 | lemma transfer_bforall_unfold: | 
| 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 huffman parents: 
47325diff
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: 
47325diff
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: 
47325diff
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: 
47627diff
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: 
47612diff
changeset | 89 | by simp | 
| 
1568dadd598a
make correspondence tactic more robust by replacing lhs with schematic variable before applying intro rules
 huffman parents: 
47612diff
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: 
52354diff
changeset | 92 | unfolding Rel_def by simp | 
| 
f4c4bcb0d564
implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
 huffman parents: 
52354diff
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: 
47684diff
changeset | 97 | lemma Rel_app: | 
| 47523 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 huffman parents: 
47503diff
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: 
47684diff
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: 
47503diff
changeset | 101 | |
| 47789 
71a526ee569a
implement transfer tactic with more scalable forward proof methods
 huffman parents: 
47684diff
changeset | 102 | lemma Rel_abs: | 
| 47523 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 huffman parents: 
47503diff
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: 
47684diff
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: 
47503diff
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: 
56085diff
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: 
56085diff
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: 
56085diff
changeset | 111 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
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: 
56085diff
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: 
56085diff
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 | ||
| 71827 | 129 | lemma left_unique_iff: "left_unique R \<longleftrightarrow> (\<forall>z. \<exists>\<^sub>\<le>\<^sub>1x. R x z)" | 
| 130 | unfolding Uniq_def left_unique_def by force | |
| 131 | ||
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 132 | 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 | 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: 
56085diff
changeset | 134 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 135 | 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 | 136 | 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 | 137 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 138 | 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 | 139 | "(\<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 | 140 | 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 | 141 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 142 | 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 | 143 | 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 | 144 | 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 | 145 | 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 | 146 | |
| 53927 | 147 | lemma bi_uniqueDr: "\<lbrakk> bi_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z" | 
| 71827 | 148 | by(simp add: bi_unique_def) | 
| 53927 | 149 | |
| 150 | lemma bi_uniqueDl: "\<lbrakk> bi_unique A; A x y; A z y \<rbrakk> \<Longrightarrow> x = z" | |
| 71827 | 151 | by(simp add: bi_unique_def) | 
| 152 | ||
| 153 | lemma bi_unique_iff: "bi_unique R \<longleftrightarrow> (\<forall>z. \<exists>\<^sub>\<le>\<^sub>1x. R x z) \<and> (\<forall>z. \<exists>\<^sub>\<le>\<^sub>1x. R z x)" | |
| 154 | unfolding Uniq_def bi_unique_def by force | |
| 155 | ||
| 156 | lemma right_unique_iff: "right_unique R \<longleftrightarrow> (\<forall>z. \<exists>\<^sub>\<le>\<^sub>1x. R z x)" | |
| 157 | unfolding Uniq_def right_unique_def by force | |
| 53927 | 158 | |
| 159 | lemma right_uniqueI: "(\<And>x y z. \<lbrakk> A x y; A x z \<rbrakk> \<Longrightarrow> y = z) \<Longrightarrow> right_unique A" | |
| 56085 | 160 | unfolding right_unique_def by fast | 
| 53927 | 161 | |
| 162 | lemma right_uniqueD: "\<lbrakk> right_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z" | |
| 56085 | 163 | unfolding right_unique_def by fast | 
| 53927 | 164 | |
| 59514 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 165 | 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 | 166 | by(simp add: right_total_def) | 
| 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 167 | |
| 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 168 | lemma right_totalE: | 
| 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 169 | assumes "right_total A" | 
| 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 170 | obtains x where "A x y" | 
| 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 171 | using assms by(auto simp add: right_total_def) | 
| 
509caf5edfa6
add intro and elim rules for right_total
 Andreas Lochbihler parents: 
59276diff
changeset | 172 | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 173 | lemma right_total_alt_def2: | 
| 71697 | 174 | "right_total R \<longleftrightarrow> ((R ===> (\<longrightarrow>)) ===> (\<longrightarrow>)) All All" (is "?lhs = ?rhs") | 
| 175 | proof | |
| 176 | assume ?lhs then show ?rhs | |
| 177 | unfolding right_total_def rel_fun_def by blast | |
| 178 | next | |
| 179 | assume \<section>: ?rhs | |
| 180 | show ?lhs | |
| 181 | using \<section> [unfolded rel_fun_def, rule_format, of "\<lambda>x. True" "\<lambda>y. \<exists>x. R x y"] | |
| 182 | unfolding right_total_def by blast | |
| 183 | qed | |
| 47325 | 184 | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 185 | lemma right_unique_alt_def2: | 
| 67399 | 186 | "right_unique R \<longleftrightarrow> (R ===> R ===> (\<longrightarrow>)) (=) (=)" | 
| 55945 | 187 | unfolding right_unique_def rel_fun_def by auto | 
| 47325 | 188 | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 189 | lemma bi_total_alt_def2: | 
| 71697 | 190 | "bi_total R \<longleftrightarrow> ((R ===> (=)) ===> (=)) All All" (is "?lhs = ?rhs") | 
| 191 | proof | |
| 192 | assume ?lhs then show ?rhs | |
| 193 | unfolding bi_total_def rel_fun_def by blast | |
| 194 | next | |
| 195 | assume \<section>: ?rhs | |
| 196 | show ?lhs | |
| 197 | using \<section> [unfolded rel_fun_def, rule_format, of "\<lambda>x. \<exists>y. R x y" "\<lambda>y. True"] | |
| 198 | using \<section> [unfolded rel_fun_def, rule_format, of "\<lambda>x. True" "\<lambda>y. \<exists>x. R x y"] | |
| 199 | by (auto simp: bi_total_def) | |
| 200 | qed | |
| 47325 | 201 | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 202 | lemma bi_unique_alt_def2: | 
| 67399 | 203 | "bi_unique R \<longleftrightarrow> (R ===> R ===> (=)) (=) (=)" | 
| 55945 | 204 | unfolding bi_unique_def rel_fun_def by auto | 
| 47325 | 205 | |
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 206 | 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 | 207 | shows left_unique_conversep: "left_unique A\<inverse>\<inverse> \<longleftrightarrow> right_unique A" | 
| 71697 | 208 | and right_unique_conversep: "right_unique A\<inverse>\<inverse> \<longleftrightarrow> left_unique A" | 
| 209 | by(auto simp add: left_unique_def right_unique_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 | 210 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 211 | 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 | 212 | shows left_total_conversep: "left_total A\<inverse>\<inverse> \<longleftrightarrow> right_total A" | 
| 71697 | 213 | and right_total_conversep: "right_total A\<inverse>\<inverse> \<longleftrightarrow> left_total A" | 
| 214 | by(simp_all add: left_total_def right_total_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 | 215 | |
| 53944 | 216 | lemma bi_unique_conversep [simp]: "bi_unique R\<inverse>\<inverse> = bi_unique R" | 
| 71697 | 217 | by(auto simp add: bi_unique_def) | 
| 53944 | 218 | |
| 219 | lemma bi_total_conversep [simp]: "bi_total R\<inverse>\<inverse> = bi_total R" | |
| 71697 | 220 | by(auto simp add: bi_total_def) | 
| 53944 | 221 | |
| 67399 | 222 | lemma right_unique_alt_def: "right_unique R = (conversep R OO R \<le> (=))" unfolding right_unique_def by blast | 
| 223 | 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: 
56520diff
changeset | 224 | |
| 67399 | 225 | lemma right_total_alt_def: "right_total R = (conversep R OO R \<ge> (=))" unfolding right_total_def by blast | 
| 226 | 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: 
56520diff
changeset | 227 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 228 | 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 | 229 | 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 | 230 | |
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 231 | 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 | 232 | 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 | 233 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 234 | 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 | 235 | 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 | 236 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 237 | 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 | 238 | unfolding bi_unique_alt_def .. | 
| 
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 | end | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 241 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 242 | |
| 70491 | 243 | lemma is_equality_lemma: "(\<And>R. is_equality R \<Longrightarrow> PROP (P R)) \<equiv> PROP (P (=))" | 
| 71697 | 244 | unfolding is_equality_def | 
| 245 | proof (rule equal_intr_rule) | |
| 246 | show "(\<And>R. R = (=) \<Longrightarrow> PROP P R) \<Longrightarrow> PROP P (=)" | |
| 247 | apply (drule meta_spec) | |
| 248 | apply (erule meta_mp [OF _ refl]) | |
| 249 | done | |
| 250 | qed simp | |
| 70491 | 251 | |
| 252 | lemma Domainp_lemma: "(\<And>R. Domainp T = R \<Longrightarrow> PROP (P R)) \<equiv> PROP (P (Domainp T))" | |
| 71697 | 253 | proof (rule equal_intr_rule) | 
| 254 | show "(\<And>R. Domainp T = R \<Longrightarrow> PROP P R) \<Longrightarrow> PROP P (Domainp T)" | |
| 255 | apply (drule meta_spec) | |
| 256 | apply (erule meta_mp [OF _ refl]) | |
| 257 | done | |
| 258 | qed simp | |
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 259 | |
| 69605 | 260 | ML_file \<open>Tools/Transfer/transfer.ML\<close> | 
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 261 | declare refl [transfer_rule] | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 262 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 263 | hide_const (open) Rel | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 264 | |
| 63343 | 265 | context includes lifting_syntax | 
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 266 | begin | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 267 | |
| 60758 | 268 | text \<open>Handling of domains\<close> | 
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 269 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 270 | 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 | 271 | by auto | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 272 | |
| 58386 | 273 | lemma Domainp_refl[transfer_domain_rule]: | 
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 274 | "Domainp T = Domainp T" .. | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 275 | |
| 67399 | 276 | 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: 
59523diff
changeset | 277 | |
| 64425 
b17acc1834e3
a more general relator domain rule for the function type
 kuncar parents: 
64014diff
changeset | 278 | lemma Domainp_pred_fun_eq[relator_domain]: | 
| 
b17acc1834e3
a more general relator domain rule for the function type
 kuncar parents: 
64014diff
changeset | 279 | assumes "left_unique T" | 
| 71697 | 280 | shows "Domainp (T ===> S) = pred_fun (Domainp T) (Domainp S)" (is "?lhs = ?rhs") | 
| 281 | proof (intro ext iffI) | |
| 282 | fix x | |
| 283 | assume "?lhs x" | |
| 284 | then show "?rhs x" | |
| 285 | using assms unfolding rel_fun_def pred_fun_def by blast | |
| 286 | next | |
| 287 | fix x | |
| 288 | assume "?rhs x" | |
| 289 | then have "\<exists>g. \<forall>y xa. T xa y \<longrightarrow> S (x xa) (g y)" | |
| 290 | using assms unfolding Domainp_iff left_unique_def pred_fun_def | |
| 291 | by (intro choice) blast | |
| 292 | then show "?lhs x" | |
| 293 | by blast | |
| 294 | qed | |
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 295 | |
| 60758 | 296 | text \<open>Properties are preserved by relation composition.\<close> | 
| 47660 | 297 | |
| 298 | lemma OO_def: "R OO S = (\<lambda>x z. \<exists>y. R x y \<and> S y z)" | |
| 299 | by auto | |
| 300 | ||
| 301 | lemma bi_total_OO: "\<lbrakk>bi_total A; bi_total B\<rbrakk> \<Longrightarrow> bi_total (A OO B)" | |
| 56085 | 302 | unfolding bi_total_def OO_def by fast | 
| 47660 | 303 | |
| 304 | lemma bi_unique_OO: "\<lbrakk>bi_unique A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A OO B)" | |
| 56085 | 305 | unfolding bi_unique_def OO_def by blast | 
| 47660 | 306 | |
| 307 | lemma right_total_OO: | |
| 308 | "\<lbrakk>right_total A; right_total B\<rbrakk> \<Longrightarrow> right_total (A OO B)" | |
| 56085 | 309 | unfolding right_total_def OO_def by fast | 
| 47660 | 310 | |
| 311 | lemma right_unique_OO: | |
| 312 | "\<lbrakk>right_unique A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A OO B)" | |
| 56085 | 313 | unfolding right_unique_def OO_def by fast | 
| 47660 | 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 left_total_OO: "left_total R \<Longrightarrow> left_total S \<Longrightarrow> left_total (R OO S)" | 
| 71697 | 316 | unfolding left_total_def OO_def by fast | 
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 317 | |
| 
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 left_unique_OO: "left_unique R \<Longrightarrow> left_unique S \<Longrightarrow> left_unique (R OO S)" | 
| 71697 | 319 | unfolding left_unique_def OO_def by blast | 
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 320 | |
| 47325 | 321 | |
| 60758 | 322 | subsection \<open>Properties of relators\<close> | 
| 47325 | 323 | |
| 67399 | 324 | 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: 
56085diff
changeset | 325 | 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 | 326 | |
| 67399 | 327 | 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: 
56085diff
changeset | 328 | 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 | 329 | |
| 67399 | 330 | lemma right_total_eq [transfer_rule]: "right_total (=)" | 
| 47325 | 331 | unfolding right_total_def by simp | 
| 332 | ||
| 67399 | 333 | lemma right_unique_eq [transfer_rule]: "right_unique (=)" | 
| 47325 | 334 | unfolding right_unique_def by simp | 
| 335 | ||
| 67399 | 336 | lemma bi_total_eq[transfer_rule]: "bi_total (=)" | 
| 47325 | 337 | unfolding bi_total_def by simp | 
| 338 | ||
| 67399 | 339 | lemma bi_unique_eq[transfer_rule]: "bi_unique (=)" | 
| 47325 | 340 | unfolding bi_unique_def by simp | 
| 341 | ||
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 342 | lemma left_total_fun[transfer_rule]: | 
| 71697 | 343 | assumes "left_unique A" "left_total B" | 
| 344 | shows "left_total (A ===> B)" | |
| 345 | unfolding left_total_def | |
| 346 | proof | |
| 347 | fix f | |
| 348 | show "Ex ((A ===> B) f)" | |
| 349 | unfolding rel_fun_def | |
| 350 | proof (intro exI strip) | |
| 351 | fix x y | |
| 352 | assume A: "A x y" | |
| 353 | have "(THE x. A x y) = x" | |
| 354 | using A assms by (simp add: left_unique_def the_equality) | |
| 355 | then show "B (f x) (SOME z. B (f (THE x. A x y)) z)" | |
| 356 | using assms by (force simp: left_total_def intro: someI_ex) | |
| 357 | qed | |
| 358 | qed | |
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 359 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 360 | 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 | 361 | "\<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 | 362 | 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 | 363 | 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 | 364 | |
| 47325 | 365 | lemma right_total_fun [transfer_rule]: | 
| 71697 | 366 | assumes "right_unique A" "right_total B" | 
| 367 | shows "right_total (A ===> B)" | |
| 368 | unfolding right_total_def | |
| 369 | proof | |
| 370 | fix g | |
| 371 | show "\<exists>x. (A ===> B) x g" | |
| 372 | unfolding rel_fun_def | |
| 373 | proof (intro exI strip) | |
| 374 | fix x y | |
| 375 | assume A: "A x y" | |
| 376 | have "(THE y. A x y) = y" | |
| 377 | using A assms by (simp add: right_unique_def the_equality) | |
| 378 | then show "B (SOME z. B z (g (THE y. A x y))) (g y)" | |
| 379 | using assms by (force simp: right_total_def intro: someI_ex) | |
| 380 | qed | |
| 381 | qed | |
| 47325 | 382 | |
| 383 | lemma right_unique_fun [transfer_rule]: | |
| 384 | "\<lbrakk>right_total A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A ===> B)" | |
| 55945 | 385 | unfolding right_total_def right_unique_def rel_fun_def | 
| 47325 | 386 | by (clarify, rule ext, fast) | 
| 387 | ||
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 388 | lemma bi_total_fun[transfer_rule]: | 
| 47325 | 389 | "\<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 | 390 | 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 | 391 | by (blast intro: right_total_fun left_total_fun) | 
| 47325 | 392 | |
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 393 | lemma bi_unique_fun[transfer_rule]: | 
| 47325 | 394 | "\<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 | 395 | 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 | 396 | by (blast intro: right_unique_fun left_unique_fun) | 
| 47325 | 397 | |
| 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 | 398 | 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 | 399 | |
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59141diff
changeset | 400 | lemma if_conn: | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59141diff
changeset | 401 | "(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 | 402 | "(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 | 403 | "(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 | 404 | "(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 | 405 | by auto | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59141diff
changeset | 406 | |
| 69605 | 407 | ML_file \<open>Tools/Transfer/transfer_bnf.ML\<close> | 
| 408 | 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: 
59141diff
changeset | 409 | |
| 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 | 410 | 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 | 411 | 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 | 412 | |
| 64425 
b17acc1834e3
a more general relator domain rule for the function type
 kuncar parents: 
64014diff
changeset | 413 | (* Delete the automated generated rule from the bnf command; | 
| 
b17acc1834e3
a more general relator domain rule for the function type
 kuncar parents: 
64014diff
changeset | 414 | 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: 
64014diff
changeset | 415 | declare fun.Domainp_rel[relator_domain del] | 
| 
b17acc1834e3
a more general relator domain rule for the function type
 kuncar parents: 
64014diff
changeset | 416 | |
| 60758 | 417 | subsection \<open>Transfer rules\<close> | 
| 47325 | 418 | |
| 63343 | 419 | 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: 
56524diff
changeset | 420 | 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 | 421 | |
| 53952 | 422 | lemma Domainp_forall_transfer [transfer_rule]: | 
| 423 | assumes "right_total A" | |
| 67399 | 424 | shows "((A ===> (=)) ===> (=)) | 
| 53952 | 425 | (transfer_bforall (Domainp A)) transfer_forall" | 
| 426 | using assms unfolding right_total_def | |
| 55945 | 427 | unfolding transfer_forall_def transfer_bforall_def rel_fun_def Domainp_iff | 
| 56085 | 428 | by fast | 
| 53952 | 429 | |
| 60758 | 430 | text \<open>Transfer rules using implication instead of equality on booleans.\<close> | 
| 47684 | 431 | |
| 52354 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 432 | lemma transfer_forall_transfer [transfer_rule]: | 
| 67399 | 433 | "bi_total A \<Longrightarrow> ((A ===> (=)) ===> (=)) transfer_forall transfer_forall" | 
| 434 | "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: 
51956diff
changeset | 435 | "right_total A \<Longrightarrow> ((A ===> implies) ===> implies) transfer_forall transfer_forall" | 
| 67399 | 436 | "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: 
51956diff
changeset | 437 | "bi_total A \<Longrightarrow> ((A ===> rev_implies) ===> rev_implies) transfer_forall transfer_forall" | 
| 55945 | 438 | unfolding transfer_forall_def rev_implies_def rel_fun_def right_total_def bi_total_def | 
| 56085 | 439 | by fast+ | 
| 52354 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 440 | |
| 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 441 | lemma transfer_implies_transfer [transfer_rule]: | 
| 67399 | 442 | "((=) ===> (=) ===> (=) ) transfer_implies transfer_implies" | 
| 52354 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 443 | "(rev_implies ===> implies ===> implies ) transfer_implies transfer_implies" | 
| 67399 | 444 | "(rev_implies ===> (=) ===> implies ) transfer_implies transfer_implies" | 
| 445 | "((=) ===> implies ===> implies ) transfer_implies transfer_implies" | |
| 446 | "((=) ===> (=) ===> implies ) transfer_implies transfer_implies" | |
| 52354 
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
 huffman parents: 
51956diff
changeset | 447 | "(implies ===> rev_implies ===> rev_implies) transfer_implies transfer_implies" | 
| 67399 | 448 | "(implies ===> (=) ===> rev_implies) transfer_implies transfer_implies" | 
| 449 | "((=) ===> rev_implies ===> rev_implies) transfer_implies transfer_implies" | |
| 450 | "((=) ===> (=) ===> rev_implies) transfer_implies transfer_implies" | |
| 55945 | 451 | 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 | 452 | |
| 47684 | 453 | lemma eq_imp_transfer [transfer_rule]: | 
| 67399 | 454 | "right_unique A \<Longrightarrow> (A ===> A ===> (\<longrightarrow>)) (=) (=)" | 
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 455 | unfolding right_unique_alt_def2 . | 
| 47684 | 456 | |
| 60758 | 457 | 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 | 458 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 459 | 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 | 460 | 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 | 461 | 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 | 462 | assumes "bi_unique A" | 
| 67399 | 463 | shows "((A ===> B ===> (=)) ===> implies) left_unique left_unique" | 
| 71697 | 464 | using assms unfolding left_unique_def right_total_def bi_unique_def rel_fun_def | 
| 465 | by metis | |
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56085diff
changeset | 466 | |
| 47636 | 467 | lemma eq_transfer [transfer_rule]: | 
| 47325 | 468 | assumes "bi_unique A" | 
| 67399 | 469 | shows "(A ===> A ===> (=)) (=) (=)" | 
| 55945 | 470 | using assms unfolding bi_unique_def rel_fun_def by auto | 
| 47325 | 471 | |
| 51956 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 kuncar parents: 
51955diff
changeset | 472 | 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 | 473 | assumes "right_total A" | 
| 67399 | 474 | shows "((A ===> (=)) ===> (=)) (Bex (Collect (Domainp A))) Ex" | 
| 71697 | 475 | using assms unfolding right_total_def Bex_def rel_fun_def Domainp_iff | 
| 476 | 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 | 477 | |
| 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 kuncar parents: 
51955diff
changeset | 478 | 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 | 479 | assumes "right_total A" | 
| 67399 | 480 | shows "((A ===> (=)) ===> (=)) (Ball (Collect (Domainp A))) All" | 
| 71697 | 481 | using assms unfolding right_total_def Ball_def rel_fun_def Domainp_iff | 
| 482 | 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 | 483 | |
| 70927 | 484 | context | 
| 485 | includes lifting_syntax | |
| 486 | begin | |
| 487 | ||
| 68521 | 488 | lemma right_total_fun_eq_transfer: | 
| 489 | assumes [transfer_rule]: "right_total A" "bi_unique B" | |
| 490 | shows "((A ===> B) ===> (A ===> B) ===> (=)) (\<lambda>f g. \<forall>x\<in>Collect(Domainp A). f x = g x) (=)" | |
| 491 | unfolding fun_eq_iff | |
| 492 | by transfer_prover | |
| 493 | ||
| 70927 | 494 | end | 
| 495 | ||
| 47636 | 496 | lemma All_transfer [transfer_rule]: | 
| 47325 | 497 | assumes "bi_total A" | 
| 67399 | 498 | shows "((A ===> (=)) ===> (=)) All All" | 
| 55945 | 499 | using assms unfolding bi_total_def rel_fun_def by fast | 
| 47325 | 500 | |
| 47636 | 501 | lemma Ex_transfer [transfer_rule]: | 
| 47325 | 502 | assumes "bi_total A" | 
| 67399 | 503 | shows "((A ===> (=)) ===> (=)) Ex Ex" | 
| 55945 | 504 | using assms unfolding bi_total_def rel_fun_def by fast | 
| 47325 | 505 | |
| 59515 | 506 | lemma Ex1_parametric [transfer_rule]: | 
| 507 | assumes [transfer_rule]: "bi_unique A" "bi_total A" | |
| 67399 | 508 | shows "((A ===> (=)) ===> (=)) Ex1 Ex1" | 
| 71697 | 509 | unfolding Ex1_def by transfer_prover | 
| 59515 | 510 | |
| 58448 | 511 | declare If_transfer [transfer_rule] | 
| 47325 | 512 | |
| 47636 | 513 | lemma Let_transfer [transfer_rule]: "(A ===> (A ===> B) ===> B) Let Let" | 
| 55945 | 514 | unfolding rel_fun_def by simp | 
| 47612 | 515 | |
| 58916 | 516 | declare id_transfer [transfer_rule] | 
| 47625 | 517 | |
| 58444 | 518 | declare comp_transfer [transfer_rule] | 
| 47325 | 519 | |
| 58916 | 520 | lemma curry_transfer [transfer_rule]: | 
| 521 | "((rel_prod A B ===> C) ===> A ===> B ===> C) curry curry" | |
| 522 | unfolding curry_def by transfer_prover | |
| 523 | ||
| 47636 | 524 | lemma fun_upd_transfer [transfer_rule]: | 
| 47325 | 525 | assumes [transfer_rule]: "bi_unique A" | 
| 526 | shows "((A ===> B) ===> A ===> B ===> A ===> B) fun_upd fun_upd" | |
| 71697 | 527 | unfolding fun_upd_def by transfer_prover | 
| 47325 | 528 | |
| 55415 | 529 | lemma case_nat_transfer [transfer_rule]: | 
| 67399 | 530 | "(A ===> ((=) ===> A) ===> (=) ===> A) case_nat case_nat" | 
| 55945 | 531 | 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 | 532 | |
| 55415 | 533 | lemma rec_nat_transfer [transfer_rule]: | 
| 67399 | 534 | "(A ===> ((=) ===> A ===> A) ===> (=) ===> A) rec_nat rec_nat" | 
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
71827diff
changeset | 535 | unfolding rel_fun_def | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
71827diff
changeset | 536 | apply safe | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
71827diff
changeset | 537 | subgoal for _ _ _ _ _ n | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
71827diff
changeset | 538 | by (induction n) simp_all | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
71827diff
changeset | 539 | done | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
71827diff
changeset | 540 | |
| 47924 | 541 | |
| 542 | lemma funpow_transfer [transfer_rule]: | |
| 67399 | 543 | "((=) ===> (A ===> A) ===> (A ===> A)) compow compow" | 
| 47924 | 544 | unfolding funpow_def by transfer_prover | 
| 545 | ||
| 53952 | 546 | lemma mono_transfer[transfer_rule]: | 
| 547 | assumes [transfer_rule]: "bi_total A" | |
| 67399 | 548 | assumes [transfer_rule]: "(A ===> A ===> (=)) (\<le>) (\<le>)" | 
| 549 | assumes [transfer_rule]: "(B ===> B ===> (=)) (\<le>) (\<le>)" | |
| 550 | shows "((A ===> B) ===> (=)) mono mono" | |
| 71697 | 551 | unfolding mono_def by transfer_prover | 
| 53952 | 552 | |
| 58182 | 553 | lemma right_total_relcompp_transfer[transfer_rule]: | 
| 53952 | 554 | assumes [transfer_rule]: "right_total B" | 
| 67399 | 555 | shows "((A ===> B ===> (=)) ===> (B ===> C ===> (=)) ===> A ===> C ===> (=)) | 
| 556 | (\<lambda>R S x z. \<exists>y\<in>Collect (Domainp B). R x y \<and> S y z) (OO)" | |
| 71697 | 557 | unfolding OO_def by transfer_prover | 
| 53952 | 558 | |
| 58182 | 559 | lemma relcompp_transfer[transfer_rule]: | 
| 53952 | 560 | assumes [transfer_rule]: "bi_total B" | 
| 67399 | 561 | shows "((A ===> B ===> (=)) ===> (B ===> C ===> (=)) ===> A ===> C ===> (=)) (OO) (OO)" | 
| 71697 | 562 | unfolding OO_def by transfer_prover | 
| 47627 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 huffman parents: 
47625diff
changeset | 563 | |
| 53952 | 564 | lemma right_total_Domainp_transfer[transfer_rule]: | 
| 565 | assumes [transfer_rule]: "right_total B" | |
| 67399 | 566 | shows "((A ===> B ===> (=)) ===> A ===> (=)) (\<lambda>T x. \<exists>y\<in>Collect(Domainp B). T x y) Domainp" | 
| 53952 | 567 | apply(subst(2) Domainp_iff[abs_def]) by transfer_prover | 
| 568 | ||
| 569 | lemma Domainp_transfer[transfer_rule]: | |
| 570 | assumes [transfer_rule]: "bi_total B" | |
| 67399 | 571 | shows "((A ===> B ===> (=)) ===> A ===> (=)) Domainp Domainp" | 
| 71697 | 572 | unfolding Domainp_iff by transfer_prover | 
| 53952 | 573 | |
| 58182 | 574 | lemma reflp_transfer[transfer_rule]: | 
| 67399 | 575 | "bi_total A \<Longrightarrow> ((A ===> A ===> (=)) ===> (=)) reflp reflp" | 
| 53952 | 576 | "right_total A \<Longrightarrow> ((A ===> A ===> implies) ===> implies) reflp reflp" | 
| 67399 | 577 | "right_total A \<Longrightarrow> ((A ===> A ===> (=)) ===> implies) reflp reflp" | 
| 53952 | 578 | "bi_total A \<Longrightarrow> ((A ===> A ===> rev_implies) ===> rev_implies) reflp reflp" | 
| 67399 | 579 | "bi_total A \<Longrightarrow> ((A ===> A ===> (=)) ===> rev_implies) reflp reflp" | 
| 71697 | 580 | unfolding reflp_def rev_implies_def bi_total_def right_total_def rel_fun_def | 
| 53952 | 581 | by fast+ | 
| 582 | ||
| 583 | lemma right_unique_transfer [transfer_rule]: | |
| 59523 | 584 | "\<lbrakk> right_total A; right_total B; bi_unique B \<rbrakk> | 
| 67399 | 585 | \<Longrightarrow> ((A ===> B ===> (=)) ===> implies) right_unique right_unique" | 
| 71697 | 586 | unfolding right_unique_def right_total_def bi_unique_def rel_fun_def | 
| 53952 | 587 | by metis | 
| 47325 | 588 | |
| 59523 | 589 | lemma left_total_parametric [transfer_rule]: | 
| 590 | assumes [transfer_rule]: "bi_total A" "bi_total B" | |
| 67399 | 591 | shows "((A ===> B ===> (=)) ===> (=)) left_total left_total" | 
| 71697 | 592 | unfolding left_total_def by transfer_prover | 
| 59523 | 593 | |
| 594 | lemma right_total_parametric [transfer_rule]: | |
| 595 | assumes [transfer_rule]: "bi_total A" "bi_total B" | |
| 67399 | 596 | shows "((A ===> B ===> (=)) ===> (=)) right_total right_total" | 
| 71697 | 597 | unfolding right_total_def by transfer_prover | 
| 59523 | 598 | |
| 599 | lemma left_unique_parametric [transfer_rule]: | |
| 600 | assumes [transfer_rule]: "bi_unique A" "bi_total A" "bi_total B" | |
| 67399 | 601 | shows "((A ===> B ===> (=)) ===> (=)) left_unique left_unique" | 
| 71697 | 602 | unfolding left_unique_def by transfer_prover | 
| 59523 | 603 | |
| 604 | lemma prod_pred_parametric [transfer_rule]: | |
| 67399 | 605 | "((A ===> (=)) ===> (B ===> (=)) ===> rel_prod A B ===> (=)) pred_prod pred_prod" | 
| 71697 | 606 | unfolding prod.pred_set Basic_BNFs.fsts_def Basic_BNFs.snds_def fstsp.simps sndsp.simps | 
| 59523 | 607 | by simp transfer_prover | 
| 608 | ||
| 609 | lemma apfst_parametric [transfer_rule]: | |
| 610 | "((A ===> B) ===> rel_prod A C ===> rel_prod B C) apfst apfst" | |
| 71697 | 611 | unfolding apfst_def by transfer_prover | 
| 59523 | 612 | |
| 67399 | 613 | 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: 
56520diff
changeset | 614 | unfolding eq_onp_def rel_fun_def by auto | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 615 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 616 | lemma rel_fun_eq_onp_rel: | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 617 | 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 | 618 | 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 | 619 | |
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 620 | lemma eq_onp_transfer [transfer_rule]: | 
| 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 621 | assumes [transfer_rule]: "bi_unique A" | 
| 67399 | 622 | shows "((A ===> (=)) ===> A ===> A ===> (=)) eq_onp eq_onp" | 
| 71697 | 623 | unfolding eq_onp_def by transfer_prover | 
| 56524 
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
 kuncar parents: 
56520diff
changeset | 624 | |
| 57599 | 625 | lemma rtranclp_parametric [transfer_rule]: | 
| 626 | assumes "bi_unique A" "bi_total A" | |
| 67399 | 627 | shows "((A ===> A ===> (=)) ===> A ===> A ===> (=)) rtranclp rtranclp" | 
| 57599 | 628 | proof(rule rel_funI iffI)+ | 
| 629 | fix R :: "'a \<Rightarrow> 'a \<Rightarrow> bool" and R' x y x' y' | |
| 67399 | 630 | assume R: "(A ===> A ===> (=)) R R'" and "A x x'" | 
| 57599 | 631 |   {
 | 
| 632 | assume "R\<^sup>*\<^sup>* x y" "A y y'" | |
| 633 | thus "R'\<^sup>*\<^sup>* x' y'" | |
| 634 | proof(induction arbitrary: y') | |
| 635 | case base | |
| 60758 | 636 | with \<open>bi_unique A\<close> \<open>A x x'\<close> have "x' = y'" by(rule bi_uniqueDr) | 
| 57599 | 637 | thus ?case by simp | 
| 638 | next | |
| 639 | case (step y z z') | |
| 60758 | 640 | from \<open>bi_total A\<close> obtain y' where "A y y'" unfolding bi_total_def by blast | 
| 57599 | 641 | hence "R'\<^sup>*\<^sup>* x' y'" by(rule step.IH) | 
| 60758 | 642 | moreover from R \<open>A y y'\<close> \<open>A z z'\<close> \<open>R y z\<close> | 
| 57599 | 643 | have "R' y' z'" by(auto dest: rel_funD) | 
| 644 | ultimately show ?case .. | |
| 645 | qed | |
| 646 | next | |
| 647 | assume "R'\<^sup>*\<^sup>* x' y'" "A y y'" | |
| 648 | thus "R\<^sup>*\<^sup>* x y" | |
| 649 | proof(induction arbitrary: y) | |
| 650 | case base | |
| 60758 | 651 | with \<open>bi_unique A\<close> \<open>A x x'\<close> have "x = y" by(rule bi_uniqueDl) | 
| 57599 | 652 | thus ?case by simp | 
| 653 | next | |
| 654 | case (step y' z' z) | |
| 60758 | 655 | from \<open>bi_total A\<close> obtain y where "A y y'" unfolding bi_total_def by blast | 
| 57599 | 656 | hence "R\<^sup>*\<^sup>* x y" by(rule step.IH) | 
| 60758 | 657 | moreover from R \<open>A y y'\<close> \<open>A z z'\<close> \<open>R' y' z'\<close> | 
| 57599 | 658 | have "R y z" by(auto dest: rel_funD) | 
| 659 | ultimately show ?case .. | |
| 660 | qed | |
| 661 | } | |
| 662 | qed | |
| 663 | ||
| 59523 | 664 | lemma right_unique_parametric [transfer_rule]: | 
| 665 | assumes [transfer_rule]: "bi_total A" "bi_unique B" "bi_total B" | |
| 67399 | 666 | shows "((A ===> B ===> (=)) ===> (=)) right_unique right_unique" | 
| 71697 | 667 | unfolding right_unique_def by transfer_prover | 
| 59523 | 668 | |
| 61630 | 669 | lemma map_fun_parametric [transfer_rule]: | 
| 670 | "((A ===> B) ===> (C ===> D) ===> (B ===> C) ===> A ===> D) map_fun map_fun" | |
| 71697 | 671 | unfolding map_fun_def by transfer_prover | 
| 61630 | 672 | |
| 47325 | 673 | end | 
| 53011 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 kuncar parents: 
52358diff
changeset | 674 | |
| 64014 | 675 | |
| 71182 | 676 | subsection \<open>\<^const>\<open>of_bool\<close> and \<^const>\<open>of_nat\<close>\<close> | 
| 677 | ||
| 678 | context | |
| 679 | includes lifting_syntax | |
| 680 | begin | |
| 681 | ||
| 682 | lemma transfer_rule_of_bool: | |
| 683 | \<open>((\<longleftrightarrow>) ===> (\<cong>)) of_bool of_bool\<close> | |
| 684 | if [transfer_rule]: \<open>0 \<cong> 0\<close> \<open>1 \<cong> 1\<close> | |
| 685 | for R :: \<open>'a::zero_neq_one \<Rightarrow> 'b::zero_neq_one \<Rightarrow> bool\<close> (infix \<open>\<cong>\<close> 50) | |
| 71697 | 686 | unfolding of_bool_def by transfer_prover | 
| 64014 | 687 | |
| 688 | lemma transfer_rule_of_nat: | |
| 71182 | 689 | "((=) ===> (\<cong>)) of_nat of_nat" | 
| 690 | if [transfer_rule]: \<open>0 \<cong> 0\<close> \<open>1 \<cong> 1\<close> | |
| 691 | \<open>((\<cong>) ===> (\<cong>) ===> (\<cong>)) (+) (+)\<close> | |
| 692 | for R :: \<open>'a::semiring_1 \<Rightarrow> 'b::semiring_1 \<Rightarrow> bool\<close> (infix \<open>\<cong>\<close> 50) | |
| 71697 | 693 | unfolding of_nat_def by transfer_prover | 
| 64014 | 694 | |
| 53011 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 kuncar parents: 
52358diff
changeset | 695 | end | 
| 71182 | 696 | |
| 697 | end |