Nominal/Ex/Foo2.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 16 Dec 2010 08:42:48 +0000
changeset 2611 3d101f2f817c
parent 2609 666ffc8a92a9
child 2616 dd7490fdd998
permissions -rw-r--r--
simple cases for strong inducts done; infrastructure for the difficult ones is there
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
2603
90779aefbf1a first tests about exhaust
Christian Urban <urbanc@in.tum.de>
parents: 2602
diff changeset
    27
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
    28
3d101f2f817c simple cases for strong inducts done; infrastructure for the difficult ones is there
Christian Urban <urbanc@in.tum.de>
parents: 2609
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
2600
ca6b4bc7a871 kept the nested structure of constructors (belonging to one datatype)
Christian Urban <urbanc@in.tum.de>
parents: 2599
diff changeset
    36
2573
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
thm foo.distinct
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
thm foo.induct
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
thm foo.inducts
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
thm foo.exhaust
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
2609
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    52
(*
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    53
lemma ex_prop:
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    54
  shows "\<exists>n::nat. Suc n = n"
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    55
sorry
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    56
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    57
lemma test:
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    58
  shows "True"
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    59
apply -
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    60
ML_prf {*
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    61
  val (x, ctxt') = Obtain.result (K (
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    62
  EVERY [print_tac "test", 
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    63
         etac exE 1,
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    64
         print_tac "end"])) 
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    65
  @{thms ex_prop} @{context};
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    66
  ProofContext.verbose := true;
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    67
  ProofContext.pretty_context ctxt'
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    68
  |> Pretty.block
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    69
  |> Pretty.writeln
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    70
*}
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    71
*)
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
    72
2586
3ebc7ecfb0dd disabled the Foo examples, because of heavy work
Christian Urban <urbanc@in.tum.de>
parents: 2585
diff changeset
    73
2600
ca6b4bc7a871 kept the nested structure of constructors (belonging to one datatype)
Christian Urban <urbanc@in.tum.de>
parents: 2599
diff changeset
    74
2590
98dc38c250bb added abs_rename_res lemma
Christian Urban <urbanc@in.tum.de>
parents: 2589
diff changeset
    75
2586
3ebc7ecfb0dd disabled the Foo examples, because of heavy work
Christian Urban <urbanc@in.tum.de>
parents: 2585
diff changeset
    76
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    77
lemma test6:
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
    78
  fixes c::"'a::fs"
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
    79
  assumes "\<And>name. y = Var name \<Longrightarrow> P" 
3d101f2f817c simple cases for strong inducts done; infrastructure for the difficult ones is there
Christian Urban <urbanc@in.tum.de>
parents: 2609
diff changeset
    80
  and     "\<And>trm1 trm2. y = App trm1 trm2 \<Longrightarrow> P"
2592
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
    81
  and     "\<And>name trm. \<lbrakk>{atom name} \<sharp>* c; y = Lam name trm\<rbrakk> \<Longrightarrow> P" 
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
    82
  and     "\<And>a1 trm1 a2 trm2.  \<lbrakk>((set (bn a1)) \<union> (set (bn a2))) \<sharp>* c; y = Let1 a1 trm1 a2 trm2\<rbrakk> \<Longrightarrow> P"
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    83
  and     "\<exists>x1 x2 trm1 trm2. {atom x1, atom x2} \<sharp>* c \<and> y = Let2 x1 x2 trm1 trm2 \<Longrightarrow> P"
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
    84
  shows "P"
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
    85
apply(rule_tac y="y" in foo.exhaust(1))
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    86
apply(rule assms(1))
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    87
apply(assumption)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    88
apply(rule assms(2))
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    89
apply(assumption)
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    90
(* lam case *)
2609
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    91
apply(subgoal_tac "\<exists>p. (p \<bullet> {atom name}) \<sharp>* (c, name, trm)")
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
    92
apply(erule exE)
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
    93
apply(insert Abs_rename_lst)[1]
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    94
apply(drule_tac x="p" in meta_spec)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    95
apply(drule_tac x="[atom name]" in meta_spec)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    96
apply(drule_tac x="trm" in meta_spec)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
    97
apply(simp only: fresh_star_Pair set.simps)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    98
apply(drule meta_mp)
2609
666ffc8a92a9 freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
Christian Urban <urbanc@in.tum.de>
parents: 2608
diff changeset
    99
apply(simp add: fresh_star_def fresh_Nil fresh_Cons fresh_atom_at_base)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   100
apply(erule exE)
2592
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   101
apply(rule assms(3))
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
   102
apply(perm_simp)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   103
apply(erule conjE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   104
apply(assumption)
2592
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   105
apply(clarify)
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   106
apply(simp (no_asm) add: foo.eq_iff)
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   107
apply(perm_simp)
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   108
apply(assumption)
2589
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   109
apply(rule at_set_avoiding1)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   110
apply(simp)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   111
apply(simp add: finite_supp)
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   112
(* let1 case *)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   113
apply(subgoal_tac "\<exists>p. (p \<bullet> ((set (bn assg1)) \<union> (set (bn assg2)))) \<sharp>* (c, bn assg1, bn assg2, trm1, trm2)")
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   114
apply(erule exE)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   115
apply(rule assms(4))
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
   116
apply(perm_simp add: foo.permute_bn)
3d101f2f817c simple cases for strong inducts done; infrastructure for the difficult ones is there
Christian Urban <urbanc@in.tum.de>
parents: 2609
diff changeset
   117
apply(simp add: fresh_star_Pair)
3d101f2f817c simple cases for strong inducts done; infrastructure for the difficult ones is there
Christian Urban <urbanc@in.tum.de>
parents: 2609
diff changeset
   118
apply(erule conjE)+
3d101f2f817c simple cases for strong inducts done; infrastructure for the difficult ones is there
Christian Urban <urbanc@in.tum.de>
parents: 2609
diff changeset
   119
apply(assumption)
3d101f2f817c simple cases for strong inducts done; infrastructure for the difficult ones is there
Christian Urban <urbanc@in.tum.de>
parents: 2609
diff changeset
   120
apply(simp only:)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   121
apply(simp only: foo.eq_iff)
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
   122
(* *)
3d101f2f817c simple cases for strong inducts done; infrastructure for the difficult ones is there
Christian Urban <urbanc@in.tum.de>
parents: 2609
diff changeset
   123
apply(insert Abs_rename_lst)[1]
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   124
apply(drule_tac x="p" in meta_spec)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   125
apply(drule_tac x="bn assg1" in meta_spec)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   126
apply(drule_tac x="trm1" in meta_spec)
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
   127
apply(insert Abs_rename_lst)[1]
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   128
apply(drule_tac x="p" in meta_spec)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   129
apply(drule_tac x="bn assg2" in meta_spec)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   130
apply(drule_tac x="trm2" in meta_spec)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   131
apply(drule meta_mp)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   132
apply(simp only: union_eqvt fresh_star_Pair set.simps fresh_star_Un, simp)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   133
apply(drule meta_mp)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   134
apply(simp only: union_eqvt fresh_star_Pair set.simps fresh_star_Un, simp)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   135
apply(erule exE)+
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   136
apply(perm_simp add: foo.permute_bn)
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
   137
apply(simp add: fresh_star_Pair)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   138
apply(erule conjE)+
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
   139
apply(rule assms(4))
3d101f2f817c simple cases for strong inducts done; infrastructure for the difficult ones is there
Christian Urban <urbanc@in.tum.de>
parents: 2609
diff changeset
   140
apply(assumption)
3d101f2f817c simple cases for strong inducts done; infrastructure for the difficult ones is there
Christian Urban <urbanc@in.tum.de>
parents: 2609
diff changeset
   141
apply(simp add: foo.eq_iff refl)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   142
apply(rule conjI)
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
   143
apply(rule refl)
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
   144
apply(rule conjI)
2594
515e5496171c automated alpha_perm_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2593
diff changeset
   145
apply(rule foo.perm_bn_alpha)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   146
apply(rule conjI)
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
   147
apply(rule refl)
2594
515e5496171c automated alpha_perm_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2593
diff changeset
   148
apply(rule foo.perm_bn_alpha)
2589
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   149
apply(rule at_set_avoiding1)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   150
apply(simp)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   151
apply(simp add: finite_supp)
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   152
(* let2 case *)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   153
apply(subgoal_tac "\<exists>p. (p \<bullet> ({atom name1} \<union> {atom name2})) \<sharp>* (c, atom name1, atom name2, trm1, trm2)")
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   154
apply(erule exE)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   155
apply(rule assms(5))
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   156
apply(simp only:)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   157
apply(simp only: foo.eq_iff)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   158
apply(insert Abs_rename_list)[1]
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   159
apply(drule_tac x="p" in meta_spec)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   160
apply(drule_tac x="[atom name1] @ [atom name2]" in meta_spec)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   161
apply(drule_tac x="trm1" in meta_spec)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   162
apply(insert Abs_rename_list)[1]
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   163
apply(drule_tac x="p" in meta_spec)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   164
apply(drule_tac x="[atom name2]" in meta_spec)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   165
apply(drule_tac x="trm2" in meta_spec)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   166
apply(drule meta_mp)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   167
apply(simp only: union_eqvt fresh_star_Pair fresh_star_list fresh_star_Un, simp)
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   168
apply(auto)[1]
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   169
apply(perm_simp)
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   170
apply(auto simp add: fresh_star_def)[1]
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   171
apply(perm_simp)
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   172
apply(auto simp add: fresh_star_def)[1]
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   173
apply(perm_simp)
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   174
apply(auto simp add: fresh_star_def)[1]
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   175
apply(drule meta_mp)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   176
apply(perm_simp)
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   177
apply(auto simp add: fresh_star_def fresh_Pair fresh_Nil fresh_Cons)[1]
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   178
apply(erule exE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   179
apply(rule exI)+
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   180
apply(perm_simp add: foo.permute_bn)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   181
apply(rule conjI)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   182
prefer 2
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   183
apply(rule conjI)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   184
apply(assumption)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   185
apply(assumption)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   186
apply(simp add: fresh_star_Pair)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   187
apply(simp add: fresh_star_def)
2589
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   188
apply(rule at_set_avoiding1)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   189
apply(simp)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   190
apply(simp add: finite_supp)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   191
done
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
   192
2585
385add25dedf slightly simplified the Foo2 tests and hint at a general lemma
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   193
lemma test5:
385add25dedf slightly simplified the Foo2 tests and hint at a general lemma
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   194
  fixes c::"'a::fs"
385add25dedf slightly simplified the Foo2 tests and hint at a general lemma
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   195
  assumes "\<And>name. y = Var name \<Longrightarrow> P" 
385add25dedf slightly simplified the Foo2 tests and hint at a general lemma
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   196
  and     "\<And>trm1 trm2. y = App trm1 trm2 \<Longrightarrow> P"
385add25dedf slightly simplified the Foo2 tests and hint at a general lemma
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   197
  and     "\<And>name trm. \<lbrakk>{atom name} \<sharp>* c; y = Lam name trm\<rbrakk> \<Longrightarrow> P" 
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   198
  and     "\<And>a1 trm1 a2 trm2. \<lbrakk>((set (bn a1)) \<union> (set (bn a2))) \<sharp>* c; y = Let1 a1 trm1 a2 trm2\<rbrakk> \<Longrightarrow> P"
2585
385add25dedf slightly simplified the Foo2 tests and hint at a general lemma
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   199
  and     "\<And>x1 x2 trm1 trm2. \<lbrakk>{atom x1, atom x2} \<sharp>* c; y = Let2 x1 x2 trm1 trm2\<rbrakk> \<Longrightarrow> P"
385add25dedf slightly simplified the Foo2 tests and hint at a general lemma
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   200
  shows "P"
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   201
apply(rule_tac y="y" in test6)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   202
apply(erule exE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   203
apply(rule assms(1))
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   204
apply(assumption)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   205
apply(erule exE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   206
apply(rule assms(2))
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   207
apply(assumption)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   208
apply(rule assms(3))
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   209
apply(auto)[2]
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   210
apply(erule exE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   211
apply(rule assms(4))
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   212
apply(auto)[2]
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   213
apply(erule exE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   214
apply(rule assms(5))
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   215
apply(auto)[2]
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   216
done
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   217
2585
385add25dedf slightly simplified the Foo2 tests and hint at a general lemma
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   218
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   219
lemma strong_induct:
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   220
  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
   221
  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
   222
  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
   223
  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
   224
  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
   225
  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
   226
    \<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
   227
    \<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
   228
  and a4: "\<And>n1 n2 t1 t2 c. 
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   229
    \<lbrakk>{atom n1, atom n2} \<sharp>* c; \<And>d. P1 d t1; \<And>d. P1 d t2\<rbrakk> \<Longrightarrow> P1 c (Let2 n1 n2 t1 t2)"
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   230
  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
   231
  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
   232
  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
   233
  using assms
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   234
  apply(induction_schema)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   235
  apply(rule_tac y="trm" and c="c" in test5)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   236
  apply(simp_all)[5]
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   237
  apply(rule_tac y="assg" in foo.exhaust(2))
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   238
  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
   239
  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
   240
  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
   241
  done
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
   242
2573
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
end
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246