| author | Christian Urban <urbanc@in.tum.de> | 
| Wed, 24 Nov 2010 02:36:21 +0000 | |
| changeset 2581 | 3696659358c8 | 
| parent 2559 | add799cf0817 | 
| child 2617 | e44551d067e6 | 
| 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  | 
| 
2454
 
9ffee4eb1ae1
renamed NewParser to Nominal2
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2442 
diff
changeset
 | 
2  | 
imports "../Nominal2"  | 
| 1594 | 3  | 
begin  | 
4  | 
||
5  | 
atom_decl name  | 
|
| 
2440
 
0a36825b16c1
"isabelle make test" makes all major examples....they work up to supp theorems (excluding)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
6  | 
declare [[STEPS = 100]]  | 
| 1594 | 7  | 
|
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
8  | 
nominal_datatype lam =  | 
| 
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
 | 
9  | 
Var "name"  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
10  | 
| App "lam" "lam"  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
11  | 
| Lam x::"name" l::"lam" bind x in l  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2425 
diff
changeset
 | 
12  | 
|
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
13  | 
thm lam.distinct  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
14  | 
thm lam.induct  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
15  | 
thm lam.exhaust  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
16  | 
thm lam.fv_defs  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
17  | 
thm lam.bn_defs  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
18  | 
thm lam.perm_simps  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
19  | 
thm lam.eq_iff  | 
| 
2442
 
1f9360daf6e1
make copies of the "old" files
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2440 
diff
changeset
 | 
20  | 
thm lam.eq_iff[folded alphas]  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
21  | 
thm lam.fv_bn_eqvt  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
22  | 
thm lam.size_eqvt  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2425 
diff
changeset
 | 
23  | 
|
| 
2424
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2311 
diff
changeset
 | 
24  | 
|
| 
2497
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
25  | 
section {* Strong Exhausts Lemma and Strong Induction Lemma via Induct_schema *}
 | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
26  | 
|
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
27  | 
lemma strong_exhaust:  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
28  | 
fixes c::"'a::fs"  | 
| 
2503
 
cc5d23547341
simplified exhaust proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2497 
diff
changeset
 | 
29  | 
assumes "\<And>name. \<lbrakk>y = Var name\<rbrakk> \<Longrightarrow> P"  | 
| 
 
cc5d23547341
simplified exhaust proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2497 
diff
changeset
 | 
30  | 
and "\<And>lam1 lam2. \<lbrakk>y = App lam1 lam2\<rbrakk> \<Longrightarrow> P"  | 
| 
 
cc5d23547341
simplified exhaust proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2497 
diff
changeset
 | 
31  | 
and "\<And>name lam. \<lbrakk>atom name \<sharp> c; y = Lam name lam\<rbrakk> \<Longrightarrow> P"  | 
| 
2497
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
32  | 
shows "P"  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
33  | 
apply(rule_tac y="y" in lam.exhaust)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
34  | 
apply(rule assms(1))  | 
| 
2503
 
cc5d23547341
simplified exhaust proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2497 
diff
changeset
 | 
35  | 
apply(assumption)  | 
| 
2497
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
36  | 
apply(rule assms(2))  | 
| 
2503
 
cc5d23547341
simplified exhaust proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2497 
diff
changeset
 | 
37  | 
apply(assumption)  | 
| 
2497
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
38  | 
apply(subgoal_tac "\<exists>q. (q \<bullet> (atom name)) \<sharp> c \<and> supp (Lam name lam) \<sharp>* q")  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
39  | 
apply(erule exE)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
40  | 
apply(erule conjE)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
41  | 
apply(rule assms(3))  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
42  | 
apply(simp add: atom_eqvt)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
43  | 
apply(clarify)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
44  | 
apply(drule supp_perm_eq[symmetric])  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
45  | 
apply(simp)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
46  | 
apply(rule at_set_avoiding2_atom)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
47  | 
apply(simp add: finite_supp)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
48  | 
apply(simp add: finite_supp)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
49  | 
apply(simp add: lam.fresh)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
50  | 
done  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
51  | 
|
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
52  | 
|
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
53  | 
lemma strong_induct:  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
54  | 
fixes c::"'a::fs"  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
55  | 
assumes a1: "\<And>name c. P c (Var name)"  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
56  | 
and a2: "\<And>lam1 lam2 c. \<lbrakk>\<And>d. P d lam1; \<And>d. P d lam2\<rbrakk> \<Longrightarrow> P c (App lam1 lam2)"  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
57  | 
and a3: "\<And>name lam c. \<lbrakk>atom name \<sharp> c; \<And>d. P d lam\<rbrakk> \<Longrightarrow> P c (Lam name lam)"  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
58  | 
shows "P c lam"  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
59  | 
using assms  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
60  | 
apply(induction_schema)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
61  | 
apply(rule_tac y="lam" in strong_exhaust)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
62  | 
apply(blast)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
63  | 
apply(blast)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
64  | 
apply(blast)  | 
| 
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
65  | 
apply(relation "measure (\<lambda>(x,y). size y)")  | 
| 
2503
 
cc5d23547341
simplified exhaust proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2497 
diff
changeset
 | 
66  | 
apply(simp_all add: lam.size)  | 
| 
2497
 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2454 
diff
changeset
 | 
67  | 
done  | 
| 
2082
 
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2041 
diff
changeset
 | 
68  | 
|
| 
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
 | 
69  | 
section {* Strong Induction Principles*}
 | 
| 1594 | 70  | 
|
| 
2041
 
3842464ee03b
Move 2 more to NewParser
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
1954 
diff
changeset
 | 
71  | 
(*  | 
| 
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
 | 
72  | 
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
 | 
73  | 
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
 | 
74  | 
*)  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
75  | 
(*  | 
| 1594 | 76  | 
lemma  | 
77  | 
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
 | 
78  | 
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
 | 
79  | 
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
 | 
80  | 
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
 | 
81  | 
shows "P c lam"  | 
| 1594 | 82  | 
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
 | 
83  | 
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
 | 
84  | 
apply(induct lam arbitrary: c rule: lam.induct)  | 
| 1805 | 85  | 
apply(perm_simp)  | 
| 1594 | 86  | 
apply(rule a1)  | 
| 1805 | 87  | 
apply(perm_simp)  | 
| 1594 | 88  | 
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
 | 
89  | 
apply(assumption)  | 
| 1594 | 90  | 
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
 | 
91  | 
apply(subgoal_tac "\<exists>new::name. (atom new) \<sharp> (c, Lam (p \<bullet> name) (p \<bullet> lam))")  | 
| 1594 | 92  | 
defer  | 
93  | 
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
 | 
94  | 
apply(rule_tac X="supp (c, Lam (p \<bullet> name) (p \<bullet> lam))" in obtain_at_base)  | 
| 1594 | 95  | 
apply(simp add: supp_Pair finite_supp)  | 
96  | 
apply(blast)  | 
|
97  | 
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
 | 
98  | 
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
 | 
99  | 
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
 | 
100  | 
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
 | 
101  | 
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
 | 
102  | 
apply(subst (2) lam.perm)  | 
| 1594 | 103  | 
apply(rule flip_fresh_fresh)  | 
104  | 
apply(simp add: fresh_def)  | 
|
105  | 
apply(simp only: supp_fn')  | 
|
106  | 
apply(simp)  | 
|
107  | 
apply(simp add: fresh_Pair)  | 
|
108  | 
apply(simp)  | 
|
109  | 
apply(rule a3)  | 
|
110  | 
apply(simp add: fresh_Pair)  | 
|
111  | 
apply(drule_tac x="((p \<bullet> name) \<leftrightarrow> new) + p" in meta_spec)  | 
|
112  | 
apply(simp)  | 
|
113  | 
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
 | 
114  | 
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
 | 
115  | 
then show "P c lam" by simp  | 
| 1594 | 116  | 
qed  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
117  | 
*)  | 
| 
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
 | 
118  | 
(*  | 
| 
 
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
 | 
119  | 
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
 | 
120  | 
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
 | 
121  | 
*)  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
122  | 
(*  | 
| 1594 | 123  | 
lemma  | 
124  | 
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
 | 
125  | 
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
 | 
126  | 
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
 | 
127  | 
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
 | 
128  | 
shows "P c lam"  | 
| 1594 | 129  | 
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
 | 
130  | 
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
 | 
131  | 
apply(induct lam arbitrary: c rule: lam.induct)  | 
| 1805 | 132  | 
apply(perm_simp)  | 
| 1594 | 133  | 
apply(rule a1)  | 
| 1805 | 134  | 
apply(perm_simp)  | 
| 1594 | 135  | 
apply(rule a2)  | 
136  | 
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
 | 
137  | 
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
 | 
138  | 
    apply(subgoal_tac "\<exists>q. (q \<bullet> {p \<bullet> atom name}) \<sharp>* c \<and> supp (p \<bullet> Lam name lam) \<sharp>* q")
 | 
| 1594 | 139  | 
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
 | 
140  | 
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
 | 
141  | 
s="q \<bullet> p \<bullet> Lam name lam" in subst)  | 
| 1594 | 142  | 
defer  | 
| 1805 | 143  | 
apply(simp)  | 
| 1594 | 144  | 
apply(rule a3)  | 
145  | 
apply(simp add: eqvts fresh_star_def)  | 
|
146  | 
apply(drule_tac x="q + p" in meta_spec)  | 
|
147  | 
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
 | 
148  | 
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
 | 
149  | 
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
 | 
150  | 
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
 | 
151  | 
apply(simp add: finite_supp)  | 
| 1805 | 152  | 
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
 | 
153  | 
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
 | 
154  | 
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
 | 
155  | 
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
 | 
156  | 
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
 | 
157  | 
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
 | 
158  | 
then show "P c lam" by simp  | 
| 1594 | 159  | 
qed  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
160  | 
*)  | 
| 1594 | 161  | 
|
| 1805 | 162  | 
section {* Typing *}
 | 
163  | 
||
164  | 
nominal_datatype ty =  | 
|
165  | 
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
 | 
166  | 
| 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
 | 
167  | 
|
| 
 
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
 | 
168  | 
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
 | 
169  | 
 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
 | 
170  | 
|
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
171  | 
(*  | 
| 
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
 | 
172  | 
declare ty.perm[eqvt]  | 
| 1805 | 173  | 
|
174  | 
inductive  | 
|
175  | 
valid :: "(name \<times> ty) list \<Rightarrow> bool"  | 
|
176  | 
where  | 
|
177  | 
"valid []"  | 
|
178  | 
| "\<lbrakk>atom x \<sharp> Gamma; valid Gamma\<rbrakk> \<Longrightarrow> valid ((x, T)#Gamma)"  | 
|
179  | 
||
| 1828 | 180  | 
inductive  | 
181  | 
  typing :: "(name\<times>ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
 | 
|
182  | 
where  | 
|
183  | 
t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"  | 
|
| 1947 | 184  | 
| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"  | 
| 1828 | 185  | 
| t_Lam[intro]: "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam x t : T1 \<rightarrow> T2"  | 
186  | 
||
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
187  | 
equivariance valid  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
188  | 
equivariance typing  | 
| 
1816
 
56cebe7f8e24
some small tunings (incompleted work in Lambda.thy)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1814 
diff
changeset
 | 
189  | 
|
| 
1831
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
190  | 
thm valid.eqvt  | 
| 
 
16653e702d89
first working version of the automatic equivariance procedure
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1828 
diff
changeset
 | 
191  | 
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
 | 
192  | 
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
 | 
193  | 
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
 | 
194  | 
|
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
195  | 
thm typing.induct[of "\<Gamma>" "t" "T", no_vars]  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
196  | 
|
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
197  | 
lemma  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
198  | 
fixes c::"'a::fs"  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
199  | 
assumes a: "\<Gamma> \<turnstile> t : T"  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
200  | 
and a1: "\<And>\<Gamma> x T c. \<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P c \<Gamma> (Var x) T"  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
201  | 
and a2: "\<And>\<Gamma> t1 T1 T2 t2 c. \<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<And>d. P d \<Gamma> t1 T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1; \<And>d. P d \<Gamma> t2 T1\<rbrakk>  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
202  | 
\<Longrightarrow> P c \<Gamma> (App t1 t2) T2"  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
203  | 
and a3: "\<And>x \<Gamma> T1 t T2 c. \<lbrakk>atom x \<sharp> c; atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2; \<And>d. P d ((x, T1) # \<Gamma>) t T2\<rbrakk>  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
204  | 
\<Longrightarrow> P c \<Gamma> (Lam x t) T1 \<rightarrow> T2"  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
205  | 
shows "P c \<Gamma> t T"  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
206  | 
proof -  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
207  | 
from a have "\<And>p c. P c (p \<bullet> \<Gamma>) (p \<bullet> t) (p \<bullet> T)"  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
208  | 
proof (induct)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
209  | 
case (t_Var \<Gamma> x T p c)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
210  | 
then show ?case  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
211  | 
apply -  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
212  | 
apply(perm_strict_simp)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
213  | 
apply(rule a1)  | 
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
214  | 
apply(drule_tac p="p" in permute_boolI)  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
215  | 
apply(perm_strict_simp add: permute_minus_cancel)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
216  | 
apply(assumption)  | 
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
217  | 
apply(rotate_tac 1)  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
218  | 
apply(drule_tac p="p" in permute_boolI)  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
219  | 
apply(perm_strict_simp add: permute_minus_cancel)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
220  | 
apply(assumption)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
221  | 
done  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
222  | 
next  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
223  | 
case (t_App \<Gamma> t1 T1 T2 t2 p c)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
224  | 
then show ?case  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
225  | 
apply -  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
226  | 
apply(perm_strict_simp)  | 
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
227  | 
apply(rule a2)  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
228  | 
apply(drule_tac p="p" in permute_boolI)  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
229  | 
apply(perm_strict_simp add: permute_minus_cancel)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
230  | 
apply(assumption)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
231  | 
apply(assumption)  | 
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
232  | 
apply(rotate_tac 2)  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
233  | 
apply(drule_tac p="p" in permute_boolI)  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
234  | 
apply(perm_strict_simp add: permute_minus_cancel)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
235  | 
apply(assumption)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
236  | 
apply(assumption)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
237  | 
done  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
238  | 
next  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
239  | 
case (t_Lam x \<Gamma> T1 t T2 p c)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
240  | 
then show ?case  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
241  | 
apply -  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
242  | 
      apply(subgoal_tac "\<exists>q. (q \<bullet> {p \<bullet> atom x}) \<sharp>* c \<and> 
 | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
243  | 
supp (p \<bullet> \<Gamma>, p \<bullet> Lam x t, p \<bullet> (T1 \<rightarrow> T2)) \<sharp>* q")  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
244  | 
apply(erule exE)  | 
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
245  | 
apply(rule_tac t="p \<bullet> \<Gamma>" and s="(q + p) \<bullet> \<Gamma>" in subst)  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
246  | 
apply(simp only: permute_plus)  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
247  | 
apply(rule supp_perm_eq)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
248  | 
apply(simp add: supp_Pair fresh_star_union)  | 
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
249  | 
apply(rule_tac t="p \<bullet> Lam x t" and s="(q + p) \<bullet> Lam x t" in subst)  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
250  | 
apply(simp only: permute_plus)  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
251  | 
apply(rule supp_perm_eq)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
252  | 
apply(simp add: supp_Pair fresh_star_union)  | 
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
253  | 
apply(rule_tac t="p \<bullet> (T1 \<rightarrow> T2)" and s="(q + p) \<bullet> (T1 \<rightarrow> T2)" in subst)  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
254  | 
apply(simp only: permute_plus)  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
255  | 
apply(rule supp_perm_eq)  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
256  | 
apply(simp add: supp_Pair fresh_star_union)  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
257  | 
apply(simp (no_asm) only: eqvts)  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
258  | 
apply(rule a3)  | 
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
259  | 
apply(simp only: eqvts permute_plus)  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
260  | 
apply(simp add: fresh_star_def)  | 
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
261  | 
apply(drule_tac p="q + p" in permute_boolI)  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
262  | 
apply(perm_strict_simp add: permute_minus_cancel)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
263  | 
apply(assumption)  | 
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
264  | 
apply(rotate_tac 1)  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
265  | 
apply(drule_tac p="q + p" in permute_boolI)  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
266  | 
apply(perm_strict_simp add: permute_minus_cancel)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
267  | 
apply(assumption)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
268  | 
apply(drule_tac x="d" in meta_spec)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
269  | 
apply(drule_tac x="q + p" in meta_spec)  | 
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
270  | 
apply(perm_strict_simp add: permute_minus_cancel)  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
271  | 
apply(assumption)  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
272  | 
apply(rule at_set_avoiding2)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
273  | 
apply(simp add: finite_supp)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
274  | 
apply(simp add: finite_supp)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
275  | 
apply(simp add: finite_supp)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
276  | 
apply(rule_tac p="-p" in permute_boolE)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
277  | 
apply(perm_strict_simp add: permute_minus_cancel)  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
278  | 
--"supplied by the user"  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
279  | 
apply(simp add: fresh_star_prod)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
280  | 
apply(simp add: fresh_star_def)  | 
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
281  | 
sorry  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
282  | 
qed  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
283  | 
then have "P c (0 \<bullet> \<Gamma>) (0 \<bullet> t) (0 \<bullet> T)" .  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
284  | 
then show "P c \<Gamma> t T" by simp  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
285  | 
qed  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
286  | 
|
| 
1833
 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1831 
diff
changeset
 | 
287  | 
*)  | 
| 
 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1831 
diff
changeset
 | 
288  | 
|
| 
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
 | 
289  | 
|
| 
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
 | 
290  | 
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
 | 
291  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-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  | 
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
 | 
293  | 
  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
 | 
294  | 
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
 | 
295  | 
"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
 | 
296  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-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  | 
(*  | 
| 
 
78fdc6b36a1c
changed the eqvt-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  | 
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
 | 
299  | 
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
 | 
300  | 
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
 | 
301  | 
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
 | 
302  | 
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
 | 
303  | 
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
 | 
304  | 
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
 | 
305  | 
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
 | 
306  | 
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
 | 
307  | 
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
 | 
308  | 
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
 | 
309  | 
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
 | 
310  | 
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
 | 
311  | 
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
 | 
312  | 
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
 | 
313  | 
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
 | 
314  | 
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
 | 
315  | 
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
 | 
316  | 
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
 | 
317  | 
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
 | 
318  | 
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
 | 
319  | 
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
 | 
320  | 
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
 | 
321  | 
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
 | 
322  | 
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
 | 
323  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-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(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
 | 
325  | 
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
 | 
326  | 
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
 | 
327  | 
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
 | 
328  | 
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
 | 
329  | 
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
 | 
330  | 
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
 | 
331  | 
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
 | 
332  | 
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
 | 
333  | 
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
 | 
334  | 
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
 | 
335  | 
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
 | 
336  | 
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
 | 
337  | 
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
 | 
338  | 
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
 | 
339  | 
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
 | 
340  | 
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
 | 
341  | 
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
 | 
342  | 
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
 | 
343  | 
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
 | 
344  | 
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
 | 
345  | 
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
 | 
346  | 
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
 | 
347  | 
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
 | 
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  | 
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
 | 
350  | 
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
 | 
351  | 
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
 | 
352  | 
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
 | 
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 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
 | 
355  | 
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
 | 
356  | 
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
 | 
357  | 
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
 | 
358  | 
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
 | 
359  | 
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
 | 
360  | 
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
 | 
361  | 
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
 | 
362  | 
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
 | 
363  | 
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
 | 
364  | 
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
 | 
365  | 
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
 | 
366  | 
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
 | 
367  | 
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
 | 
368  | 
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
 | 
369  | 
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
 | 
370  | 
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
 | 
371  | 
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
 | 
372  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-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  | 
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
 | 
374  | 
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
 | 
375  | 
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
 | 
376  | 
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
 | 
377  | 
|
| 
 
78fdc6b36a1c
changed the eqvt-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  | 
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
 | 
379  | 
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
 | 
380  | 
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
 | 
381  | 
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
 | 
382  | 
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
 | 
383  | 
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
 | 
384  | 
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
 | 
385  | 
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
 | 
386  | 
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
 | 
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  | 
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
 | 
389  | 
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
 | 
390  | 
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
 | 
391  | 
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
 | 
392  | 
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
 | 
393  | 
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
 | 
394  | 
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
 | 
395  | 
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
 | 
396  | 
*)  | 
| 
 
78fdc6b36a1c
changed the eqvt-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  | 
|
| 
1954
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
398  | 
ML {*
 | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
399  | 
fun mk_avoids ctxt params name set =  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
400  | 
let  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
401  | 
val (_, ctxt') = ProofContext.add_fixes  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
402  | 
(map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params) ctxt;  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
403  | 
fun mk s =  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
404  | 
let  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
405  | 
val t = Syntax.read_term ctxt' s;  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
406  | 
val t' = list_abs_free (params, t) |>  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
407  | 
funpow (length params) (fn Abs (_, _, t) => t)  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
408  | 
in (t', HOLogic.dest_setT (fastype_of t)) end  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
409  | 
handle TERM _ =>  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
410  | 
        error ("Expression " ^ quote s ^ " to be avoided in case " ^
 | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
411  | 
quote name ^ " is not a set type");  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
412  | 
fun add_set p [] = [p]  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
413  | 
| add_set (t, T) ((u, U) :: ps) =  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
414  | 
if T = U then  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
415  | 
let val S = HOLogic.mk_setT T  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
416  | 
            in (Const (@{const_name sup}, S --> S --> S) $ u $ t, T) :: ps
 | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
417  | 
end  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
418  | 
else (u, U) :: add_set (t, T) ps  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
419  | 
in  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
420  | 
(mk #> add_set) set  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
421  | 
end;  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
422  | 
*}  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
423  | 
|
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
424  | 
|
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
425  | 
ML {* 
 | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
426  | 
  writeln (commas (map (Syntax.string_of_term @{context} o fst) 
 | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
427  | 
    (mk_avoids @{context} [] "t_Var" "{x}" []))) 
 | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
428  | 
*}  | 
| 
 
23480003f9c5
some changes to the paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1950 
diff
changeset
 | 
429  | 
|
| 1947 | 430  | 
|
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
431  | 
ML {*
 | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
432  | 
|
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
433  | 
fun prove_strong_ind (pred_name, avoids) ctxt =  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
434  | 
Proof.theorem NONE (K I) [] ctxt  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
435  | 
|
| 2169 | 436  | 
local structure P = Parse and K = Keyword in  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
437  | 
|
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
438  | 
val _ =  | 
| 2169 | 439  | 
Outer_Syntax.local_theory_to_proof "nominal_inductive"  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
440  | 
"proves strong induction theorem for inductive predicate involving nominal datatypes" K.thy_goal  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
441  | 
(P.xname -- (Scan.optional (P.$$$ "avoids" |-- P.enum1 "|" (P.name --  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
442  | 
(P.$$$ ":" |-- P.and_list1 P.term))) []) >> prove_strong_ind)  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
443  | 
|
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
444  | 
end;  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
445  | 
|
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
446  | 
*}  | 
| 
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
447  | 
|
| 
1950
 
7de54c9f81ac
eliminated command so that all compiles
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1949 
diff
changeset
 | 
448  | 
(*  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
449  | 
nominal_inductive typing  | 
| 
1950
 
7de54c9f81ac
eliminated command so that all compiles
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1949 
diff
changeset
 | 
450  | 
*)  | 
| 
1949
 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1947 
diff
changeset
 | 
451  | 
|
| 2157 | 452  | 
(* Substitution *)  | 
| 
2159
 
ce00205e07ab
Single variable substitution
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2158 
diff
changeset
 | 
453  | 
|
| 2165 | 454  | 
primrec match_Var_raw where  | 
455  | 
"match_Var_raw (Var_raw x) = Some x"  | 
|
456  | 
| "match_Var_raw (App_raw x y) = None"  | 
|
457  | 
| "match_Var_raw (Lam_raw n t) = None"  | 
|
458  | 
||
459  | 
quotient_definition  | 
|
460  | 
"match_Var :: lam \<Rightarrow> name option"  | 
|
461  | 
is match_Var_raw  | 
|
462  | 
||
463  | 
lemma [quot_respect]: "(alpha_lam_raw ===> op =) match_Var_raw match_Var_raw"  | 
|
464  | 
apply rule  | 
|
| 
2559
 
add799cf0817
adapted to changes by Florian on the quotient package and removed local fix for function package
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2503 
diff
changeset
 | 
465  | 
apply (induct_tac x y rule: alpha_lam_raw.induct)  | 
| 2165 | 466  | 
apply simp_all  | 
467  | 
done  | 
|
468  | 
||
469  | 
lemmas match_Var_simps = match_Var_raw.simps[quot_lifted]  | 
|
470  | 
||
471  | 
primrec match_App_raw where  | 
|
472  | 
"match_App_raw (Var_raw x) = None"  | 
|
473  | 
| "match_App_raw (App_raw x y) = Some (x, y)"  | 
|
474  | 
| "match_App_raw (Lam_raw n t) = None"  | 
|
475  | 
||
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
476  | 
(*  | 
| 2165 | 477  | 
quotient_definition  | 
478  | 
"match_App :: lam \<Rightarrow> (lam \<times> lam) option"  | 
|
479  | 
is match_App_raw  | 
|
480  | 
||
481  | 
lemma [quot_respect]:  | 
|
482  | 
"(alpha_lam_raw ===> option_rel (prod_rel alpha_lam_raw alpha_lam_raw)) match_App_raw match_App_raw"  | 
|
483  | 
apply (intro fun_relI)  | 
|
484  | 
apply (induct_tac a b rule: alpha_lam_raw.induct)  | 
|
485  | 
apply simp_all  | 
|
486  | 
done  | 
|
487  | 
||
488  | 
lemmas match_App_simps = match_App_raw.simps[quot_lifted]  | 
|
489  | 
||
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
490  | 
definition new where  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
491  | 
"new (s :: 'a :: fs) = (THE x. \<forall>a \<in> supp s. atom x \<noteq> a)"  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
492  | 
|
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
493  | 
definition  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
494  | 
"match_Lam (S :: 'a :: fs) t = (if (\<exists>n s. (t = Lam n s)) then  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
495  | 
(let z = new (S, t) in Some (z, THE s. t = Lam z s)) else None)"  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
496  | 
|
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
497  | 
lemma lam_half_inj: "(Lam z s = Lam z sa) = (s = sa)"  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
498  | 
apply auto  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
499  | 
apply (simp only: lam.eq_iff alphas)  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
500  | 
apply clarify  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
501  | 
apply (simp add: eqvts)  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
502  | 
sorry  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
503  | 
|
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
504  | 
lemma match_Lam_simps:  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
505  | 
"match_Lam S (Var n) = None"  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
506  | 
"match_Lam S (App l r) = None"  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
507  | 
"z = new (S, (Lam z s)) \<Longrightarrow> match_Lam S (Lam z s) = Some (z, s)"  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
508  | 
apply (simp_all add: match_Lam_def)  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
509  | 
apply (simp add: lam_half_inj)  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
510  | 
apply auto  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
511  | 
done  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
512  | 
*)  | 
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
513  | 
(*  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
514  | 
lemma match_Lam_simps2:  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
515  | 
"atom n \<sharp> ((S :: 'a :: fs), Lam n s) \<Longrightarrow> match_Lam S (Lam n s) = Some (n, s)"  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
516  | 
apply (rule_tac t="Lam n s"  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
517  | 
and s="Lam (new (S, (Lam n s))) ((n \<leftrightarrow> (new (S, (Lam n s)))) \<bullet> s)" in subst)  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
518  | 
defer  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
519  | 
apply (subst match_Lam_simps(3))  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
520  | 
defer  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
521  | 
apply simp  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
522  | 
*)  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
523  | 
|
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
524  | 
(*primrec match_Lam_raw where  | 
| 2165 | 525  | 
"match_Lam_raw (S :: atom set) (Var_raw x) = None"  | 
526  | 
| "match_Lam_raw S (App_raw x y) = None"  | 
|
527  | 
| "match_Lam_raw S (Lam_raw n t) = (let z = new (S \<union> (fv_lam_raw t - {atom n})) in Some (z, (n \<leftrightarrow> z) \<bullet> t))"
 | 
|
528  | 
||
529  | 
quotient_definition  | 
|
530  | 
"match_Lam :: (atom set) \<Rightarrow> lam \<Rightarrow> (name \<times> lam) option"  | 
|
531  | 
is match_Lam_raw  | 
|
532  | 
||
| 
2172
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
533  | 
lemma swap_fresh:  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
534  | 
assumes a: "fv_lam_raw t \<sharp>* p"  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
535  | 
shows "alpha_lam_raw (p \<bullet> t) t"  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
536  | 
using a apply (induct t)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
537  | 
apply (simp add: supp_at_base fresh_star_def)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
538  | 
apply (rule alpha_lam_raw.intros)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
539  | 
apply (metis Rep_name_inverse atom_eqvt atom_name_def fresh_perm)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
540  | 
apply (simp)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
541  | 
apply (simp only: fresh_star_union)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
542  | 
apply clarify  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
543  | 
apply (rule alpha_lam_raw.intros)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
544  | 
apply simp  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
545  | 
apply simp  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
546  | 
apply simp  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
547  | 
apply (rule alpha_lam_raw.intros)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
548  | 
sorry  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
549  | 
|
| 2165 | 550  | 
lemma [quot_respect]:  | 
551  | 
"(op = ===> alpha_lam_raw ===> option_rel (prod_rel op = alpha_lam_raw)) match_Lam_raw match_Lam_raw"  | 
|
552  | 
proof (intro fun_relI, clarify)  | 
|
553  | 
fix S t s  | 
|
554  | 
assume a: "alpha_lam_raw t s"  | 
|
555  | 
show "option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S t) (match_Lam_raw S s)"  | 
|
556  | 
using a proof (induct t s rule: alpha_lam_raw.induct)  | 
|
557  | 
case goal1 show ?case by simp  | 
|
558  | 
next  | 
|
559  | 
case goal2 show ?case by simp  | 
|
560  | 
next  | 
|
561  | 
case (goal3 x t y s)  | 
|
562  | 
      then obtain p where "({atom x}, t) \<approx>gen (\<lambda>x1 x2. alpha_lam_raw x1 x2 \<and>
 | 
|
563  | 
option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S x1)  | 
|
564  | 
                               (match_Lam_raw S x2)) fv_lam_raw p ({atom y}, s)" ..
 | 
|
565  | 
then have  | 
|
566  | 
        c: "fv_lam_raw t - {atom x} = fv_lam_raw s - {atom y}" and
 | 
|
567  | 
        d: "(fv_lam_raw t - {atom x}) \<sharp>* p" and
 | 
|
568  | 
e: "alpha_lam_raw (p \<bullet> t) s" and  | 
|
569  | 
f: "option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S (p \<bullet> t)) (match_Lam_raw S s)" and  | 
|
570  | 
        g: "p \<bullet> {atom x} = {atom y}" unfolding alphas(1) by - (elim conjE, assumption)+
 | 
|
571  | 
      let ?z = "new (S \<union> (fv_lam_raw t - {atom x}))"
 | 
|
572  | 
      have h: "?z = new (S \<union> (fv_lam_raw s - {atom y}))" using c by simp
 | 
|
573  | 
show ?case  | 
|
574  | 
unfolding match_Lam_raw.simps Let_def option_rel.simps prod_rel.simps split_conv  | 
|
575  | 
proof  | 
|
576  | 
        show "?z = new (S \<union> (fv_lam_raw s - {atom y}))" by (fact h)
 | 
|
577  | 
next  | 
|
578  | 
have "atom y \<sharp> p" sorry  | 
|
579  | 
have "fv_lam_raw t \<sharp>* ((x \<leftrightarrow> y) \<bullet> p)" sorry  | 
|
| 
2172
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
580  | 
then have "alpha_lam_raw (((x \<leftrightarrow> y) \<bullet> p) \<bullet> t) t" using swap_fresh by auto  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
581  | 
then have "alpha_lam_raw (p \<bullet> t) ((x \<leftrightarrow> y) \<bullet> t)" sorry  | 
| 2165 | 582  | 
have "alpha_lam_raw t ((x \<leftrightarrow> y) \<bullet> s)" sorry  | 
583  | 
then have "alpha_lam_raw ((x \<leftrightarrow> ?z) \<bullet> t) ((y \<leftrightarrow> ?z) \<bullet> s)" using eqvts(15) sorry  | 
|
584  | 
        then show "alpha_lam_raw ((x \<leftrightarrow> new (S \<union> (fv_lam_raw t - {atom x}))) \<bullet> t)
 | 
|
585  | 
                  ((y \<leftrightarrow> new (S \<union> (fv_lam_raw s - {atom y}))) \<bullet> s)" unfolding h .
 | 
|
586  | 
qed  | 
|
587  | 
qed  | 
|
588  | 
qed  | 
|
589  | 
||
590  | 
lemmas match_Lam_simps = match_Lam_raw.simps[quot_lifted]  | 
|
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
591  | 
*)  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
592  | 
(*  | 
| 2165 | 593  | 
lemma app_some: "match_App x = Some (a, b) \<Longrightarrow> x = App a b"  | 
594  | 
by (induct x rule: lam.induct) (simp_all add: match_App_simps)  | 
|
595  | 
||
596  | 
lemma lam_some: "match_Lam S x = Some (z, s) \<Longrightarrow> x = Lam z s \<and> atom z \<sharp> S"  | 
|
597  | 
apply (induct x rule: lam.induct)  | 
|
598  | 
apply (simp_all add: match_Lam_simps)  | 
|
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
599  | 
apply (thin_tac "match_Lam S lam = Some (z, s) \<Longrightarrow> lam = Lam z s \<and> atom z \<sharp> S")  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
600  | 
apply (simp add: match_Lam_def)  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
601  | 
apply (subgoal_tac "\<exists>n s. Lam name lam = Lam n s")  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
602  | 
prefer 2  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
603  | 
apply auto[1]  | 
| 2165 | 604  | 
apply (simp add: Let_def)  | 
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
605  | 
apply (thin_tac "\<exists>n s. Lam name lam = Lam n s")  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
606  | 
apply clarify  | 
| 2165 | 607  | 
apply (rule conjI)  | 
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
608  | 
apply (rule_tac t="THE s. Lam name lam = Lam (new (S, Lam name lam)) s" and  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
609  | 
s="(name \<leftrightarrow> (new (S, Lam name lam))) \<bullet> lam" in subst)  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
610  | 
defer  | 
| 2165 | 611  | 
apply (simp add: lam.eq_iff)  | 
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
612  | 
apply (rule_tac x="(name \<leftrightarrow> (new (S, Lam name lam)))" in exI)  | 
| 2165 | 613  | 
apply (simp add: alphas)  | 
614  | 
apply (simp add: eqvts)  | 
|
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
615  | 
apply (rule conjI)  | 
| 
2172
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
616  | 
sorry  | 
| 2165 | 617  | 
|
618  | 
function subst where  | 
|
619  | 
"subst v s t = (  | 
|
620  | 
case match_Var t of Some n \<Rightarrow> if n = v then s else Var n | None \<Rightarrow>  | 
|
621  | 
case match_App t of Some (l, r) \<Rightarrow> App (subst v s l) (subst v s r) | None \<Rightarrow>  | 
|
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
622  | 
case match_Lam (v,s) t of Some (n, t) \<Rightarrow> Lam n (subst v s t) | None \<Rightarrow> undefined)"  | 
| 2165 | 623  | 
by pat_completeness auto  | 
624  | 
||
625  | 
termination apply (relation "measure (\<lambda>(_, _, t). size t)")  | 
|
| 
2172
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
626  | 
apply auto[1]  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
627  | 
apply (case_tac a) apply simp  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
628  | 
apply (frule lam_some) apply simp  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
629  | 
apply (case_tac a) apply simp  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
630  | 
apply (frule app_some) apply simp  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
631  | 
apply (case_tac a) apply simp  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
632  | 
apply (frule app_some) apply simp  | 
| 2165 | 633  | 
done  | 
634  | 
||
635  | 
lemmas lam_exhaust = lam_raw.exhaust[quot_lifted]  | 
|
636  | 
||
637  | 
lemma subst_eqvt:  | 
|
638  | 
"p \<bullet> (subst v s t) = subst (p \<bullet> v) (p \<bullet> s) (p \<bullet> t)"  | 
|
639  | 
proof (induct v s t rule: subst.induct)  | 
|
640  | 
case (1 v s t)  | 
|
641  | 
show ?case proof (cases t rule: lam_exhaust)  | 
|
642  | 
fix n  | 
|
643  | 
assume "t = Var n"  | 
|
644  | 
then show ?thesis by (simp add: match_Var_simps)  | 
|
645  | 
next  | 
|
646  | 
fix l r  | 
|
647  | 
assume "t = App l r"  | 
|
648  | 
then show ?thesis  | 
|
649  | 
apply (simp only:)  | 
|
650  | 
apply (subst subst.simps)  | 
|
651  | 
apply (subst match_Var_simps)  | 
|
652  | 
apply (simp only: option.cases)  | 
|
653  | 
apply (subst match_App_simps)  | 
|
654  | 
apply (simp only: option.cases)  | 
|
655  | 
apply (simp only: prod.cases)  | 
|
656  | 
apply (simp only: lam.perm)  | 
|
657  | 
apply (subst (3) subst.simps)  | 
|
658  | 
apply (subst match_Var_simps)  | 
|
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
659  | 
apply (simp only: option.cases)  | 
| 2165 | 660  | 
apply (subst match_App_simps)  | 
661  | 
apply (simp only: option.cases)  | 
|
662  | 
apply (simp only: prod.cases)  | 
|
663  | 
apply (subst 1(2)[of "(l, r)" "l" "r"])  | 
|
664  | 
apply (simp add: match_Var_simps)  | 
|
665  | 
apply (simp add: match_App_simps)  | 
|
666  | 
apply (rule refl)  | 
|
667  | 
apply (subst 1(3)[of "(l, r)" "l" "r"])  | 
|
668  | 
apply (simp add: match_Var_simps)  | 
|
669  | 
apply (simp add: match_App_simps)  | 
|
670  | 
apply (rule refl)  | 
|
671  | 
apply (rule refl)  | 
|
672  | 
done  | 
|
673  | 
next  | 
|
674  | 
fix n t'  | 
|
675  | 
assume "t = Lam n t'"  | 
|
676  | 
then show ?thesis  | 
|
677  | 
apply (simp only: )  | 
|
678  | 
apply (simp only: lam.perm)  | 
|
679  | 
apply (subst subst.simps)  | 
|
680  | 
apply (subst match_Var_simps)  | 
|
681  | 
apply (simp only: option.cases)  | 
|
682  | 
apply (subst match_App_simps)  | 
|
683  | 
apply (simp only: option.cases)  | 
|
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
684  | 
apply (rule_tac t="Lam n t'" and s="Lam (new ((v, s), Lam n t')) ((n \<leftrightarrow> new ((v, s), Lam n t')) \<bullet> t')" in subst)  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
685  | 
defer  | 
| 2165 | 686  | 
apply (subst match_Lam_simps)  | 
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
687  | 
defer  | 
| 2165 | 688  | 
apply (simp only: option.cases)  | 
689  | 
apply (simp only: prod.cases)  | 
|
690  | 
apply (subst (2) subst.simps)  | 
|
691  | 
apply (subst match_Var_simps)  | 
|
692  | 
apply (simp only: option.cases)  | 
|
693  | 
apply (subst match_App_simps)  | 
|
694  | 
apply (simp only: option.cases)  | 
|
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
695  | 
apply (rule_tac t="Lam (p \<bullet> n) (p \<bullet> t')" and s="Lam (new ((p \<bullet> v, p \<bullet> s), Lam (p \<bullet> n) (p \<bullet> t'))) (((p \<bullet> n) \<leftrightarrow> new ((p \<bullet> v, p \<bullet> s), Lam (p \<bullet> n) (p \<bullet> t'))) \<bullet> t')" in subst)  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
696  | 
defer  | 
| 2165 | 697  | 
apply (subst match_Lam_simps)  | 
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
698  | 
defer  | 
| 2165 | 699  | 
apply (simp only: option.cases)  | 
700  | 
apply (simp only: prod.cases)  | 
|
701  | 
apply (simp only: lam.perm)  | 
|
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
702  | 
thm 1(1)  | 
| 2165 | 703  | 
sorry  | 
704  | 
qed  | 
|
705  | 
qed  | 
|
706  | 
||
| 
2172
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
707  | 
lemma subst_proper_eqs:  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
708  | 
"subst y s (Var x) = (if x = y then s else (Var x))"  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
709  | 
"subst y s (App l r) = App (subst y s l) (subst y s r)"  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
710  | 
"atom x \<sharp> (t, s) \<Longrightarrow> subst y s (Lam x t) = Lam x (subst y s t)"  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
711  | 
apply (subst subst.simps)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
712  | 
apply (simp only: match_Var_simps)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
713  | 
apply (simp only: option.simps)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
714  | 
apply (subst subst.simps)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
715  | 
apply (simp only: match_App_simps)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
716  | 
apply (simp only: option.simps)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
717  | 
apply (simp only: prod.simps)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
718  | 
apply (simp only: match_Var_simps)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
719  | 
apply (simp only: option.simps)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
720  | 
apply (subst subst.simps)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
721  | 
apply (simp only: match_Var_simps)  | 
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
722  | 
apply (simp only: option.simps)  | 
| 
2172
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
723  | 
apply (simp only: match_App_simps)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
724  | 
apply (simp only: option.simps)  | 
| 
2173
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
725  | 
apply (rule_tac t="Lam x t" and s="Lam (new ((y, s), Lam x t)) ((x \<leftrightarrow> new ((y, s), Lam x t)) \<bullet> t)" in subst)  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
726  | 
defer  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
727  | 
apply (subst match_Lam_simps)  | 
| 
 
477293d841e8
Match_Lam defined on Quotient Level.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2172 
diff
changeset
 | 
728  | 
defer  | 
| 
2172
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
729  | 
apply (simp only: option.simps)  | 
| 
 
fd5eec72c3f5
More on Function-defined subst.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2170 
diff
changeset
 | 
730  | 
apply (simp only: prod.simps)  | 
| 2157 | 731  | 
sorry  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
732  | 
*)  | 
| 1594 | 733  | 
end  | 
734  | 
||
735  | 
||
736  |