| author | Christian Urban <urbanc@in.tum.de> | 
| Wed, 14 Apr 2010 13:21:11 +0200 | |
| changeset 1831 | 16653e702d89 | 
| parent 1828 | f374ffd50c7c | 
| child 1833 | 2050b5723c04 | 
| permissions | -rw-r--r-- | 
| 
1797
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
1  | 
theory Lambda  | 
| 
1773
 
c0eac04ae3b4
added README and moved examples into separate directory
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1599 
diff
changeset
 | 
2  | 
imports "../Parser"  | 
| 1594 | 3  | 
begin  | 
4  | 
||
5  | 
atom_decl name  | 
|
6  | 
||
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
7  | 
nominal_datatype lam =  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
8  | 
Var "name"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
9  | 
| App "lam" "lam"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
10  | 
| Lam x::"name" l::"lam" bind x in l  | 
| 1594 | 11  | 
|
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
12  | 
lemmas supp_fn' = lam.fv[simplified lam.supp]  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
13  | 
|
| 1805 | 14  | 
declare lam.perm[eqvt]  | 
15  | 
||
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
16  | 
section {* Strong Induction Principles*}
 | 
| 1594 | 17  | 
|
| 
1797
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
18  | 
(*  | 
| 
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
19  | 
Old way of establishing strong induction  | 
| 
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
20  | 
principles by chosing a fresh name.  | 
| 
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
21  | 
*)  | 
| 1594 | 22  | 
lemma  | 
23  | 
fixes c::"'a::fs"  | 
|
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
24  | 
assumes a1: "\<And>name c. P c (Var name)"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
25  | 
and a2: "\<And>lam1 lam2 c. \<lbrakk>\<And>d. P d lam1; \<And>d. P d lam2\<rbrakk> \<Longrightarrow> P c (App lam1 lam2)"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
26  | 
and a3: "\<And>name lam c. \<lbrakk>atom name \<sharp> c; \<And>d. P d lam\<rbrakk> \<Longrightarrow> P c (Lam name lam)"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
27  | 
shows "P c lam"  | 
| 1594 | 28  | 
proof -  | 
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
29  | 
have "\<And>p. P c (p \<bullet> lam)"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
30  | 
apply(induct lam arbitrary: c rule: lam.induct)  | 
| 1805 | 31  | 
apply(perm_simp)  | 
| 1594 | 32  | 
apply(rule a1)  | 
| 1805 | 33  | 
apply(perm_simp)  | 
| 1594 | 34  | 
apply(rule a2)  | 
| 
1797
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
35  | 
apply(assumption)  | 
| 1594 | 36  | 
apply(assumption)  | 
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
37  | 
apply(subgoal_tac "\<exists>new::name. (atom new) \<sharp> (c, Lam (p \<bullet> name) (p \<bullet> lam))")  | 
| 1594 | 38  | 
defer  | 
39  | 
apply(simp add: fresh_def)  | 
|
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
40  | 
apply(rule_tac X="supp (c, Lam (p \<bullet> name) (p \<bullet> lam))" in obtain_at_base)  | 
| 1594 | 41  | 
apply(simp add: supp_Pair finite_supp)  | 
42  | 
apply(blast)  | 
|
43  | 
apply(erule exE)  | 
|
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
44  | 
apply(rule_tac t="p \<bullet> Lam name lam" and  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
45  | 
s="(((p \<bullet> name) \<leftrightarrow> new) + p) \<bullet> Lam name lam" in subst)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
46  | 
apply(simp del: lam.perm)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
47  | 
apply(subst lam.perm)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
48  | 
apply(subst (2) lam.perm)  | 
| 1594 | 49  | 
apply(rule flip_fresh_fresh)  | 
50  | 
apply(simp add: fresh_def)  | 
|
51  | 
apply(simp only: supp_fn')  | 
|
52  | 
apply(simp)  | 
|
53  | 
apply(simp add: fresh_Pair)  | 
|
54  | 
apply(simp)  | 
|
55  | 
apply(rule a3)  | 
|
56  | 
apply(simp add: fresh_Pair)  | 
|
57  | 
apply(drule_tac x="((p \<bullet> name) \<leftrightarrow> new) + p" in meta_spec)  | 
|
58  | 
apply(simp)  | 
|
59  | 
done  | 
|
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
60  | 
then have "P c (0 \<bullet> lam)" by blast  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
61  | 
then show "P c lam" by simp  | 
| 1594 | 62  | 
qed  | 
63  | 
||
| 
1797
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
64  | 
(*  | 
| 
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
65  | 
New way of establishing strong induction  | 
| 
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
66  | 
principles by using a appropriate permutation.  | 
| 
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
67  | 
*)  | 
| 1594 | 68  | 
lemma  | 
69  | 
fixes c::"'a::fs"  | 
|
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
70  | 
assumes a1: "\<And>name c. P c (Var name)"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
71  | 
and a2: "\<And>lam1 lam2 c. \<lbrakk>\<And>d. P d lam1; \<And>d. P d lam2\<rbrakk> \<Longrightarrow> P c (App lam1 lam2)"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
72  | 
and a3: "\<And>name lam c. \<lbrakk>atom name \<sharp> c; \<And>d. P d lam\<rbrakk> \<Longrightarrow> P c (Lam name lam)"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
73  | 
shows "P c lam"  | 
| 1594 | 74  | 
proof -  | 
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
75  | 
have "\<And>p. P c (p \<bullet> lam)"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
76  | 
apply(induct lam arbitrary: c rule: lam.induct)  | 
| 1805 | 77  | 
apply(perm_simp)  | 
| 1594 | 78  | 
apply(rule a1)  | 
| 1805 | 79  | 
apply(perm_simp)  | 
| 1594 | 80  | 
apply(rule a2)  | 
81  | 
apply(assumption)  | 
|
| 
1797
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
82  | 
apply(assumption)  | 
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
83  | 
    apply(subgoal_tac "\<exists>q. (q \<bullet> {p \<bullet> atom name}) \<sharp>* c \<and> supp (p \<bullet> Lam name lam) \<sharp>* q")
 | 
| 1594 | 84  | 
apply(erule exE)  | 
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
85  | 
apply(rule_tac t="p \<bullet> Lam name lam" and  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
86  | 
s="q \<bullet> p \<bullet> Lam name lam" in subst)  | 
| 1594 | 87  | 
defer  | 
| 1805 | 88  | 
apply(simp)  | 
| 1594 | 89  | 
apply(rule a3)  | 
90  | 
apply(simp add: eqvts fresh_star_def)  | 
|
91  | 
apply(drule_tac x="q + p" in meta_spec)  | 
|
92  | 
apply(simp)  | 
|
| 
1797
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
93  | 
apply(rule at_set_avoiding2)  | 
| 
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
94  | 
apply(simp add: finite_supp)  | 
| 
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
95  | 
apply(simp add: finite_supp)  | 
| 
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
96  | 
apply(simp add: finite_supp)  | 
| 1805 | 97  | 
apply(perm_simp)  | 
| 
1797
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
98  | 
apply(simp add: fresh_star_def fresh_def supp_fn')  | 
| 
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
99  | 
apply(rule supp_perm_eq)  | 
| 
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
100  | 
apply(simp)  | 
| 
 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1773 
diff
changeset
 | 
101  | 
done  | 
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
102  | 
then have "P c (0 \<bullet> lam)" by blast  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
103  | 
then show "P c lam" by simp  | 
| 1594 | 104  | 
qed  | 
105  | 
||
| 1805 | 106  | 
section {* Typing *}
 | 
107  | 
||
108  | 
nominal_datatype ty =  | 
|
109  | 
TVar string  | 
|
| 
1810
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
110  | 
| TFun ty ty  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
111  | 
|
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
112  | 
notation  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
113  | 
 TFun ("_ \<rightarrow> _") 
 | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
114  | 
|
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
115  | 
declare ty.perm[eqvt]  | 
| 1805 | 116  | 
|
117  | 
inductive  | 
|
118  | 
valid :: "(name \<times> ty) list \<Rightarrow> bool"  | 
|
119  | 
where  | 
|
120  | 
"valid []"  | 
|
121  | 
| "\<lbrakk>atom x \<sharp> Gamma; valid Gamma\<rbrakk> \<Longrightarrow> valid ((x, T)#Gamma)"  | 
|
122  | 
||
| 1828 | 123  | 
inductive  | 
124  | 
  typing :: "(name\<times>ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
 | 
|
125  | 
where  | 
|
126  | 
t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"  | 
|
127  | 
| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2 \<or> \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"  | 
|
128  | 
| t_Lam[intro]: "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam x t : T1 \<rightarrow> T2"  | 
|
129  | 
||
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
130  | 
ML {*
 | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
131  | 
val inductive_atomize = @{thms induct_atomize};
 | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
132  | 
|
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
133  | 
val atomize_conv =  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
134  | 
MetaSimplifier.rewrite_cterm (true, false, false) (K (K NONE))  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
135  | 
(HOL_basic_ss addsimps inductive_atomize);  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
136  | 
val atomize_intr = Conv.fconv_rule (Conv.prems_conv ~1 atomize_conv);  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
137  | 
fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
138  | 
(Conv.params_conv ~1 (K (Conv.prems_conv ~1 atomize_conv)) ctxt));  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
139  | 
*}  | 
| 1828 | 140  | 
|
| 1805 | 141  | 
ML {*
 | 
| 1828 | 142  | 
fun map_term f t =  | 
143  | 
(case f t of  | 
|
144  | 
NONE => map_term' f t  | 
|
145  | 
| x => x)  | 
|
146  | 
and map_term' f (t $ u) =  | 
|
147  | 
(case (map_term f t, map_term f u) of  | 
|
148  | 
(NONE, NONE) => NONE  | 
|
149  | 
| (SOME t'', NONE) => SOME (t'' $ u)  | 
|
150  | 
| (NONE, SOME u'') => SOME (t $ u'')  | 
|
151  | 
| (SOME t'', SOME u'') => SOME (t'' $ u''))  | 
|
152  | 
| map_term' f (Abs (s, T, t)) =  | 
|
153  | 
(case map_term f t of  | 
|
154  | 
NONE => NONE  | 
|
155  | 
| SOME t'' => SOME (Abs (s, T, t'')))  | 
|
156  | 
| map_term' _ _ = NONE;  | 
|
157  | 
||
158  | 
fun map_thm_tac ctxt tac thm =  | 
|
159  | 
let  | 
|
160  | 
val monos = Inductive.get_monos ctxt  | 
|
161  | 
in  | 
|
162  | 
EVERY [cut_facts_tac [thm] 1, etac rev_mp 1,  | 
|
163  | 
REPEAT_DETERM (FIRSTGOAL (resolve_tac monos)),  | 
|
164  | 
REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))]  | 
|
165  | 
end  | 
|
166  | 
||
167  | 
(*  | 
|
168  | 
proves F[f t] from F[t] where F[t] is the given theorem  | 
|
169  | 
||
170  | 
- F needs to be monotone  | 
|
171  | 
- f returns either SOME for a term it fires  | 
|
172  | 
and NONE elsewhere  | 
|
173  | 
*)  | 
|
174  | 
fun map_thm ctxt f tac thm =  | 
|
175  | 
let  | 
|
176  | 
val opt_goal_trm = map_term f (prop_of thm)  | 
|
177  | 
fun prove goal =  | 
|
178  | 
Goal.prove ctxt [] [] goal (fn _ => map_thm_tac ctxt tac thm)  | 
|
179  | 
in  | 
|
180  | 
case opt_goal_trm of  | 
|
181  | 
NONE => thm  | 
|
182  | 
| SOME goal => prove goal  | 
|
183  | 
end  | 
|
184  | 
||
185  | 
fun transform_prem ctxt names thm =  | 
|
186  | 
let  | 
|
187  | 
  fun split_conj names (Const ("op &", _) $ p $ q) = 
 | 
|
188  | 
(case head_of p of  | 
|
189  | 
Const (name, _) => if name mem names then SOME q else NONE  | 
|
190  | 
| _ => NONE)  | 
|
191  | 
| split_conj _ _ = NONE;  | 
|
192  | 
in  | 
|
193  | 
map_thm ctxt (split_conj names) (etac conjunct2 1) thm  | 
|
194  | 
end  | 
|
| 1805 | 195  | 
*}  | 
196  | 
||
| 1828 | 197  | 
ML {*
 | 
198  | 
open Nominal_Permeq  | 
|
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
199  | 
open Nominal_ThmDecls  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
200  | 
*}  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
201  | 
|
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
202  | 
ML {*
 | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
203  | 
fun mk_perm p trm =  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
204  | 
let  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
205  | 
val ty = fastype_of trm  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
206  | 
in  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
207  | 
  Const (@{const_name "permute"}, @{typ "perm"} --> ty --> ty) $ p $ trm
 | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
208  | 
end  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
209  | 
|
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
210  | 
fun mk_minus p =  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
211  | 
 Const (@{const_name "uminus"}, @{typ "perm => perm"}) $ p   
 | 
| 1828 | 212  | 
*}  | 
213  | 
||
214  | 
ML {* 
 | 
|
215  | 
fun single_case_tac ctxt pred_names pi intro =  | 
|
216  | 
let  | 
|
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
217  | 
val thy = ProofContext.theory_of ctxt  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
218  | 
val cpi = Thm.cterm_of thy (mk_minus pi)  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
219  | 
  val rule = Drule.instantiate' [] [SOME cpi] @{thm permute_boolE}
 | 
| 1828 | 220  | 
in  | 
221  | 
eqvt_strict_tac ctxt [] [] THEN'  | 
|
222  | 
  SUBPROOF (fn {prems, context as ctxt, ...} =>
 | 
|
223  | 
let  | 
|
224  | 
val prems' = map (transform_prem ctxt pred_names) prems  | 
|
225  | 
val side_cond_tac = EVERY'  | 
|
226  | 
[ rtac rule,  | 
|
227  | 
          eqvt_strict_tac ctxt @{thms permute_minus_cancel(2)} [],
 | 
|
228  | 
resolve_tac prems' ]  | 
|
229  | 
in  | 
|
230  | 
HEADGOAL (rtac intro THEN_ALL_NEW (resolve_tac prems' ORELSE' side_cond_tac))  | 
|
231  | 
end) ctxt  | 
|
232  | 
end  | 
|
233  | 
*}  | 
|
234  | 
||
235  | 
ML {*
 | 
|
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
236  | 
fun prepare_pred params_no pi pred =  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
237  | 
let  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
238  | 
val (c, xs) = strip_comb pred;  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
239  | 
val (xs1, xs2) = chop params_no xs  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
240  | 
in  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
241  | 
HOLogic.mk_imp  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
242  | 
(pred, list_comb (c, xs1 @ map (mk_perm pi) xs2))  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
243  | 
end  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
244  | 
*}  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
245  | 
|
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
246  | 
ML {*
 | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
247  | 
fun transp ([] :: _) = []  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
248  | 
| transp xs = map hd xs :: transp (map tl xs);  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
249  | 
*}  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
250  | 
|
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
251  | 
ML {* 
 | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
252  | 
Local_Theory.note;  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
253  | 
Local_Theory.notes;  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
254  | 
fold_map  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
255  | 
*}  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
256  | 
|
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
257  | 
ML {*
 | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
258  | 
fun note_named_thm (name, thm) ctxt =  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
259  | 
let  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
260  | 
val thm_name = Binding.qualified_name  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
261  | 
(Long_Name.qualify (Long_Name.base_name name) "eqvt")  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
262  | 
val attr = Attrib.internal (K eqvt_add)  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
263  | 
in  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
264  | 
Local_Theory.note ((thm_name, [attr]), [thm]) ctxt  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
265  | 
end  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
266  | 
*}  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
267  | 
|
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
268  | 
ML {*
 | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
269  | 
fun eqvt_rel_tac pred_name ctxt =  | 
| 1828 | 270  | 
let  | 
271  | 
val thy = ProofContext.theory_of ctxt  | 
|
272  | 
  val ({names, ...}, {raw_induct, intrs, ...}) =
 | 
|
273  | 
Inductive.the_inductive ctxt (Sign.intern_const thy pred_name)  | 
|
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
274  | 
val raw_induct = atomize_induct ctxt raw_induct;  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
275  | 
val intros = map atomize_intr intrs;  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
276  | 
val params_no = length (Inductive.params_of raw_induct)  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
277  | 
val (([raw_concl], [raw_pi]), ctxt') =  | 
| 1828 | 278  | 
ctxt |> Variable.import_terms false [concl_of raw_induct]  | 
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
279  | 
||>> Variable.variant_fixes ["pi"]  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
280  | 
  val pi = Free (raw_pi, @{typ perm})
 | 
| 1828 | 281  | 
val preds = map (fst o HOLogic.dest_imp)  | 
282  | 
(HOLogic.dest_conj (HOLogic.dest_Trueprop raw_concl));  | 
|
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
283  | 
val goal = HOLogic.mk_Trueprop  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
284  | 
(foldr1 HOLogic.mk_conj (map (prepare_pred params_no pi) preds))  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
285  | 
  val thm = Goal.prove ctxt' [] [] goal (fn {context,...} => 
 | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
286  | 
HEADGOAL (EVERY' (rtac raw_induct :: map (single_case_tac context names pi) intros)))  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
287  | 
|> singleton (ProofContext.export ctxt' ctxt)  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
288  | 
val thms = map (fn th => zero_var_indexes (th RS mp)) (Datatype_Aux.split_conj_thm thm)  | 
| 1828 | 289  | 
in  | 
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
290  | 
ctxt |> fold_map note_named_thm (names ~~ thms)  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
291  | 
|> snd  | 
| 1828 | 292  | 
end  | 
293  | 
*}  | 
|
294  | 
||
295  | 
||
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
296  | 
ML {*
 | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
297  | 
local structure P = OuterParse and K = OuterKeyword in  | 
| 1814 | 298  | 
|
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
299  | 
val _ =  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
300  | 
OuterSyntax.local_theory "equivariance"  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
301  | 
"prove equivariance for inductive predicate involving nominal datatypes" K.thy_decl  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
302  | 
(P.xname >> eqvt_rel_tac);  | 
| 1814 | 303  | 
|
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
304  | 
end;  | 
| 
1816
 
56cebe7f8e24
some small tunings (incompleted work in Lambda.thy)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1814 
diff
changeset
 | 
305  | 
*}  | 
| 
 
56cebe7f8e24
some small tunings (incompleted work in Lambda.thy)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1814 
diff
changeset
 | 
306  | 
|
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
307  | 
equivariance valid  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
308  | 
equivariance typing  | 
| 
1816
 
56cebe7f8e24
some small tunings (incompleted work in Lambda.thy)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1814 
diff
changeset
 | 
309  | 
|
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
310  | 
thm valid.eqvt  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
311  | 
thm typing.eqvt  | 
| 
1811
 
ae176476b525
implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1810 
diff
changeset
 | 
312  | 
thm eqvts  | 
| 
 
ae176476b525
implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1810 
diff
changeset
 | 
313  | 
thm eqvts_raw  | 
| 
 
ae176476b525
implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1810 
diff
changeset
 | 
314  | 
|
| 
1810
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
315  | 
|
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
316  | 
inductive  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
317  | 
alpha_lam_raw'  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
318  | 
where  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
319  | 
"name = namea \<Longrightarrow> alpha_lam_raw' (Var_raw name) (Var_raw namea)"  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
320  | 
| "\<lbrakk>alpha_lam_raw' lam_raw1 lam_raw1a; alpha_lam_raw' lam_raw2 lam_raw2a\<rbrakk> \<Longrightarrow>  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
321  | 
alpha_lam_raw' (App_raw lam_raw1 lam_raw2) (App_raw lam_raw1a lam_raw2a)"  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
322  | 
| "\<exists>pi. ({atom name}, lam_raw) \<approx>gen alpha_lam_raw fv_lam_raw pi ({atom namea}, lam_rawa) \<Longrightarrow>
 | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
323  | 
alpha_lam_raw' (Lam_raw name lam_raw) (Lam_raw namea lam_rawa)"  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
324  | 
|
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
325  | 
declare permute_lam_raw.simps[eqvt]  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
326  | 
(*declare alpha_gen_real_eqvt[eqvt]*)  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
327  | 
(*equivariance alpha_lam_raw'*)  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
328  | 
|
| 
1810
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
329  | 
lemma  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
330  | 
assumes a: "alpha_lam_raw' t1 t2"  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
331  | 
shows "alpha_lam_raw' (p \<bullet> t1) (p \<bullet> t2)"  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
332  | 
using a  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
333  | 
apply(induct)  | 
| 
 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1805 
diff
changeset
 | 
334  | 
oops  | 
| 1805 | 335  | 
|
336  | 
||
| 
1800
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
337  | 
section {* size function *}
 | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
338  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
339  | 
lemma size_eqvt_raw:  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
340  | 
fixes t::"lam_raw"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
341  | 
shows "size (pi \<bullet> t) = size t"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
342  | 
apply (induct rule: lam_raw.inducts)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
343  | 
apply simp_all  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
344  | 
done  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
345  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
346  | 
instantiation lam :: size  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
347  | 
begin  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
348  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
349  | 
quotient_definition  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
350  | 
"size_lam :: lam \<Rightarrow> nat"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
351  | 
is  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
352  | 
"size :: lam_raw \<Rightarrow> nat"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
353  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
354  | 
lemma size_rsp:  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
355  | 
"alpha_lam_raw x y \<Longrightarrow> size x = size y"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
356  | 
apply (induct rule: alpha_lam_raw.inducts)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
357  | 
apply (simp_all only: lam_raw.size)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
358  | 
apply (simp_all only: alphas)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
359  | 
apply clarify  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
360  | 
apply (simp_all only: size_eqvt_raw)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
361  | 
done  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
362  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
363  | 
lemma [quot_respect]:  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
364  | 
"(alpha_lam_raw ===> op =) size size"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
365  | 
by (simp_all add: size_rsp)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
366  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
367  | 
lemma [quot_preserve]:  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
368  | 
"(rep_lam ---> id) size = size"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
369  | 
by (simp_all add: size_lam_def)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
370  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
371  | 
instance  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
372  | 
by default  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
373  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
374  | 
end  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
375  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
376  | 
lemmas size_lam[simp] =  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
377  | 
lam_raw.size(4)[quot_lifted]  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
378  | 
lam_raw.size(5)[quot_lifted]  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
379  | 
lam_raw.size(6)[quot_lifted]  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
380  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
381  | 
(* is this needed? *)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
382  | 
lemma [measure_function]:  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
383  | 
"is_measure (size::lam\<Rightarrow>nat)"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
384  | 
by (rule is_measure_trivial)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
385  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
386  | 
section {* Matching *}
 | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
387  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
388  | 
definition  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
389  | 
  MATCH :: "('c::pt \<Rightarrow> (bool * 'a::pt * 'b::pt)) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
390  | 
where  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
391  | 
"MATCH M d x \<equiv> if (\<exists>!r. \<exists>q. M q = (True, x, r)) then (THE r. \<exists>q. M q = (True, x, r)) else d"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
392  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
393  | 
(*  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
394  | 
lemma MATCH_eqvt:  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
395  | 
shows "p \<bullet> (MATCH M d x) = MATCH (p \<bullet> M) (p \<bullet> d) (p \<bullet> x)"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
396  | 
unfolding MATCH_def  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
397  | 
apply(perm_simp the_eqvt)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
398  | 
apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
 | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
399  | 
apply(simp)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
400  | 
thm eqvts_raw  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
401  | 
apply(subst if_eqvt)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
402  | 
apply(subst ex1_eqvt)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
403  | 
apply(subst permute_fun_def)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
404  | 
apply(subst ex_eqvt)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
405  | 
apply(subst permute_fun_def)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
406  | 
apply(subst eq_eqvt)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
407  | 
apply(subst permute_fun_app_eq[where f="M"])  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
408  | 
apply(simp only: permute_minus_cancel)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
409  | 
apply(subst permute_prod.simps)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
410  | 
apply(subst permute_prod.simps)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
411  | 
apply(simp only: permute_minus_cancel)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
412  | 
apply(simp only: permute_bool_def)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
413  | 
apply(simp)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
414  | 
apply(subst ex1_eqvt)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
415  | 
apply(subst permute_fun_def)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
416  | 
apply(subst ex_eqvt)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
417  | 
apply(subst permute_fun_def)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
418  | 
apply(subst eq_eqvt)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
419  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
420  | 
apply(simp only: eqvts)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
421  | 
apply(simp)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
422  | 
apply(subgoal_tac "(p \<bullet> (\<exists>!r. \<exists>q. M q = (True, x, r))) = (\<exists>!r. \<exists>q. (p \<bullet> M) q = (True, p \<bullet> x, r))")  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
423  | 
apply(drule sym)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
424  | 
apply(simp)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
425  | 
apply(rule impI)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
426  | 
apply(simp add: perm_bool)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
427  | 
apply(rule trans)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
428  | 
apply(rule pt_the_eqvt[OF pta at])  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
429  | 
apply(assumption)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
430  | 
apply(simp add: pt_ex_eqvt[OF pt at])  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
431  | 
apply(simp add: pt_eq_eqvt[OF ptb at])  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
432  | 
apply(rule cheat)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
433  | 
apply(rule trans)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
434  | 
apply(rule pt_ex1_eqvt)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
435  | 
apply(rule pta)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
436  | 
apply(rule at)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
437  | 
apply(simp add: pt_ex_eqvt[OF pt at])  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
438  | 
apply(simp add: pt_eq_eqvt[OF ptb at])  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
439  | 
apply(subst pt_pi_rev[OF pta at])  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
440  | 
apply(subst pt_fun_app_eq[OF pt at])  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
441  | 
apply(subst pt_pi_rev[OF pt at])  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
442  | 
apply(simp)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
443  | 
done  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
444  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
445  | 
lemma MATCH_cng:  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
446  | 
assumes a: "M1 = M2" "d1 = d2"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
447  | 
shows "MATCH M1 d1 x = MATCH M2 d2 x"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
448  | 
using a by simp  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
449  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
450  | 
lemma MATCH_eq:  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
451  | 
assumes a: "t = l x" "G x" "\<And>x'. t = l x' \<Longrightarrow> G x' \<Longrightarrow> r x' = r x"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
452  | 
shows "MATCH (\<lambda>x. (G x, l x, r x)) d t = r x"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
453  | 
using a  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
454  | 
unfolding MATCH_def  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
455  | 
apply(subst if_P)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
456  | 
apply(rule_tac a="r x" in ex1I)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
457  | 
apply(rule_tac x="x" in exI)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
458  | 
apply(blast)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
459  | 
apply(erule exE)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
460  | 
apply(drule_tac x="q" in meta_spec)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
461  | 
apply(auto)[1]  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
462  | 
apply(rule the_equality)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
463  | 
apply(blast)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
464  | 
apply(erule exE)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
465  | 
apply(drule_tac x="q" in meta_spec)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
466  | 
apply(auto)[1]  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
467  | 
done  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
468  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
469  | 
lemma MATCH_eq2:  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
470  | 
assumes a: "t = l x1 x2" "G x1 x2" "\<And>x1' x2'. t = l x1' x2' \<Longrightarrow> G x1' x2' \<Longrightarrow> r x1' x2' = r x1 x2"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
471  | 
shows "MATCH (\<lambda>(x1,x2). (G x1 x2, l x1 x2, r x1 x2)) d t = r x1 x2"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
472  | 
sorry  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
473  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
474  | 
lemma MATCH_neq:  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
475  | 
assumes a: "\<And>x. t = l x \<Longrightarrow> G x \<Longrightarrow> False"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
476  | 
shows "MATCH (\<lambda>x. (G x, l x, r x)) d t = d"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
477  | 
using a  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
478  | 
unfolding MATCH_def  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
479  | 
apply(subst if_not_P)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
480  | 
apply(blast)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
481  | 
apply(rule refl)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
482  | 
done  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
483  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
484  | 
lemma MATCH_neq2:  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
485  | 
assumes a: "\<And>x1 x2. t = l x1 x2 \<Longrightarrow> G x1 x2 \<Longrightarrow> False"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
486  | 
shows "MATCH (\<lambda>(x1,x2). (G x1 x2, l x1 x2, r x1 x2)) d t = d"  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
487  | 
using a  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
488  | 
unfolding MATCH_def  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
489  | 
apply(subst if_not_P)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
490  | 
apply(auto)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
491  | 
done  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
492  | 
*)  | 
| 
 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1797 
diff
changeset
 | 
493  | 
|
| 1594 | 494  | 
end  | 
495  | 
||
496  | 
||
497  |