author | Christian Urban <urbanc@in.tum.de> |
Thu, 23 Dec 2010 00:46:06 +0000 | |
changeset 2626 | d1bdc281be2b |
parent 2616 | dd7490fdd998 |
child 2628 | 16ffbc8442ca |
permissions | -rw-r--r-- |
2573 | 1 |
theory Foo2 |
2 |
imports "../Nominal2" |
|
3 |
begin |
|
4 |
||
5 |
(* |
|
6 |
Contrived example that has more than one |
|
7 |
binding clause |
|
8 |
*) |
|
9 |
||
10 |
atom_decl name |
|
11 |
||
12 |
nominal_datatype foo: trm = |
|
13 |
Var "name" |
|
14 |
| App "trm" "trm" |
|
15 |
| Lam x::"name" t::"trm" bind x in t |
|
16 |
| Let1 a1::"assg" t1::"trm" a2::"assg" t2::"trm" bind "bn a1" in t1, bind "bn a2" in t2 |
|
17 |
| Let2 x::"name" y::"name" t1::"trm" t2::"trm" bind x y in t1, bind y in t2 |
|
18 |
and assg = |
|
19 |
As_Nil |
|
20 |
| As "name" x::"name" t::"trm" "assg" |
|
21 |
binder |
|
22 |
bn::"assg \<Rightarrow> atom list" |
|
23 |
where |
|
2611
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
Christian Urban <urbanc@in.tum.de>
parents:
2609
diff
changeset
|
24 |
"bn (As x y t a) = [atom x] @ bn a" |
2573 | 25 |
| "bn (As_Nil) = []" |
26 |
||
2598
b136721eedb2
automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents:
2594
diff
changeset
|
27 |
thm foo.bn_defs |
b136721eedb2
automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents:
2594
diff
changeset
|
28 |
thm foo.permute_bn |
2593
25dcb2b1329e
ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
29 |
thm foo.perm_bn_alpha |
2573 | 30 |
thm foo.perm_bn_simps |
2593
25dcb2b1329e
ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
31 |
thm foo.bn_finite |
2573 | 32 |
|
2600
ca6b4bc7a871
kept the nested structure of constructors (belonging to one datatype)
Christian Urban <urbanc@in.tum.de>
parents:
2599
diff
changeset
|
33 |
|
2573 | 34 |
thm foo.distinct |
35 |
thm foo.induct |
|
36 |
thm foo.inducts |
|
37 |
thm foo.exhaust |
|
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
parents:
2616
diff
changeset
|
38 |
thm foo.strong_exhaust |
2573 | 39 |
thm foo.fv_defs |
40 |
thm foo.bn_defs |
|
41 |
thm foo.perm_simps |
|
2607
7430e07a5d61
moved setify and listify functions into the library; introduced versions that have a type argument
Christian Urban <urbanc@in.tum.de>
parents:
2603
diff
changeset
|
42 |
thm foo.eq_iff |
2573 | 43 |
thm foo.fv_bn_eqvt |
44 |
thm foo.size_eqvt |
|
45 |
thm foo.supports |
|
46 |
thm foo.fsupp |
|
47 |
thm foo.supp |
|
48 |
thm foo.fresh |
|
49 |
||
2588
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
50 |
lemma strong_induct: |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
51 |
fixes c :: "'a :: fs" |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
52 |
and assg :: assg and trm :: trm |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
53 |
assumes a0: "\<And>name c. P1 c (Var name)" |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
54 |
and a1: "\<And>trm1 trm2 c. \<lbrakk>\<And>d. P1 d trm1; \<And>d. P1 d trm2\<rbrakk> \<Longrightarrow> P1 c (App trm1 trm2)" |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
55 |
and a2: "\<And>name trm c. (\<And>d. P1 d trm) \<Longrightarrow> P1 c (Lam name trm)" |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
56 |
and a3: "\<And>a1 t1 a2 t2 c. |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
57 |
\<lbrakk>((set (bn a1)) \<union> (set (bn a2))) \<sharp>* c; \<And>d. P2 c a1; \<And>d. P1 d t1; \<And>d. P2 d a2; \<And>d. P1 d t2\<rbrakk> |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
58 |
\<Longrightarrow> P1 c (Let1 a1 t1 a2 t2)" |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
59 |
and a4: "\<And>n1 n2 t1 t2 c. |
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
parents:
2616
diff
changeset
|
60 |
\<lbrakk>({atom n1} \<union> {atom n2}) \<sharp>* c; \<And>d. P1 d t1; \<And>d. P1 d t2\<rbrakk> \<Longrightarrow> P1 c (Let2 n1 n2 t1 t2)" |
2588
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
61 |
and a5: "\<And>c. P2 c As_Nil" |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
62 |
and a6: "\<And>name1 name2 trm assg c. \<lbrakk>\<And>d. P1 d trm; \<And>d. P2 d assg\<rbrakk> \<Longrightarrow> P2 c (As name1 name2 trm assg)" |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
63 |
shows "P1 c trm" "P2 c assg" |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
64 |
using assms |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
65 |
apply(induction_schema) |
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
parents:
2616
diff
changeset
|
66 |
apply(rule_tac y="trm" and c="c" in foo.strong_exhaust(1)) |
2588
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
67 |
apply(simp_all)[5] |
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
parents:
2616
diff
changeset
|
68 |
apply(rule_tac ya="assg" in foo.strong_exhaust(2)) |
2588
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
69 |
apply(simp_all)[2] |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
70 |
apply(relation "measure (sum_case (size o snd) (size o snd))") |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
71 |
apply(simp_all add: foo.size) |
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2586
diff
changeset
|
72 |
done |
2584
1eac050a36f4
completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents:
2579
diff
changeset
|
73 |
|
2573 | 74 |
end |
75 |
||
76 |
||
77 |