author | Christian Urban <urbanc@in.tum.de> |
Wed, 05 Jan 2011 17:33:43 +0000 | |
changeset 2639 | a8fc346deda3 |
parent 2638 | e1e2ca92760b |
permissions | -rw-r--r-- |
2636
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
theory Lambda |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
imports "../Nominal2" |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
begin |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
atom_decl name |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
nominal_datatype lam = |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
Var "name" |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
| App "lam" "lam" |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
| Lam x::"name" l::"lam" bind x in l |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
thm lam.distinct |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
thm lam.induct |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
thm lam.exhaust lam.strong_exhaust |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
thm lam.fv_defs |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
thm lam.bn_defs |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
thm lam.perm_simps |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
thm lam.eq_iff |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
thm lam.fv_bn_eqvt |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
thm lam.size_eqvt |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
22 |
|
2638
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
23 |
|
2636
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
24 |
section {* Typing *} |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
25 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
26 |
nominal_datatype ty = |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
27 |
TVar string |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
| TFun ty ty ("_ \<rightarrow> _") |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
30 |
lemma ty_fresh: |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
31 |
fixes x::"name" |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
32 |
and T::"ty" |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
33 |
shows "atom x \<sharp> T" |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
apply (nominal_induct T rule: ty.strong_induct) |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
35 |
apply (simp_all add: ty.fresh pure_fresh) |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
36 |
done |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
37 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
38 |
inductive |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
39 |
valid :: "(name \<times> ty) list \<Rightarrow> bool" |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
40 |
where |
2638
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
41 |
v_Nil[intro]: "valid []" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
42 |
| v_Cons[intro]: "\<lbrakk>atom x \<sharp> Gamma; valid Gamma\<rbrakk> \<Longrightarrow> valid ((x, T)#Gamma)" |
2636
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
43 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
44 |
inductive |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
45 |
typing :: "(name\<times>ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60) |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
46 |
where |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
47 |
t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T" |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
48 |
| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2" |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
49 |
| t_Lam[intro]: "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam x t : T1 \<rightarrow> T2" |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
50 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
51 |
thm typing.intros |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
52 |
thm typing.induct |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
53 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
54 |
equivariance valid |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
55 |
equivariance typing |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
57 |
nominal_inductive typing |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
58 |
avoids t_Lam: "x" |
2639
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
59 |
by (simp_all add: fresh_star_def ty_fresh lam.fresh) |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
60 |
|
2636
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
61 |
|
2638
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
62 |
thm typing.strong_induct |
2636
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
63 |
|
2638
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
64 |
abbreviation |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
65 |
"sub_context" :: "(name \<times> ty) list \<Rightarrow> (name \<times> ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [60,60] 60) |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
66 |
where |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
67 |
"\<Gamma>1 \<subseteq> \<Gamma>2 \<equiv> \<forall>x T. (x, T) \<in> set \<Gamma>1 \<longrightarrow> (x, T) \<in> set \<Gamma>2" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
68 |
|
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
69 |
text {* Now it comes: The Weakening Lemma *} |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
70 |
|
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
71 |
text {* |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
72 |
The first version is, after setting up the induction, |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
73 |
completely automatic except for use of atomize. *} |
2636
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
74 |
|
2638
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
75 |
lemma weakening_version2: |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
76 |
fixes \<Gamma>1 \<Gamma>2::"(name \<times> ty) list" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
77 |
and t ::"lam" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
78 |
and \<tau> ::"ty" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
79 |
assumes a: "\<Gamma>1 \<turnstile> t : T" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
80 |
and b: "valid \<Gamma>2" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
81 |
and c: "\<Gamma>1 \<subseteq> \<Gamma>2" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
82 |
shows "\<Gamma>2 \<turnstile> t : T" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
83 |
using a b c |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
84 |
proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct) |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
85 |
case (t_Var \<Gamma>1 x T) (* variable case *) |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
86 |
have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
87 |
moreover |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
88 |
have "valid \<Gamma>2" by fact |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
89 |
moreover |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
90 |
have "(x,T)\<in> set \<Gamma>1" by fact |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
91 |
ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
92 |
next |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
93 |
case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *) |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
94 |
have vc: "atom x \<sharp> \<Gamma>2" by fact (* variable convention *) |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
95 |
have ih: "\<lbrakk>valid ((x, T1) # \<Gamma>2); (x, T1) # \<Gamma>1 \<subseteq> (x, T1) # \<Gamma>2\<rbrakk> \<Longrightarrow> (x, T1) # \<Gamma>2 \<turnstile> t : T2" by fact |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
96 |
have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
97 |
then have "(x, T1) # \<Gamma>1 \<subseteq> (x, T1) # \<Gamma>2" by simp |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
98 |
moreover |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
99 |
have "valid \<Gamma>2" by fact |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
100 |
then have "valid ((x, T1) # \<Gamma>2)" using vc by (simp add: v_Cons) |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
101 |
ultimately have "(x, T1) # \<Gamma>2 \<turnstile> t : T2" using ih by simp |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
102 |
with vc show "\<Gamma>2 \<turnstile> Lam x t : T1 \<rightarrow> T2" by auto |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
103 |
qed (auto) (* app case *) |
2636
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
104 |
|
2638
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
105 |
lemma weakening_version1: |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
106 |
fixes \<Gamma>1 \<Gamma>2::"(name \<times> ty) list" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
107 |
assumes a: "\<Gamma>1 \<turnstile> t : T" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
108 |
and b: "valid \<Gamma>2" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
109 |
and c: "\<Gamma>1 \<subseteq> \<Gamma>2" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
110 |
shows "\<Gamma>2 \<turnstile> t : T" |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
111 |
using a b c |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
112 |
apply (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct) |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
113 |
apply (auto | atomize)+ |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
114 |
done |
e1e2ca92760b
strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents:
2637
diff
changeset
|
115 |
|
2636
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
116 |
|
2639
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
117 |
inductive |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
118 |
Acc :: "('a::pt \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
119 |
where |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
120 |
AccI: "(\<And>y. R y x \<Longrightarrow> Acc R y) \<Longrightarrow> Acc R x" |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
121 |
|
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
122 |
|
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
123 |
lemma Acc_eqvt [eqvt]: |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
124 |
fixes p::"perm" |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
125 |
assumes a: "Acc R x" |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
126 |
shows "Acc (p \<bullet> R) (p \<bullet> x)" |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
127 |
using a |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
128 |
apply(induct) |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
129 |
apply(rule AccI) |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
130 |
apply(rotate_tac 1) |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
131 |
apply(drule_tac x="-p \<bullet> y" in meta_spec) |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
132 |
apply(simp) |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
133 |
apply(drule meta_mp) |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
134 |
apply(rule_tac p="p" in permute_boolE) |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
135 |
apply(perm_simp add: permute_minus_cancel) |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
136 |
apply(assumption) |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
137 |
apply(assumption) |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
138 |
done |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
139 |
|
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
140 |
|
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
141 |
nominal_inductive Acc . |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
142 |
|
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
143 |
thm Acc.strong_induct |
a8fc346deda3
exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
2638
diff
changeset
|
144 |
|
2636
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
145 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
146 |
end |
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
147 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
148 |
|
0865caafbfe6
file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
149 |