Nominal/Ex/Foo2.thy
author Christian Urban <urbanc@in.tum.de>
Sun, 26 Dec 2010 16:35:16 +0000
changeset 2628 16ffbc8442ca
parent 2626 d1bdc281be2b
child 2629 ffb5a181844b
permissions -rw-r--r--
generated goals for strong induction theorems.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2573
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Foo2
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports "../Nominal2" 
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
(* 
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
  Contrived example that has more than one
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
  binding clause
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
*)
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
atom_decl name
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
nominal_datatype foo: trm =
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
  Var "name"
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
| App "trm" "trm"
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
| Lam x::"name" t::"trm"  bind x in t
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
| Let1 a1::"assg" t1::"trm" a2::"assg" t2::"trm" bind "bn a1" in t1, bind "bn a2" in t2
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
| Let2 x::"name" y::"name" t1::"trm" t2::"trm" bind x y in t1, bind y in t2
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
and assg =
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  As_Nil
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
| As "name" x::"name" t::"trm" "assg" 
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
binder
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  bn::"assg \<Rightarrow> atom list"
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    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
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
| "bn (As_Nil) = []"
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
2628
16ffbc8442ca generated goals for strong induction theorems.
Christian Urban <urbanc@in.tum.de>
parents: 2626
diff changeset
    27
16ffbc8442ca generated goals for strong induction theorems.
Christian Urban <urbanc@in.tum.de>
parents: 2626
diff changeset
    28
16ffbc8442ca generated goals for strong induction theorems.
Christian Urban <urbanc@in.tum.de>
parents: 2626
diff changeset
    29
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    30
thm foo.bn_defs
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    31
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
    32
thm foo.perm_bn_alpha
2573
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
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
    34
thm foo.bn_finite
2573
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
thm foo.distinct
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
thm foo.induct
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
thm foo.inducts
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
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
    40
thm foo.strong_exhaust
2573
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
thm foo.fv_defs
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
thm foo.bn_defs
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
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
    44
thm foo.eq_iff
2573
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
thm foo.fv_bn_eqvt
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
thm foo.size_eqvt
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
thm foo.supports
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
thm foo.fsupp
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
thm foo.supp
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
thm foo.fresh
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    52
lemma strong_induct:
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    53
  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
    54
  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
    55
  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
    56
  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
    57
  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
    58
  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
    59
    \<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
    60
    \<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
    61
  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
    62
    \<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
    63
  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
    64
  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
    65
  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
    66
  using assms
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    67
  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
    68
  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
    69
  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
    70
  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
    71
  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
    72
  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
    73
  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
    74
  done
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
    75
2573
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
end
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79