author | blanchet |
Tue, 16 Sep 2014 19:23:37 +0200 | |
changeset 58352 | 37745650a3f4 |
parent 58182 | 82478e6c60cb |
child 58386 | 6999f55645ea |
permissions | -rw-r--r-- |
47325 | 1 |
(* Title: HOL/Transfer.thy |
2 |
Author: Brian Huffman, TU Muenchen |
|
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51955
diff
changeset
|
3 |
Author: Ondrej Kuncar, TU Muenchen |
47325 | 4 |
*) |
5 |
||
6 |
header {* Generic theorem transfer using relations *} |
|
7 |
||
8 |
theory Transfer |
|
58128 | 9 |
imports Hilbert_Choice Metis Option |
47325 | 10 |
begin |
11 |
||
58182 | 12 |
(* We import Option here although it's not needed here. |
13 |
By doing this, we avoid a diamond problem for BNF and |
|
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
14 |
FP sugar interpretation defined in this file. *) |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
15 |
|
47325 | 16 |
subsection {* Relator for function space *} |
17 |
||
53011
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52358
diff
changeset
|
18 |
locale lifting_syntax |
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52358
diff
changeset
|
19 |
begin |
55945 | 20 |
notation rel_fun (infixr "===>" 55) |
53011
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52358
diff
changeset
|
21 |
notation map_fun (infixr "--->" 55) |
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52358
diff
changeset
|
22 |
end |
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52358
diff
changeset
|
23 |
|
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52358
diff
changeset
|
24 |
context |
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52358
diff
changeset
|
25 |
begin |
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52358
diff
changeset
|
26 |
interpretation lifting_syntax . |
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52358
diff
changeset
|
27 |
|
55945 | 28 |
lemma rel_funD2: |
29 |
assumes "rel_fun A B f g" and "A x x" |
|
47937
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47924
diff
changeset
|
30 |
shows "B (f x) (g x)" |
55945 | 31 |
using assms by (rule rel_funD) |
47937
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47924
diff
changeset
|
32 |
|
55945 | 33 |
lemma rel_funE: |
34 |
assumes "rel_fun A B f g" and "A x y" |
|
47325 | 35 |
obtains "B (f x) (g y)" |
55945 | 36 |
using assms by (simp add: rel_fun_def) |
47325 | 37 |
|
55945 | 38 |
lemmas rel_fun_eq = fun.rel_eq |
47325 | 39 |
|
55945 | 40 |
lemma rel_fun_eq_rel: |
41 |
shows "rel_fun (op =) R = (\<lambda>f g. \<forall>x. R (f x) (g x))" |
|
42 |
by (simp add: rel_fun_def) |
|
47325 | 43 |
|
44 |
||
45 |
subsection {* Transfer method *} |
|
46 |
||
47789
71a526ee569a
implement transfer tactic with more scalable forward proof methods
huffman
parents:
47684
diff
changeset
|
47 |
text {* Explicit tag for relation membership allows for |
71a526ee569a
implement transfer tactic with more scalable forward proof methods
huffman
parents:
47684
diff
changeset
|
48 |
backward proof methods. *} |
47325 | 49 |
|
50 |
definition Rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" |
|
51 |
where "Rel r \<equiv> r" |
|
52 |
||
49975
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents:
48891
diff
changeset
|
53 |
text {* Handling of equality relations *} |
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents:
48891
diff
changeset
|
54 |
|
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents:
48891
diff
changeset
|
55 |
definition is_equality :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" |
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents:
48891
diff
changeset
|
56 |
where "is_equality R \<longleftrightarrow> R = (op =)" |
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents:
48891
diff
changeset
|
57 |
|
51437
8739f8abbecb
fixing transfer tactic - unfold fully identity relation by using relator_eq
kuncar
parents:
51112
diff
changeset
|
58 |
lemma is_equality_eq: "is_equality (op =)" |
8739f8abbecb
fixing transfer tactic - unfold fully identity relation by using relator_eq
kuncar
parents:
51112
diff
changeset
|
59 |
unfolding is_equality_def by simp |
8739f8abbecb
fixing transfer tactic - unfold fully identity relation by using relator_eq
kuncar
parents:
51112
diff
changeset
|
60 |
|
52354
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
61 |
text {* Reverse implication for monotonicity rules *} |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
62 |
|
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
63 |
definition rev_implies where |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
64 |
"rev_implies x y \<longleftrightarrow> (y \<longrightarrow> x)" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
65 |
|
47325 | 66 |
text {* Handling of meta-logic connectives *} |
67 |
||
68 |
definition transfer_forall where |
|
69 |
"transfer_forall \<equiv> All" |
|
70 |
||
71 |
definition transfer_implies where |
|
72 |
"transfer_implies \<equiv> op \<longrightarrow>" |
|
73 |
||
47355
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
huffman
parents:
47325
diff
changeset
|
74 |
definition transfer_bforall :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" |
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
huffman
parents:
47325
diff
changeset
|
75 |
where "transfer_bforall \<equiv> (\<lambda>P Q. \<forall>x. P x \<longrightarrow> Q x)" |
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
huffman
parents:
47325
diff
changeset
|
76 |
|
47325 | 77 |
lemma transfer_forall_eq: "(\<And>x. P x) \<equiv> Trueprop (transfer_forall (\<lambda>x. P x))" |
78 |
unfolding atomize_all transfer_forall_def .. |
|
79 |
||
80 |
lemma transfer_implies_eq: "(A \<Longrightarrow> B) \<equiv> Trueprop (transfer_implies A B)" |
|
81 |
unfolding atomize_imp transfer_implies_def .. |
|
82 |
||
47355
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
huffman
parents:
47325
diff
changeset
|
83 |
lemma transfer_bforall_unfold: |
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
huffman
parents:
47325
diff
changeset
|
84 |
"Trueprop (transfer_bforall P (\<lambda>x. Q x)) \<equiv> (\<And>x. P x \<Longrightarrow> Q x)" |
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
huffman
parents:
47325
diff
changeset
|
85 |
unfolding transfer_bforall_def atomize_imp atomize_all .. |
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
huffman
parents:
47325
diff
changeset
|
86 |
|
47658
7631f6f7873d
enable variant of transfer method that proves an implication instead of an equivalence
huffman
parents:
47637
diff
changeset
|
87 |
lemma transfer_start: "\<lbrakk>P; Rel (op =) P Q\<rbrakk> \<Longrightarrow> Q" |
47325 | 88 |
unfolding Rel_def by simp |
89 |
||
47658
7631f6f7873d
enable variant of transfer method that proves an implication instead of an equivalence
huffman
parents:
47637
diff
changeset
|
90 |
lemma transfer_start': "\<lbrakk>P; Rel (op \<longrightarrow>) P Q\<rbrakk> \<Longrightarrow> Q" |
47325 | 91 |
unfolding Rel_def by simp |
92 |
||
47635
ebb79474262c
rename 'correspondence' method to 'transfer_prover'
huffman
parents:
47627
diff
changeset
|
93 |
lemma transfer_prover_start: "\<lbrakk>x = x'; Rel R x' y\<rbrakk> \<Longrightarrow> Rel R x y" |
47618
1568dadd598a
make correspondence tactic more robust by replacing lhs with schematic variable before applying intro rules
huffman
parents:
47612
diff
changeset
|
94 |
by simp |
1568dadd598a
make correspondence tactic more robust by replacing lhs with schematic variable before applying intro rules
huffman
parents:
47612
diff
changeset
|
95 |
|
52358
f4c4bcb0d564
implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents:
52354
diff
changeset
|
96 |
lemma untransfer_start: "\<lbrakk>Q; Rel (op =) P Q\<rbrakk> \<Longrightarrow> P" |
f4c4bcb0d564
implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents:
52354
diff
changeset
|
97 |
unfolding Rel_def by simp |
f4c4bcb0d564
implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents:
52354
diff
changeset
|
98 |
|
47325 | 99 |
lemma Rel_eq_refl: "Rel (op =) x x" |
100 |
unfolding Rel_def .. |
|
101 |
||
47789
71a526ee569a
implement transfer tactic with more scalable forward proof methods
huffman
parents:
47684
diff
changeset
|
102 |
lemma Rel_app: |
47523
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
huffman
parents:
47503
diff
changeset
|
103 |
assumes "Rel (A ===> B) f g" and "Rel A x y" |
47789
71a526ee569a
implement transfer tactic with more scalable forward proof methods
huffman
parents:
47684
diff
changeset
|
104 |
shows "Rel B (f x) (g y)" |
55945 | 105 |
using assms unfolding Rel_def rel_fun_def by fast |
47523
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
huffman
parents:
47503
diff
changeset
|
106 |
|
47789
71a526ee569a
implement transfer tactic with more scalable forward proof methods
huffman
parents:
47684
diff
changeset
|
107 |
lemma Rel_abs: |
47523
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
huffman
parents:
47503
diff
changeset
|
108 |
assumes "\<And>x y. Rel A x y \<Longrightarrow> Rel B (f x) (g y)" |
47789
71a526ee569a
implement transfer tactic with more scalable forward proof methods
huffman
parents:
47684
diff
changeset
|
109 |
shows "Rel (A ===> B) (\<lambda>x. f x) (\<lambda>y. g y)" |
55945 | 110 |
using assms unfolding Rel_def rel_fun_def by fast |
47523
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
huffman
parents:
47503
diff
changeset
|
111 |
|
47325 | 112 |
subsection {* Predicates on relations, i.e. ``class constraints'' *} |
113 |
||
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
114 |
definition left_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
115 |
where "left_total R \<longleftrightarrow> (\<forall>x. \<exists>y. R x y)" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
116 |
|
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
117 |
definition left_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
118 |
where "left_unique R \<longleftrightarrow> (\<forall>x y z. R x z \<longrightarrow> R y z \<longrightarrow> x = y)" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
119 |
|
47325 | 120 |
definition right_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" |
121 |
where "right_total R \<longleftrightarrow> (\<forall>y. \<exists>x. R x y)" |
|
122 |
||
123 |
definition right_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" |
|
124 |
where "right_unique R \<longleftrightarrow> (\<forall>x y z. R x y \<longrightarrow> R x z \<longrightarrow> y = z)" |
|
125 |
||
126 |
definition bi_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" |
|
127 |
where "bi_total R \<longleftrightarrow> (\<forall>x. \<exists>y. R x y) \<and> (\<forall>y. \<exists>x. R x y)" |
|
128 |
||
129 |
definition bi_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" |
|
130 |
where "bi_unique R \<longleftrightarrow> |
|
131 |
(\<forall>x y z. R x y \<longrightarrow> R x z \<longrightarrow> y = z) \<and> |
|
132 |
(\<forall>x y z. R x z \<longrightarrow> R y z \<longrightarrow> x = y)" |
|
133 |
||
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
134 |
lemma left_uniqueI: "(\<And>x y z. \<lbrakk> A x z; A y z \<rbrakk> \<Longrightarrow> x = y) \<Longrightarrow> left_unique A" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
135 |
unfolding left_unique_def by blast |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
136 |
|
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
137 |
lemma left_uniqueD: "\<lbrakk> left_unique A; A x z; A y z \<rbrakk> \<Longrightarrow> x = y" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
138 |
unfolding left_unique_def by blast |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
139 |
|
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
140 |
lemma left_totalI: |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
141 |
"(\<And>x. \<exists>y. R x y) \<Longrightarrow> left_total R" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
142 |
unfolding left_total_def by blast |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
143 |
|
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
144 |
lemma left_totalE: |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
145 |
assumes "left_total R" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
146 |
obtains "(\<And>x. \<exists>y. R x y)" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
147 |
using assms unfolding left_total_def by blast |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
148 |
|
53927 | 149 |
lemma bi_uniqueDr: "\<lbrakk> bi_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z" |
150 |
by(simp add: bi_unique_def) |
|
151 |
||
152 |
lemma bi_uniqueDl: "\<lbrakk> bi_unique A; A x y; A z y \<rbrakk> \<Longrightarrow> x = z" |
|
153 |
by(simp add: bi_unique_def) |
|
154 |
||
155 |
lemma right_uniqueI: "(\<And>x y z. \<lbrakk> A x y; A x z \<rbrakk> \<Longrightarrow> y = z) \<Longrightarrow> right_unique A" |
|
56085 | 156 |
unfolding right_unique_def by fast |
53927 | 157 |
|
158 |
lemma right_uniqueD: "\<lbrakk> right_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z" |
|
56085 | 159 |
unfolding right_unique_def by fast |
53927 | 160 |
|
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
161 |
lemma right_total_alt_def2: |
47325 | 162 |
"right_total R \<longleftrightarrow> ((R ===> op \<longrightarrow>) ===> op \<longrightarrow>) All All" |
55945 | 163 |
unfolding right_total_def rel_fun_def |
47325 | 164 |
apply (rule iffI, fast) |
165 |
apply (rule allI) |
|
166 |
apply (drule_tac x="\<lambda>x. True" in spec) |
|
167 |
apply (drule_tac x="\<lambda>y. \<exists>x. R x y" in spec) |
|
168 |
apply fast |
|
169 |
done |
|
170 |
||
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
171 |
lemma right_unique_alt_def2: |
47325 | 172 |
"right_unique R \<longleftrightarrow> (R ===> R ===> op \<longrightarrow>) (op =) (op =)" |
55945 | 173 |
unfolding right_unique_def rel_fun_def by auto |
47325 | 174 |
|
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
175 |
lemma bi_total_alt_def2: |
47325 | 176 |
"bi_total R \<longleftrightarrow> ((R ===> op =) ===> op =) All All" |
55945 | 177 |
unfolding bi_total_def rel_fun_def |
47325 | 178 |
apply (rule iffI, fast) |
179 |
apply safe |
|
180 |
apply (drule_tac x="\<lambda>x. \<exists>y. R x y" in spec) |
|
181 |
apply (drule_tac x="\<lambda>y. True" in spec) |
|
182 |
apply fast |
|
183 |
apply (drule_tac x="\<lambda>x. True" in spec) |
|
184 |
apply (drule_tac x="\<lambda>y. \<exists>x. R x y" in spec) |
|
185 |
apply fast |
|
186 |
done |
|
187 |
||
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
188 |
lemma bi_unique_alt_def2: |
47325 | 189 |
"bi_unique R \<longleftrightarrow> (R ===> R ===> op =) (op =) (op =)" |
55945 | 190 |
unfolding bi_unique_def rel_fun_def by auto |
47325 | 191 |
|
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
192 |
lemma [simp]: |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
193 |
shows left_unique_conversep: "left_unique A\<inverse>\<inverse> \<longleftrightarrow> right_unique A" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
194 |
and right_unique_conversep: "right_unique A\<inverse>\<inverse> \<longleftrightarrow> left_unique A" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
195 |
by(auto simp add: left_unique_def right_unique_def) |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
196 |
|
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
197 |
lemma [simp]: |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
198 |
shows left_total_conversep: "left_total A\<inverse>\<inverse> \<longleftrightarrow> right_total A" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
199 |
and right_total_conversep: "right_total A\<inverse>\<inverse> \<longleftrightarrow> left_total A" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
200 |
by(simp_all add: left_total_def right_total_def) |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
201 |
|
53944 | 202 |
lemma bi_unique_conversep [simp]: "bi_unique R\<inverse>\<inverse> = bi_unique R" |
203 |
by(auto simp add: bi_unique_def) |
|
204 |
||
205 |
lemma bi_total_conversep [simp]: "bi_total R\<inverse>\<inverse> = bi_total R" |
|
206 |
by(auto simp add: bi_total_def) |
|
207 |
||
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
208 |
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:
56520
diff
changeset
|
209 |
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:
56520
diff
changeset
|
210 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
211 |
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:
56520
diff
changeset
|
212 |
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:
56520
diff
changeset
|
213 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
214 |
lemma bi_total_alt_def: "bi_total A = (left_total A \<and> right_total A)" |
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
215 |
unfolding left_total_def right_total_def bi_total_def by blast |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
216 |
|
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
217 |
lemma bi_unique_alt_def: "bi_unique A = (left_unique A \<and> right_unique A)" |
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
218 |
unfolding left_unique_def right_unique_def bi_unique_def by blast |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
219 |
|
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
220 |
lemma bi_totalI: "left_total R \<Longrightarrow> right_total R \<Longrightarrow> bi_total R" |
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
221 |
unfolding bi_total_alt_def .. |
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
222 |
|
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
223 |
lemma bi_uniqueI: "left_unique R \<Longrightarrow> right_unique R \<Longrightarrow> bi_unique R" |
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
224 |
unfolding bi_unique_alt_def .. |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
225 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
226 |
end |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
227 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
228 |
subsection {* Equality restricted by a predicate *} |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
229 |
|
58182 | 230 |
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:
56520
diff
changeset
|
231 |
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:
56520
diff
changeset
|
232 |
|
58182 | 233 |
lemma eq_onp_Grp: "eq_onp P = BNF_Def.Grp (Collect P) id" |
234 |
unfolding eq_onp_def Grp_def by auto |
|
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
235 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
236 |
lemma eq_onp_to_eq: |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
237 |
assumes "eq_onp P x y" |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
238 |
shows "x = y" |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
239 |
using assms by (simp add: eq_onp_def) |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
240 |
|
58182 | 241 |
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:
56654
diff
changeset
|
242 |
by (simp add: eq_onp_def) |
660ffb526069
predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents:
56654
diff
changeset
|
243 |
|
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
244 |
lemma eq_onp_same_args: |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
245 |
shows "eq_onp P x x = P x" |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
246 |
using assms by (auto simp add: eq_onp_def) |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
247 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
248 |
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:
56677
diff
changeset
|
249 |
by auto |
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
250 |
|
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
251 |
ML_file "Tools/Transfer/transfer.ML" |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
252 |
setup Transfer.setup |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
253 |
declare refl [transfer_rule] |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
254 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
255 |
hide_const (open) Rel |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
256 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
257 |
context |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
258 |
begin |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
259 |
interpretation lifting_syntax . |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
260 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
261 |
text {* Handling of domains *} |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
262 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
263 |
lemma Domainp_iff: "Domainp T x \<longleftrightarrow> (\<exists>y. T x y)" |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
264 |
by auto |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
265 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
266 |
lemma Domaimp_refl[transfer_domain_rule]: |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
267 |
"Domainp T = Domainp T" .. |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
268 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
269 |
lemma Domainp_prod_fun_eq[relator_domain]: |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
270 |
"Domainp (op= ===> T) = (\<lambda>f. \<forall>x. (Domainp T) (f x))" |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
271 |
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:
56085
diff
changeset
|
272 |
|
47660 | 273 |
text {* Properties are preserved by relation composition. *} |
274 |
||
275 |
lemma OO_def: "R OO S = (\<lambda>x z. \<exists>y. R x y \<and> S y z)" |
|
276 |
by auto |
|
277 |
||
278 |
lemma bi_total_OO: "\<lbrakk>bi_total A; bi_total B\<rbrakk> \<Longrightarrow> bi_total (A OO B)" |
|
56085 | 279 |
unfolding bi_total_def OO_def by fast |
47660 | 280 |
|
281 |
lemma bi_unique_OO: "\<lbrakk>bi_unique A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A OO B)" |
|
56085 | 282 |
unfolding bi_unique_def OO_def by blast |
47660 | 283 |
|
284 |
lemma right_total_OO: |
|
285 |
"\<lbrakk>right_total A; right_total B\<rbrakk> \<Longrightarrow> right_total (A OO B)" |
|
56085 | 286 |
unfolding right_total_def OO_def by fast |
47660 | 287 |
|
288 |
lemma right_unique_OO: |
|
289 |
"\<lbrakk>right_unique A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A OO B)" |
|
56085 | 290 |
unfolding right_unique_def OO_def by fast |
47660 | 291 |
|
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
292 |
lemma left_total_OO: "left_total R \<Longrightarrow> left_total S \<Longrightarrow> left_total (R OO S)" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
293 |
unfolding left_total_def OO_def by fast |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
294 |
|
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
295 |
lemma left_unique_OO: "left_unique R \<Longrightarrow> left_unique S \<Longrightarrow> left_unique (R OO S)" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
296 |
unfolding left_unique_def OO_def by blast |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
297 |
|
47325 | 298 |
|
299 |
subsection {* Properties of relators *} |
|
300 |
||
58182 | 301 |
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:
56085
diff
changeset
|
302 |
unfolding left_total_def by blast |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
303 |
|
58182 | 304 |
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:
56085
diff
changeset
|
305 |
unfolding left_unique_def by blast |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
306 |
|
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
307 |
lemma right_total_eq [transfer_rule]: "right_total op=" |
47325 | 308 |
unfolding right_total_def by simp |
309 |
||
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
310 |
lemma right_unique_eq [transfer_rule]: "right_unique op=" |
47325 | 311 |
unfolding right_unique_def by simp |
312 |
||
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
313 |
lemma bi_total_eq[transfer_rule]: "bi_total (op =)" |
47325 | 314 |
unfolding bi_total_def by simp |
315 |
||
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
316 |
lemma bi_unique_eq[transfer_rule]: "bi_unique (op =)" |
47325 | 317 |
unfolding bi_unique_def by simp |
318 |
||
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
319 |
lemma left_total_fun[transfer_rule]: |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
320 |
"\<lbrakk>left_unique A; left_total B\<rbrakk> \<Longrightarrow> left_total (A ===> B)" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
321 |
unfolding left_total_def rel_fun_def |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
322 |
apply (rule allI, rename_tac f) |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
323 |
apply (rule_tac x="\<lambda>y. SOME z. B (f (THE x. A x y)) z" in exI) |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
324 |
apply clarify |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
325 |
apply (subgoal_tac "(THE x. A x y) = x", simp) |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
326 |
apply (rule someI_ex) |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
327 |
apply (simp) |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
328 |
apply (rule the_equality) |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
329 |
apply assumption |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
330 |
apply (simp add: left_unique_def) |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
331 |
done |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
332 |
|
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
333 |
lemma left_unique_fun[transfer_rule]: |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
334 |
"\<lbrakk>left_total A; left_unique B\<rbrakk> \<Longrightarrow> left_unique (A ===> B)" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
335 |
unfolding left_total_def left_unique_def rel_fun_def |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
336 |
by (clarify, rule ext, fast) |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
337 |
|
47325 | 338 |
lemma right_total_fun [transfer_rule]: |
339 |
"\<lbrakk>right_unique A; right_total B\<rbrakk> \<Longrightarrow> right_total (A ===> B)" |
|
55945 | 340 |
unfolding right_total_def rel_fun_def |
47325 | 341 |
apply (rule allI, rename_tac g) |
342 |
apply (rule_tac x="\<lambda>x. SOME z. B z (g (THE y. A x y))" in exI) |
|
343 |
apply clarify |
|
344 |
apply (subgoal_tac "(THE y. A x y) = y", simp) |
|
345 |
apply (rule someI_ex) |
|
346 |
apply (simp) |
|
347 |
apply (rule the_equality) |
|
348 |
apply assumption |
|
349 |
apply (simp add: right_unique_def) |
|
350 |
done |
|
351 |
||
352 |
lemma right_unique_fun [transfer_rule]: |
|
353 |
"\<lbrakk>right_total A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A ===> B)" |
|
55945 | 354 |
unfolding right_total_def right_unique_def rel_fun_def |
47325 | 355 |
by (clarify, rule ext, fast) |
356 |
||
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
357 |
lemma bi_total_fun[transfer_rule]: |
47325 | 358 |
"\<lbrakk>bi_unique A; bi_total B\<rbrakk> \<Longrightarrow> bi_total (A ===> B)" |
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
359 |
unfolding bi_unique_alt_def bi_total_alt_def |
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
360 |
by (blast intro: right_total_fun left_total_fun) |
47325 | 361 |
|
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
362 |
lemma bi_unique_fun[transfer_rule]: |
47325 | 363 |
"\<lbrakk>bi_total A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A ===> B)" |
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
364 |
unfolding bi_unique_alt_def bi_total_alt_def |
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
365 |
by (blast intro: right_unique_fun left_unique_fun) |
47325 | 366 |
|
56543
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
kuncar
parents:
56524
diff
changeset
|
367 |
end |
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
kuncar
parents:
56524
diff
changeset
|
368 |
|
58182 | 369 |
ML_file "Tools/Transfer/transfer_bnf.ML" |
56543
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
kuncar
parents:
56524
diff
changeset
|
370 |
|
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
kuncar
parents:
56524
diff
changeset
|
371 |
declare pred_fun_def [simp] |
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
kuncar
parents:
56524
diff
changeset
|
372 |
declare rel_fun_eq [relator_eq] |
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
kuncar
parents:
56524
diff
changeset
|
373 |
|
47635
ebb79474262c
rename 'correspondence' method to 'transfer_prover'
huffman
parents:
47627
diff
changeset
|
374 |
subsection {* Transfer rules *} |
47325 | 375 |
|
56543
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
kuncar
parents:
56524
diff
changeset
|
376 |
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:
56524
diff
changeset
|
377 |
begin |
9bd56f2e4c10
bi_unique and co. rules from the BNF hook must be introduced after bi_unique op= and co. rules are introduced
kuncar
parents:
56524
diff
changeset
|
378 |
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:
56524
diff
changeset
|
379 |
|
53952 | 380 |
lemma Domainp_forall_transfer [transfer_rule]: |
381 |
assumes "right_total A" |
|
382 |
shows "((A ===> op =) ===> op =) |
|
383 |
(transfer_bforall (Domainp A)) transfer_forall" |
|
384 |
using assms unfolding right_total_def |
|
55945 | 385 |
unfolding transfer_forall_def transfer_bforall_def rel_fun_def Domainp_iff |
56085 | 386 |
by fast |
53952 | 387 |
|
47684 | 388 |
text {* Transfer rules using implication instead of equality on booleans. *} |
389 |
||
52354
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
390 |
lemma transfer_forall_transfer [transfer_rule]: |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
391 |
"bi_total A \<Longrightarrow> ((A ===> op =) ===> op =) transfer_forall transfer_forall" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
392 |
"right_total A \<Longrightarrow> ((A ===> op =) ===> implies) transfer_forall transfer_forall" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
393 |
"right_total A \<Longrightarrow> ((A ===> implies) ===> implies) transfer_forall transfer_forall" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
394 |
"bi_total A \<Longrightarrow> ((A ===> op =) ===> rev_implies) transfer_forall transfer_forall" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
395 |
"bi_total A \<Longrightarrow> ((A ===> rev_implies) ===> rev_implies) transfer_forall transfer_forall" |
55945 | 396 |
unfolding transfer_forall_def rev_implies_def rel_fun_def right_total_def bi_total_def |
56085 | 397 |
by fast+ |
52354
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
398 |
|
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
399 |
lemma transfer_implies_transfer [transfer_rule]: |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
400 |
"(op = ===> op = ===> op = ) transfer_implies transfer_implies" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
401 |
"(rev_implies ===> implies ===> implies ) transfer_implies transfer_implies" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
402 |
"(rev_implies ===> op = ===> implies ) transfer_implies transfer_implies" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
403 |
"(op = ===> implies ===> implies ) transfer_implies transfer_implies" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
404 |
"(op = ===> op = ===> implies ) transfer_implies transfer_implies" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
405 |
"(implies ===> rev_implies ===> rev_implies) transfer_implies transfer_implies" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
406 |
"(implies ===> op = ===> rev_implies) transfer_implies transfer_implies" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
407 |
"(op = ===> rev_implies ===> rev_implies) transfer_implies transfer_implies" |
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
408 |
"(op = ===> op = ===> rev_implies) transfer_implies transfer_implies" |
55945 | 409 |
unfolding transfer_implies_def rev_implies_def rel_fun_def by auto |
52354
acb4f932dd24
implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents:
51956
diff
changeset
|
410 |
|
47684 | 411 |
lemma eq_imp_transfer [transfer_rule]: |
412 |
"right_unique A \<Longrightarrow> (A ===> A ===> op \<longrightarrow>) (op =) (op =)" |
|
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
413 |
unfolding right_unique_alt_def2 . |
47684 | 414 |
|
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
415 |
text {* Transfer rules using equality. *} |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
416 |
|
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
417 |
lemma left_unique_transfer [transfer_rule]: |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
418 |
assumes "right_total A" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
419 |
assumes "right_total B" |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
420 |
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:
56085
diff
changeset
|
421 |
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:
56085
diff
changeset
|
422 |
using assms unfolding left_unique_def[abs_def] right_total_def bi_unique_def rel_fun_def |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
423 |
by metis |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56085
diff
changeset
|
424 |
|
47636 | 425 |
lemma eq_transfer [transfer_rule]: |
47325 | 426 |
assumes "bi_unique A" |
427 |
shows "(A ===> A ===> op =) (op =) (op =)" |
|
55945 | 428 |
using assms unfolding bi_unique_def rel_fun_def by auto |
47325 | 429 |
|
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51955
diff
changeset
|
430 |
lemma right_total_Ex_transfer[transfer_rule]: |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51955
diff
changeset
|
431 |
assumes "right_total A" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51955
diff
changeset
|
432 |
shows "((A ===> op=) ===> op=) (Bex (Collect (Domainp A))) Ex" |
55945 | 433 |
using assms unfolding right_total_def Bex_def rel_fun_def Domainp_iff[abs_def] |
56085 | 434 |
by fast |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51955
diff
changeset
|
435 |
|
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51955
diff
changeset
|
436 |
lemma right_total_All_transfer[transfer_rule]: |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51955
diff
changeset
|
437 |
assumes "right_total A" |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51955
diff
changeset
|
438 |
shows "((A ===> op =) ===> op =) (Ball (Collect (Domainp A))) All" |
55945 | 439 |
using assms unfolding right_total_def Ball_def rel_fun_def Domainp_iff[abs_def] |
56085 | 440 |
by fast |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51955
diff
changeset
|
441 |
|
47636 | 442 |
lemma All_transfer [transfer_rule]: |
47325 | 443 |
assumes "bi_total A" |
444 |
shows "((A ===> op =) ===> op =) All All" |
|
55945 | 445 |
using assms unfolding bi_total_def rel_fun_def by fast |
47325 | 446 |
|
47636 | 447 |
lemma Ex_transfer [transfer_rule]: |
47325 | 448 |
assumes "bi_total A" |
449 |
shows "((A ===> op =) ===> op =) Ex Ex" |
|
55945 | 450 |
using assms unfolding bi_total_def rel_fun_def by fast |
47325 | 451 |
|
47636 | 452 |
lemma If_transfer [transfer_rule]: "(op = ===> A ===> A ===> A) If If" |
55945 | 453 |
unfolding rel_fun_def by simp |
47325 | 454 |
|
47636 | 455 |
lemma Let_transfer [transfer_rule]: "(A ===> (A ===> B) ===> B) Let Let" |
55945 | 456 |
unfolding rel_fun_def by simp |
47612 | 457 |
|
47636 | 458 |
lemma id_transfer [transfer_rule]: "(A ===> A) id id" |
55945 | 459 |
unfolding rel_fun_def by simp |
47625 | 460 |
|
47636 | 461 |
lemma comp_transfer [transfer_rule]: |
47325 | 462 |
"((B ===> C) ===> (A ===> B) ===> (A ===> C)) (op \<circ>) (op \<circ>)" |
55945 | 463 |
unfolding rel_fun_def by simp |
47325 | 464 |
|
47636 | 465 |
lemma fun_upd_transfer [transfer_rule]: |
47325 | 466 |
assumes [transfer_rule]: "bi_unique A" |
467 |
shows "((A ===> B) ===> A ===> B ===> A ===> B) fun_upd fun_upd" |
|
47635
ebb79474262c
rename 'correspondence' method to 'transfer_prover'
huffman
parents:
47627
diff
changeset
|
468 |
unfolding fun_upd_def [abs_def] by transfer_prover |
47325 | 469 |
|
55415 | 470 |
lemma case_nat_transfer [transfer_rule]: |
471 |
"(A ===> (op = ===> A) ===> op = ===> A) case_nat case_nat" |
|
55945 | 472 |
unfolding rel_fun_def by (simp split: nat.split) |
47627
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
huffman
parents:
47625
diff
changeset
|
473 |
|
55415 | 474 |
lemma rec_nat_transfer [transfer_rule]: |
475 |
"(A ===> (op = ===> A ===> A) ===> op = ===> A) rec_nat rec_nat" |
|
55945 | 476 |
unfolding rel_fun_def by (clarsimp, rename_tac n, induct_tac n, simp_all) |
47924 | 477 |
|
478 |
lemma funpow_transfer [transfer_rule]: |
|
479 |
"(op = ===> (A ===> A) ===> (A ===> A)) compow compow" |
|
480 |
unfolding funpow_def by transfer_prover |
|
481 |
||
53952 | 482 |
lemma mono_transfer[transfer_rule]: |
483 |
assumes [transfer_rule]: "bi_total A" |
|
484 |
assumes [transfer_rule]: "(A ===> A ===> op=) op\<le> op\<le>" |
|
485 |
assumes [transfer_rule]: "(B ===> B ===> op=) op\<le> op\<le>" |
|
486 |
shows "((A ===> B) ===> op=) mono mono" |
|
487 |
unfolding mono_def[abs_def] by transfer_prover |
|
488 |
||
58182 | 489 |
lemma right_total_relcompp_transfer[transfer_rule]: |
53952 | 490 |
assumes [transfer_rule]: "right_total B" |
58182 | 491 |
shows "((A ===> B ===> op=) ===> (B ===> C ===> op=) ===> A ===> C ===> op=) |
53952 | 492 |
(\<lambda>R S x z. \<exists>y\<in>Collect (Domainp B). R x y \<and> S y z) op OO" |
493 |
unfolding OO_def[abs_def] by transfer_prover |
|
494 |
||
58182 | 495 |
lemma relcompp_transfer[transfer_rule]: |
53952 | 496 |
assumes [transfer_rule]: "bi_total B" |
497 |
shows "((A ===> B ===> op=) ===> (B ===> C ===> op=) ===> A ===> C ===> op=) op OO op OO" |
|
498 |
unfolding OO_def[abs_def] by transfer_prover |
|
47627
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
huffman
parents:
47625
diff
changeset
|
499 |
|
53952 | 500 |
lemma right_total_Domainp_transfer[transfer_rule]: |
501 |
assumes [transfer_rule]: "right_total B" |
|
502 |
shows "((A ===> B ===> op=) ===> A ===> op=) (\<lambda>T x. \<exists>y\<in>Collect(Domainp B). T x y) Domainp" |
|
503 |
apply(subst(2) Domainp_iff[abs_def]) by transfer_prover |
|
504 |
||
505 |
lemma Domainp_transfer[transfer_rule]: |
|
506 |
assumes [transfer_rule]: "bi_total B" |
|
507 |
shows "((A ===> B ===> op=) ===> A ===> op=) Domainp Domainp" |
|
508 |
unfolding Domainp_iff[abs_def] by transfer_prover |
|
509 |
||
58182 | 510 |
lemma reflp_transfer[transfer_rule]: |
53952 | 511 |
"bi_total A \<Longrightarrow> ((A ===> A ===> op=) ===> op=) reflp reflp" |
512 |
"right_total A \<Longrightarrow> ((A ===> A ===> implies) ===> implies) reflp reflp" |
|
513 |
"right_total A \<Longrightarrow> ((A ===> A ===> op=) ===> implies) reflp reflp" |
|
514 |
"bi_total A \<Longrightarrow> ((A ===> A ===> rev_implies) ===> rev_implies) reflp reflp" |
|
515 |
"bi_total A \<Longrightarrow> ((A ===> A ===> op=) ===> rev_implies) reflp reflp" |
|
58182 | 516 |
using assms unfolding reflp_def[abs_def] rev_implies_def bi_total_def right_total_def rel_fun_def |
53952 | 517 |
by fast+ |
518 |
||
519 |
lemma right_unique_transfer [transfer_rule]: |
|
520 |
assumes [transfer_rule]: "right_total A" |
|
521 |
assumes [transfer_rule]: "right_total B" |
|
522 |
assumes [transfer_rule]: "bi_unique B" |
|
523 |
shows "((A ===> B ===> op=) ===> implies) right_unique right_unique" |
|
55945 | 524 |
using assms unfolding right_unique_def[abs_def] right_total_def bi_unique_def rel_fun_def |
53952 | 525 |
by metis |
47325 | 526 |
|
56524
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
527 |
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:
56520
diff
changeset
|
528 |
unfolding eq_onp_def rel_fun_def by auto |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
529 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
530 |
lemma rel_fun_eq_onp_rel: |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
531 |
shows "((eq_onp R) ===> S) = (\<lambda>f g. \<forall>x. R x \<longrightarrow> S (f x) (g x))" |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
532 |
by (auto simp add: eq_onp_def rel_fun_def) |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
533 |
|
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
534 |
lemma eq_onp_transfer [transfer_rule]: |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
535 |
assumes [transfer_rule]: "bi_unique A" |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
536 |
shows "((A ===> op=) ===> A ===> A ===> op=) eq_onp eq_onp" |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
537 |
unfolding eq_onp_def[abs_def] by transfer_prover |
f4ba736040fa
setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
56520
diff
changeset
|
538 |
|
57599 | 539 |
lemma rtranclp_parametric [transfer_rule]: |
540 |
assumes "bi_unique A" "bi_total A" |
|
541 |
shows "((A ===> A ===> op =) ===> A ===> A ===> op =) rtranclp rtranclp" |
|
542 |
proof(rule rel_funI iffI)+ |
|
543 |
fix R :: "'a \<Rightarrow> 'a \<Rightarrow> bool" and R' x y x' y' |
|
544 |
assume R: "(A ===> A ===> op =) R R'" and "A x x'" |
|
545 |
{ |
|
546 |
assume "R\<^sup>*\<^sup>* x y" "A y y'" |
|
547 |
thus "R'\<^sup>*\<^sup>* x' y'" |
|
548 |
proof(induction arbitrary: y') |
|
549 |
case base |
|
550 |
with `bi_unique A` `A x x'` have "x' = y'" by(rule bi_uniqueDr) |
|
551 |
thus ?case by simp |
|
552 |
next |
|
553 |
case (step y z z') |
|
554 |
from `bi_total A` obtain y' where "A y y'" unfolding bi_total_def by blast |
|
555 |
hence "R'\<^sup>*\<^sup>* x' y'" by(rule step.IH) |
|
556 |
moreover from R `A y y'` `A z z'` `R y z` |
|
557 |
have "R' y' z'" by(auto dest: rel_funD) |
|
558 |
ultimately show ?case .. |
|
559 |
qed |
|
560 |
next |
|
561 |
assume "R'\<^sup>*\<^sup>* x' y'" "A y y'" |
|
562 |
thus "R\<^sup>*\<^sup>* x y" |
|
563 |
proof(induction arbitrary: y) |
|
564 |
case base |
|
565 |
with `bi_unique A` `A x x'` have "x = y" by(rule bi_uniqueDl) |
|
566 |
thus ?case by simp |
|
567 |
next |
|
568 |
case (step y' z' z) |
|
569 |
from `bi_total A` obtain y where "A y y'" unfolding bi_total_def by blast |
|
570 |
hence "R\<^sup>*\<^sup>* x y" by(rule step.IH) |
|
571 |
moreover from R `A y y'` `A z z'` `R' y' z'` |
|
572 |
have "R y z" by(auto dest: rel_funD) |
|
573 |
ultimately show ?case .. |
|
574 |
qed |
|
575 |
} |
|
576 |
qed |
|
577 |
||
47325 | 578 |
end |
53011
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52358
diff
changeset
|
579 |
|
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
kuncar
parents:
52358
diff
changeset
|
580 |
end |