| author | huffman | 
| Sun, 22 Apr 2012 21:32:35 +0200 | |
| changeset 47684 | ead185e60b8c | 
| parent 47660 | 7a5c681c0265 | 
| child 47789 | 71a526ee569a | 
| permissions | -rw-r--r-- | 
| 47325 | 1  | 
(* Title: HOL/Transfer.thy  | 
2  | 
Author: Brian Huffman, TU Muenchen  | 
|
3  | 
*)  | 
|
4  | 
||
5  | 
header {* Generic theorem transfer using relations *}
 | 
|
6  | 
||
7  | 
theory Transfer  | 
|
8  | 
imports Plain Hilbert_Choice  | 
|
9  | 
uses ("Tools/transfer.ML")
 | 
|
10  | 
begin  | 
|
11  | 
||
12  | 
subsection {* Relator for function space *}
 | 
|
13  | 
||
14  | 
definition  | 
|
15  | 
  fun_rel :: "('a \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('c \<Rightarrow> 'd) \<Rightarrow> bool" (infixr "===>" 55)
 | 
|
16  | 
where  | 
|
17  | 
"fun_rel A B = (\<lambda>f g. \<forall>x y. A x y \<longrightarrow> B (f x) (g y))"  | 
|
18  | 
||
19  | 
lemma fun_relI [intro]:  | 
|
20  | 
assumes "\<And>x y. A x y \<Longrightarrow> B (f x) (g y)"  | 
|
21  | 
shows "(A ===> B) f g"  | 
|
22  | 
using assms by (simp add: fun_rel_def)  | 
|
23  | 
||
24  | 
lemma fun_relD:  | 
|
25  | 
assumes "(A ===> B) f g" and "A x y"  | 
|
26  | 
shows "B (f x) (g y)"  | 
|
27  | 
using assms by (simp add: fun_rel_def)  | 
|
28  | 
||
29  | 
lemma fun_relE:  | 
|
30  | 
assumes "(A ===> B) f g" and "A x y"  | 
|
31  | 
obtains "B (f x) (g y)"  | 
|
32  | 
using assms by (simp add: fun_rel_def)  | 
|
33  | 
||
34  | 
lemma fun_rel_eq:  | 
|
35  | 
shows "((op =) ===> (op =)) = (op =)"  | 
|
36  | 
by (auto simp add: fun_eq_iff elim: fun_relE)  | 
|
37  | 
||
38  | 
lemma fun_rel_eq_rel:  | 
|
39  | 
shows "((op =) ===> R) = (\<lambda>f g. \<forall>x. R (f x) (g x))"  | 
|
40  | 
by (simp add: fun_rel_def)  | 
|
41  | 
||
42  | 
||
43  | 
subsection {* Transfer method *}
 | 
|
44  | 
||
45  | 
text {* Explicit tags for application, abstraction, and relation
 | 
|
46  | 
membership allow for backward proof methods. *}  | 
|
47  | 
||
48  | 
definition App :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
|
49  | 
where "App f \<equiv> f"  | 
|
50  | 
||
51  | 
definition Abs :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
|
52  | 
where "Abs f \<equiv> f"  | 
|
53  | 
||
54  | 
definition Rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
 | 
|
55  | 
where "Rel r \<equiv> r"  | 
|
56  | 
||
57  | 
text {* Handling of meta-logic connectives *}
 | 
|
58  | 
||
59  | 
definition transfer_forall where  | 
|
60  | 
"transfer_forall \<equiv> All"  | 
|
61  | 
||
62  | 
definition transfer_implies where  | 
|
63  | 
"transfer_implies \<equiv> op \<longrightarrow>"  | 
|
64  | 
||
| 
47355
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
65  | 
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
 | 
66  | 
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
 | 
67  | 
|
| 47325 | 68  | 
lemma transfer_forall_eq: "(\<And>x. P x) \<equiv> Trueprop (transfer_forall (\<lambda>x. P x))"  | 
69  | 
unfolding atomize_all transfer_forall_def ..  | 
|
70  | 
||
71  | 
lemma transfer_implies_eq: "(A \<Longrightarrow> B) \<equiv> Trueprop (transfer_implies A B)"  | 
|
72  | 
unfolding atomize_imp transfer_implies_def ..  | 
|
73  | 
||
| 
47355
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
74  | 
lemma transfer_bforall_unfold:  | 
| 
 
3d9d98e0f1a4
add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
 
huffman 
parents: 
47325 
diff
changeset
 | 
75  | 
"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
 | 
76  | 
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
 | 
77  | 
|
| 
47658
 
7631f6f7873d
enable variant of transfer method that proves an implication instead of an equivalence
 
huffman 
parents: 
47637 
diff
changeset
 | 
78  | 
lemma transfer_start: "\<lbrakk>P; Rel (op =) P Q\<rbrakk> \<Longrightarrow> Q"  | 
| 47325 | 79  | 
unfolding Rel_def by simp  | 
80  | 
||
| 
47658
 
7631f6f7873d
enable variant of transfer method that proves an implication instead of an equivalence
 
huffman 
parents: 
47637 
diff
changeset
 | 
81  | 
lemma transfer_start': "\<lbrakk>P; Rel (op \<longrightarrow>) P Q\<rbrakk> \<Longrightarrow> Q"  | 
| 47325 | 82  | 
unfolding Rel_def by simp  | 
83  | 
||
| 
47635
 
ebb79474262c
rename 'correspondence' method to 'transfer_prover'
 
huffman 
parents: 
47627 
diff
changeset
 | 
84  | 
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
 | 
85  | 
by simp  | 
| 
 
1568dadd598a
make correspondence tactic more robust by replacing lhs with schematic variable before applying intro rules
 
huffman 
parents: 
47612 
diff
changeset
 | 
86  | 
|
| 47325 | 87  | 
lemma Rel_eq_refl: "Rel (op =) x x"  | 
88  | 
unfolding Rel_def ..  | 
|
89  | 
||
| 
47523
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
90  | 
lemma Rel_App:  | 
| 
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
91  | 
assumes "Rel (A ===> B) f g" and "Rel A x y"  | 
| 
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
92  | 
shows "Rel B (App f x) (App g y)"  | 
| 
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
93  | 
using assms unfolding Rel_def App_def fun_rel_def by fast  | 
| 
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
94  | 
|
| 
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
95  | 
lemma Rel_Abs:  | 
| 
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
96  | 
assumes "\<And>x y. Rel A x y \<Longrightarrow> Rel B (f x) (g y)"  | 
| 
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
97  | 
shows "Rel (A ===> B) (Abs (\<lambda>x. f x)) (Abs (\<lambda>y. g y))"  | 
| 
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
98  | 
using assms unfolding Rel_def Abs_def fun_rel_def by fast  | 
| 
 
1bf0e92c1ca0
make transfer method more deterministic by using SOLVED' on some subgoals
 
huffman 
parents: 
47503 
diff
changeset
 | 
99  | 
|
| 47325 | 100  | 
use "Tools/transfer.ML"  | 
101  | 
||
102  | 
setup Transfer.setup  | 
|
103  | 
||
| 47503 | 104  | 
declare fun_rel_eq [relator_eq]  | 
105  | 
||
| 47325 | 106  | 
hide_const (open) App Abs Rel  | 
107  | 
||
108  | 
||
109  | 
subsection {* Predicates on relations, i.e. ``class constraints'' *}
 | 
|
110  | 
||
111  | 
definition right_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
|
112  | 
where "right_total R \<longleftrightarrow> (\<forall>y. \<exists>x. R x y)"  | 
|
113  | 
||
114  | 
definition right_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
|
115  | 
where "right_unique R \<longleftrightarrow> (\<forall>x y z. R x y \<longrightarrow> R x z \<longrightarrow> y = z)"  | 
|
116  | 
||
117  | 
definition bi_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
|
118  | 
where "bi_total R \<longleftrightarrow> (\<forall>x. \<exists>y. R x y) \<and> (\<forall>y. \<exists>x. R x y)"  | 
|
119  | 
||
120  | 
definition bi_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
|
121  | 
where "bi_unique R \<longleftrightarrow>  | 
|
122  | 
(\<forall>x y z. R x y \<longrightarrow> R x z \<longrightarrow> y = z) \<and>  | 
|
123  | 
(\<forall>x y z. R x z \<longrightarrow> R y z \<longrightarrow> x = y)"  | 
|
124  | 
||
125  | 
lemma right_total_alt_def:  | 
|
126  | 
"right_total R \<longleftrightarrow> ((R ===> op \<longrightarrow>) ===> op \<longrightarrow>) All All"  | 
|
127  | 
unfolding right_total_def fun_rel_def  | 
|
128  | 
apply (rule iffI, fast)  | 
|
129  | 
apply (rule allI)  | 
|
130  | 
apply (drule_tac x="\<lambda>x. True" in spec)  | 
|
131  | 
apply (drule_tac x="\<lambda>y. \<exists>x. R x y" in spec)  | 
|
132  | 
apply fast  | 
|
133  | 
done  | 
|
134  | 
||
135  | 
lemma right_unique_alt_def:  | 
|
136  | 
"right_unique R \<longleftrightarrow> (R ===> R ===> op \<longrightarrow>) (op =) (op =)"  | 
|
137  | 
unfolding right_unique_def fun_rel_def by auto  | 
|
138  | 
||
139  | 
lemma bi_total_alt_def:  | 
|
140  | 
"bi_total R \<longleftrightarrow> ((R ===> op =) ===> op =) All All"  | 
|
141  | 
unfolding bi_total_def fun_rel_def  | 
|
142  | 
apply (rule iffI, fast)  | 
|
143  | 
apply safe  | 
|
144  | 
apply (drule_tac x="\<lambda>x. \<exists>y. R x y" in spec)  | 
|
145  | 
apply (drule_tac x="\<lambda>y. True" in spec)  | 
|
146  | 
apply fast  | 
|
147  | 
apply (drule_tac x="\<lambda>x. True" in spec)  | 
|
148  | 
apply (drule_tac x="\<lambda>y. \<exists>x. R x y" in spec)  | 
|
149  | 
apply fast  | 
|
150  | 
done  | 
|
151  | 
||
152  | 
lemma bi_unique_alt_def:  | 
|
153  | 
"bi_unique R \<longleftrightarrow> (R ===> R ===> op =) (op =) (op =)"  | 
|
154  | 
unfolding bi_unique_def fun_rel_def by auto  | 
|
155  | 
||
| 47660 | 156  | 
text {* Properties are preserved by relation composition. *}
 | 
157  | 
||
158  | 
lemma OO_def: "R OO S = (\<lambda>x z. \<exists>y. R x y \<and> S y z)"  | 
|
159  | 
by auto  | 
|
160  | 
||
161  | 
lemma bi_total_OO: "\<lbrakk>bi_total A; bi_total B\<rbrakk> \<Longrightarrow> bi_total (A OO B)"  | 
|
162  | 
unfolding bi_total_def OO_def by metis  | 
|
163  | 
||
164  | 
lemma bi_unique_OO: "\<lbrakk>bi_unique A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A OO B)"  | 
|
165  | 
unfolding bi_unique_def OO_def by metis  | 
|
166  | 
||
167  | 
lemma right_total_OO:  | 
|
168  | 
"\<lbrakk>right_total A; right_total B\<rbrakk> \<Longrightarrow> right_total (A OO B)"  | 
|
169  | 
unfolding right_total_def OO_def by metis  | 
|
170  | 
||
171  | 
lemma right_unique_OO:  | 
|
172  | 
"\<lbrakk>right_unique A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A OO B)"  | 
|
173  | 
unfolding right_unique_def OO_def by metis  | 
|
174  | 
||
| 47325 | 175  | 
|
176  | 
subsection {* Properties of relators *}
 | 
|
177  | 
||
178  | 
lemma right_total_eq [transfer_rule]: "right_total (op =)"  | 
|
179  | 
unfolding right_total_def by simp  | 
|
180  | 
||
181  | 
lemma right_unique_eq [transfer_rule]: "right_unique (op =)"  | 
|
182  | 
unfolding right_unique_def by simp  | 
|
183  | 
||
184  | 
lemma bi_total_eq [transfer_rule]: "bi_total (op =)"  | 
|
185  | 
unfolding bi_total_def by simp  | 
|
186  | 
||
187  | 
lemma bi_unique_eq [transfer_rule]: "bi_unique (op =)"  | 
|
188  | 
unfolding bi_unique_def by simp  | 
|
189  | 
||
190  | 
lemma right_total_fun [transfer_rule]:  | 
|
191  | 
"\<lbrakk>right_unique A; right_total B\<rbrakk> \<Longrightarrow> right_total (A ===> B)"  | 
|
192  | 
unfolding right_total_def fun_rel_def  | 
|
193  | 
apply (rule allI, rename_tac g)  | 
|
194  | 
apply (rule_tac x="\<lambda>x. SOME z. B z (g (THE y. A x y))" in exI)  | 
|
195  | 
apply clarify  | 
|
196  | 
apply (subgoal_tac "(THE y. A x y) = y", simp)  | 
|
197  | 
apply (rule someI_ex)  | 
|
198  | 
apply (simp)  | 
|
199  | 
apply (rule the_equality)  | 
|
200  | 
apply assumption  | 
|
201  | 
apply (simp add: right_unique_def)  | 
|
202  | 
done  | 
|
203  | 
||
204  | 
lemma right_unique_fun [transfer_rule]:  | 
|
205  | 
"\<lbrakk>right_total A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A ===> B)"  | 
|
206  | 
unfolding right_total_def right_unique_def fun_rel_def  | 
|
207  | 
by (clarify, rule ext, fast)  | 
|
208  | 
||
209  | 
lemma bi_total_fun [transfer_rule]:  | 
|
210  | 
"\<lbrakk>bi_unique A; bi_total B\<rbrakk> \<Longrightarrow> bi_total (A ===> B)"  | 
|
211  | 
unfolding bi_total_def fun_rel_def  | 
|
212  | 
apply safe  | 
|
213  | 
apply (rename_tac f)  | 
|
214  | 
apply (rule_tac x="\<lambda>y. SOME z. B (f (THE x. A x y)) z" in exI)  | 
|
215  | 
apply clarify  | 
|
216  | 
apply (subgoal_tac "(THE x. A x y) = x", simp)  | 
|
217  | 
apply (rule someI_ex)  | 
|
218  | 
apply (simp)  | 
|
219  | 
apply (rule the_equality)  | 
|
220  | 
apply assumption  | 
|
221  | 
apply (simp add: bi_unique_def)  | 
|
222  | 
apply (rename_tac g)  | 
|
223  | 
apply (rule_tac x="\<lambda>x. SOME z. B z (g (THE y. A x y))" in exI)  | 
|
224  | 
apply clarify  | 
|
225  | 
apply (subgoal_tac "(THE y. A x y) = y", simp)  | 
|
226  | 
apply (rule someI_ex)  | 
|
227  | 
apply (simp)  | 
|
228  | 
apply (rule the_equality)  | 
|
229  | 
apply assumption  | 
|
230  | 
apply (simp add: bi_unique_def)  | 
|
231  | 
done  | 
|
232  | 
||
233  | 
lemma bi_unique_fun [transfer_rule]:  | 
|
234  | 
"\<lbrakk>bi_total A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A ===> B)"  | 
|
235  | 
unfolding bi_total_def bi_unique_def fun_rel_def fun_eq_iff  | 
|
236  | 
by (safe, metis, fast)  | 
|
237  | 
||
238  | 
||
| 
47635
 
ebb79474262c
rename 'correspondence' method to 'transfer_prover'
 
huffman 
parents: 
47627 
diff
changeset
 | 
239  | 
subsection {* Transfer rules *}
 | 
| 47325 | 240  | 
|
| 47684 | 241  | 
text {* Transfer rules using implication instead of equality on booleans. *}
 | 
242  | 
||
243  | 
lemma eq_imp_transfer [transfer_rule]:  | 
|
244  | 
"right_unique A \<Longrightarrow> (A ===> A ===> op \<longrightarrow>) (op =) (op =)"  | 
|
245  | 
unfolding right_unique_alt_def .  | 
|
246  | 
||
247  | 
lemma forall_imp_transfer [transfer_rule]:  | 
|
248  | 
"right_total A \<Longrightarrow> ((A ===> op \<longrightarrow>) ===> op \<longrightarrow>) transfer_forall transfer_forall"  | 
|
249  | 
unfolding right_total_alt_def transfer_forall_def .  | 
|
250  | 
||
| 47636 | 251  | 
lemma eq_transfer [transfer_rule]:  | 
| 47325 | 252  | 
assumes "bi_unique A"  | 
253  | 
shows "(A ===> A ===> op =) (op =) (op =)"  | 
|
254  | 
using assms unfolding bi_unique_def fun_rel_def by auto  | 
|
255  | 
||
| 47636 | 256  | 
lemma All_transfer [transfer_rule]:  | 
| 47325 | 257  | 
assumes "bi_total A"  | 
258  | 
shows "((A ===> op =) ===> op =) All All"  | 
|
259  | 
using assms unfolding bi_total_def fun_rel_def by fast  | 
|
260  | 
||
| 47636 | 261  | 
lemma Ex_transfer [transfer_rule]:  | 
| 47325 | 262  | 
assumes "bi_total A"  | 
263  | 
shows "((A ===> op =) ===> op =) Ex Ex"  | 
|
264  | 
using assms unfolding bi_total_def fun_rel_def by fast  | 
|
265  | 
||
| 47636 | 266  | 
lemma If_transfer [transfer_rule]: "(op = ===> A ===> A ===> A) If If"  | 
| 47325 | 267  | 
unfolding fun_rel_def by simp  | 
268  | 
||
| 47636 | 269  | 
lemma Let_transfer [transfer_rule]: "(A ===> (A ===> B) ===> B) Let Let"  | 
| 47612 | 270  | 
unfolding fun_rel_def by simp  | 
271  | 
||
| 47636 | 272  | 
lemma id_transfer [transfer_rule]: "(A ===> A) id id"  | 
| 47625 | 273  | 
unfolding fun_rel_def by simp  | 
274  | 
||
| 47636 | 275  | 
lemma comp_transfer [transfer_rule]:  | 
| 47325 | 276  | 
"((B ===> C) ===> (A ===> B) ===> (A ===> C)) (op \<circ>) (op \<circ>)"  | 
277  | 
unfolding fun_rel_def by simp  | 
|
278  | 
||
| 47636 | 279  | 
lemma fun_upd_transfer [transfer_rule]:  | 
| 47325 | 280  | 
assumes [transfer_rule]: "bi_unique A"  | 
281  | 
shows "((A ===> B) ===> A ===> B ===> A ===> B) fun_upd fun_upd"  | 
|
| 
47635
 
ebb79474262c
rename 'correspondence' method to 'transfer_prover'
 
huffman 
parents: 
47627 
diff
changeset
 | 
282  | 
unfolding fun_upd_def [abs_def] by transfer_prover  | 
| 47325 | 283  | 
|
| 47637 | 284  | 
lemma nat_case_transfer [transfer_rule]:  | 
285  | 
"(A ===> (op = ===> A) ===> op = ===> A) nat_case nat_case"  | 
|
286  | 
unfolding fun_rel_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
 | 
287  | 
|
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
288  | 
text {* Fallback rule for transferring universal quantifiers over
 | 
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
289  | 
correspondence relations that are not bi-total, and do not have  | 
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
290  | 
custom transfer rules (e.g. relations between function types). *}  | 
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
291  | 
|
| 47637 | 292  | 
lemma Domainp_iff: "Domainp T x \<longleftrightarrow> (\<exists>y. T x y)"  | 
293  | 
by auto  | 
|
294  | 
||
| 
47627
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
295  | 
lemma Domainp_forall_transfer [transfer_rule]:  | 
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
296  | 
assumes "right_total A"  | 
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
297  | 
shows "((A ===> op =) ===> op =)  | 
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
298  | 
(transfer_bforall (Domainp A)) transfer_forall"  | 
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
299  | 
using assms unfolding right_total_def  | 
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
300  | 
unfolding transfer_forall_def transfer_bforall_def fun_rel_def Domainp_iff  | 
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
301  | 
by metis  | 
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
302  | 
|
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
303  | 
text {* Preferred rule for transferring universal quantifiers over
 | 
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
304  | 
bi-total correspondence relations (later rules are tried first). *}  | 
| 
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
305  | 
|
| 47636 | 306  | 
lemma forall_transfer [transfer_rule]:  | 
| 
47627
 
2b1d3eda59eb
add secondary transfer rule for universal quantifiers on non-bi-total relations
 
huffman 
parents: 
47625 
diff
changeset
 | 
307  | 
"bi_total A \<Longrightarrow> ((A ===> op =) ===> op =) transfer_forall transfer_forall"  | 
| 47636 | 308  | 
unfolding transfer_forall_def by (rule All_transfer)  | 
| 47325 | 309  | 
|
310  | 
end  |