author | Christian Urban <urbanc@in.tum.de> |
Wed, 14 Apr 2010 16:05:58 +0200 | |
changeset 1835 | 636de31888a6 |
parent 1833 | 2050b5723c04 |
child 1845 | b7423c6b5564 |
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 |
||
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
130 |
inductive |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
131 |
typing' :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ \<Turnstile> _ : _" [60,60,60] 60) |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
132 |
where |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
133 |
t'_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<Turnstile> Var x : T" |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
134 |
| t'_App[intro]: "\<lbrakk>\<Gamma> \<Turnstile> t1 : T1\<rightarrow>T2 \<and> \<Gamma> \<Turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<Turnstile> App t1 t2 : T2" |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
135 |
| t'_Lam[intro]: "\<lbrakk>atom x\<sharp>\<Gamma>;(x,T1)#\<Gamma> \<Turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<Turnstile> Lam x t : T1\<rightarrow>T2" |
1828 | 136 |
|
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
137 |
inductive |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
138 |
typing2' :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ 2\<Turnstile> _ : _" [60,60,60] 60) |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
139 |
where |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
140 |
t2'_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> 2\<Turnstile> Var x : T" |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
141 |
| t2'_App[intro]: "\<lbrakk>\<Gamma> 2\<Turnstile> t1 : T1\<rightarrow>T2 \<or> \<Gamma> 2\<Turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> 2\<Turnstile> App t1 t2 : T2" |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
142 |
| t2'_Lam[intro]: "\<lbrakk>atom x\<sharp>\<Gamma>;(x,T1)#\<Gamma> 2\<Turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> 2\<Turnstile> Lam x t : T1\<rightarrow>T2" |
1831
16653e702d89
first working version of the automatic equivariance procedure
Christian Urban <urbanc@in.tum.de>
parents:
1828
diff
changeset
|
143 |
|
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
144 |
inductive |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
145 |
typing'' :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ |\<Turnstile> _ : _" [60,60,60] 60) |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
146 |
and valid' :: "(name\<times>ty) list \<Rightarrow> bool" |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
147 |
where |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
148 |
v1[intro]: "valid' []" |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
149 |
| v2[intro]: "\<lbrakk>valid' \<Gamma>;atom x\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid' ((x,T)#\<Gamma>)" |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
150 |
| t'_Var[intro]: "\<lbrakk>valid' \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> |\<Turnstile> Var x : T" |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
151 |
| t'_App[intro]: "\<lbrakk>\<Gamma> |\<Turnstile> t1 : T1\<rightarrow>T2; \<Gamma> |\<Turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> |\<Turnstile> App t1 t2 : T2" |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
152 |
| t'_Lam[intro]: "\<lbrakk>atom x\<sharp>\<Gamma>;(x,T1)#\<Gamma> |\<Turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> |\<Turnstile> Lam x t : T1\<rightarrow>T2" |
1831
16653e702d89
first working version of the automatic equivariance procedure
Christian Urban <urbanc@in.tum.de>
parents:
1828
diff
changeset
|
153 |
|
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
154 |
use "../../Nominal-General/nominal_eqvt.ML" |
1816
56cebe7f8e24
some small tunings (incompleted work in Lambda.thy)
Christian Urban <urbanc@in.tum.de>
parents:
1814
diff
changeset
|
155 |
|
1831
16653e702d89
first working version of the automatic equivariance procedure
Christian Urban <urbanc@in.tum.de>
parents:
1828
diff
changeset
|
156 |
equivariance valid |
16653e702d89
first working version of the automatic equivariance procedure
Christian Urban <urbanc@in.tum.de>
parents:
1828
diff
changeset
|
157 |
equivariance typing |
1816
56cebe7f8e24
some small tunings (incompleted work in Lambda.thy)
Christian Urban <urbanc@in.tum.de>
parents:
1814
diff
changeset
|
158 |
|
1831
16653e702d89
first working version of the automatic equivariance procedure
Christian Urban <urbanc@in.tum.de>
parents:
1828
diff
changeset
|
159 |
thm valid.eqvt |
16653e702d89
first working version of the automatic equivariance procedure
Christian Urban <urbanc@in.tum.de>
parents:
1828
diff
changeset
|
160 |
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
|
161 |
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
|
162 |
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
|
163 |
|
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
164 |
equivariance typing' |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
165 |
equivariance typing2' |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
166 |
equivariance typing'' |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
167 |
|
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
168 |
ML {* |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
169 |
fun mk_minus p = |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
170 |
Const (@{const_name "uminus"}, @{typ "perm => perm"}) $ p |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
171 |
*} |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
172 |
|
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
173 |
inductive |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
174 |
tt :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool)" |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
175 |
for r :: "('a \<Rightarrow> 'a \<Rightarrow> bool)" |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
176 |
where |
1835
636de31888a6
tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents:
1833
diff
changeset
|
177 |
aa: "tt r a a" |
636de31888a6
tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents:
1833
diff
changeset
|
178 |
| bb: "tt r a b ==> tt r a c" |
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
179 |
|
1835
636de31888a6
tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents:
1833
diff
changeset
|
180 |
(* PROBLEM: derived eqvt-theorem does not conform with [eqvt] |
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
181 |
equivariance tt |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
182 |
*) |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
1831
diff
changeset
|
183 |
|
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
|
184 |
|
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
|
185 |
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
|
186 |
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
|
187 |
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
|
188 |
"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
|
189 |
| "\<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
|
190 |
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
|
191 |
| "\<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
|
192 |
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
|
193 |
|
1831
16653e702d89
first working version of the automatic equivariance procedure
Christian Urban <urbanc@in.tum.de>
parents:
1828
diff
changeset
|
194 |
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
|
195 |
(*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
|
196 |
(*equivariance alpha_lam_raw'*) |
16653e702d89
first working version of the automatic equivariance procedure
Christian Urban <urbanc@in.tum.de>
parents:
1828
diff
changeset
|
197 |
|
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
|
198 |
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
|
199 |
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
|
200 |
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
|
201 |
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
|
202 |
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
|
203 |
oops |
1805 | 204 |
|
205 |
||
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
|
206 |
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
|
207 |
|
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
|
208 |
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
|
209 |
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
|
210 |
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
|
211 |
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
|
212 |
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
|
213 |
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
|
214 |
|
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
|
215 |
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
|
216 |
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
|
217 |
|
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
|
218 |
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
|
219 |
"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
|
220 |
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
|
221 |
"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
|
222 |
|
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
|
223 |
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
|
224 |
"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
|
225 |
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
|
226 |
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
|
227 |
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
|
228 |
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
|
229 |
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
|
230 |
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
|
231 |
|
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
|
232 |
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
|
233 |
"(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
|
234 |
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
|
235 |
|
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
|
236 |
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
|
237 |
"(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
|
238 |
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
|
239 |
|
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
|
240 |
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
|
241 |
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
|
242 |
|
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
|
243 |
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
|
244 |
|
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
|
245 |
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
|
246 |
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
|
247 |
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
|
248 |
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
|
249 |
|
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
|
250 |
(* 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
|
251 |
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
|
252 |
"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
|
253 |
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
|
254 |
|
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
|
255 |
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
|
256 |
|
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
|
257 |
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
|
258 |
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
|
259 |
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
|
260 |
"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
|
261 |
|
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
|
262 |
(* |
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
|
263 |
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
|
264 |
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
|
265 |
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
|
266 |
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
|
267 |
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
|
268 |
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
|
269 |
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
|
270 |
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
|
271 |
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
|
272 |
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
|
273 |
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
|
274 |
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
|
275 |
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
|
276 |
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
|
277 |
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
|
278 |
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
|
279 |
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
|
280 |
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
|
281 |
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
|
282 |
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
|
283 |
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
|
284 |
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
|
285 |
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
|
286 |
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
|
287 |
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
|
288 |
|
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
|
289 |
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
|
290 |
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
|
291 |
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
|
292 |
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
|
293 |
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
|
294 |
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
|
295 |
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
|
296 |
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
|
297 |
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
|
298 |
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
|
299 |
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
|
300 |
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
|
301 |
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
|
302 |
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
|
303 |
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
|
304 |
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
|
305 |
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
|
306 |
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
|
307 |
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
|
308 |
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
|
309 |
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
|
310 |
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
|
311 |
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
|
312 |
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
|
313 |
|
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
|
314 |
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
|
315 |
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
|
316 |
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
|
317 |
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
|
318 |
|
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
|
319 |
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
|
320 |
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
|
321 |
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
|
322 |
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
|
323 |
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
|
324 |
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
|
325 |
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
|
326 |
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
|
327 |
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
|
328 |
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
|
329 |
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
|
330 |
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
|
331 |
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
|
332 |
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
|
333 |
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
|
334 |
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
|
335 |
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
|
336 |
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
|
337 |
|
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 |
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
|
339 |
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
|
340 |
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
|
341 |
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
|
342 |
|
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 |
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
|
344 |
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
|
345 |
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
|
346 |
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
|
347 |
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
|
348 |
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
|
349 |
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
|
350 |
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
|
351 |
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
|
352 |
|
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 |
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
|
354 |
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
|
355 |
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
|
356 |
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
|
357 |
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
|
358 |
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
|
359 |
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
|
360 |
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
|
361 |
*) |
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 |
|
1594 | 363 |
end |
364 |
||
365 |
||
366 |