Nominal/Ex/Foo2.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 07 Dec 2010 14:27:39 +0000
changeset 2598 b136721eedb2
parent 2594 515e5496171c
child 2599 d6fe94028a5d
permissions -rw-r--r--
automated permute_bn 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
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
 "bn (As x y t a) = [atom x] @ bn a"
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
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
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    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
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
thm foo.distinct
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
thm foo.induct
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
thm foo.inducts
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
thm foo.exhaust
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
thm foo.fv_defs
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
thm foo.bn_defs
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
thm foo.perm_simps
2575
b1d38940040a single rename in let2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2573
diff changeset
    40
thm foo.eq_iff(5)
2573
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
thm foo.fv_bn_eqvt
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
thm foo.size_eqvt
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
thm foo.supports
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
thm foo.fsupp
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
thm foo.supp
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
thm foo.fresh
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
    48
text {*
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
    49
  tests by cu
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
    50
*}
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
    51
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    52
lemma set_renaming_perm:
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    53
  assumes a: "p \<bullet> bs \<inter> bs = {}" 
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    54
  and     b: "finite bs"
2586
3ebc7ecfb0dd disabled the Foo examples, because of heavy work
Christian Urban <urbanc@in.tum.de>
parents: 2585
diff changeset
    55
  shows "\<exists>q. q \<bullet> bs = p \<bullet> bs \<and> supp q \<subseteq> bs \<union> (p \<bullet> bs)"
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
    56
using b a
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    57
proof (induct)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    58
  case empty
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    59
  have "0 \<bullet> {} = p \<bullet> {} \<and> supp (0::perm) \<subseteq> {} \<union> p \<bullet> {}"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    60
    by (simp add: permute_set_eq supp_perm)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    61
  then show "\<exists>q. q \<bullet> {} = p \<bullet> {} \<and> supp q \<subseteq> {} \<union> p \<bullet> {}" by blast
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    62
next
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    63
  case (insert a bs)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    64
  then have " \<exists>q. q \<bullet> bs = p \<bullet> bs \<and> supp q \<subseteq> bs \<union> p \<bullet> bs" 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    65
    by (perm_simp) (auto)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    66
  then obtain q where *: "q \<bullet> bs = p \<bullet> bs" and **: "supp q \<subseteq> bs \<union> p \<bullet> bs" by blast 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    67
  { assume 1: "q \<bullet> a = p \<bullet> a"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    68
    have "q \<bullet> insert a bs = p \<bullet> insert a bs" using 1 * by (simp add: insert_eqvt)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    69
    moreover 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    70
    have "supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    71
      using ** by (auto simp add: insert_eqvt)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    72
    ultimately 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    73
    have "\<exists>q. q \<bullet> insert a bs = p \<bullet> insert a bs \<and> supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" by blast
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    74
  }
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    75
  moreover
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    76
  { assume 2: "q \<bullet> a \<noteq> p \<bullet> a"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    77
    def q' \<equiv> "((q \<bullet> a) \<rightleftharpoons> (p \<bullet> a)) + q"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    78
    { have "(q \<bullet> a) \<notin> (p \<bullet> bs)" using `a \<notin> bs` *[symmetric] by (simp add: mem_permute_iff)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    79
      moreover 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    80
      have "(p \<bullet> a) \<notin> (p \<bullet> bs)" using `a \<notin> bs` by (simp add: mem_permute_iff)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    81
      ultimately 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    82
      have "q' \<bullet> insert a bs = p \<bullet> insert a bs" using 2 * unfolding q'_def 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    83
        by (simp add: insert_eqvt  swap_set_not_in) 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    84
    }
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    85
    moreover 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    86
    { have "{q \<bullet> a, p \<bullet> a} \<subseteq> insert a bs \<union> p \<bullet> insert a bs"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    87
	using ** `a \<notin> bs` `p \<bullet> insert a bs \<inter> insert a bs = {}`
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    88
	by (auto simp add: supp_perm insert_eqvt)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    89
      then have "supp (q \<bullet> a \<rightleftharpoons> p \<bullet> a) \<subseteq> insert a bs \<union> p \<bullet> insert a bs" by (simp add: supp_swap)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    90
      moreover
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    91
      have "supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    92
	using ** by (auto simp add: insert_eqvt)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    93
      ultimately 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    94
      have "supp q' \<subseteq> insert a bs \<union> p \<bullet> insert a bs" 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    95
        unfolding q'_def using supp_plus_perm by blast
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    96
    }
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    97
    ultimately 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    98
    have "\<exists>q. q \<bullet> insert a bs = p \<bullet> insert a bs \<and> supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" by blast
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    99
  }
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   100
  ultimately show "\<exists>q. q \<bullet> insert a bs = p \<bullet> insert a bs \<and> supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   101
    by blast
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   102
qed
2592
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   103
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   104
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   105
lemma Abs_rename_set:
2586
3ebc7ecfb0dd disabled the Foo examples, because of heavy work
Christian Urban <urbanc@in.tum.de>
parents: 2585
diff changeset
   106
  fixes x::"'a::fs"
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   107
  assumes a: "(p \<bullet> bs) \<sharp>* (bs, x)" 
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   108
  and     b: "finite bs"
2586
3ebc7ecfb0dd disabled the Foo examples, because of heavy work
Christian Urban <urbanc@in.tum.de>
parents: 2585
diff changeset
   109
  shows "\<exists>y. [bs]set. x = [p \<bullet> bs]set. y"
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   110
proof -
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   111
  from a b have "p \<bullet> bs \<inter> bs = {}" using at_fresh_star_inter by (auto simp add: fresh_star_Pair)   
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   112
  with set_renaming_perm 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   113
  obtain q where *: "q \<bullet> bs = p \<bullet> bs" and **: "supp q \<subseteq> bs \<union> (p \<bullet> bs)" using b by metis
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   114
  have "[bs]set. x =  q \<bullet> ([bs]set. x)"
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   115
    apply(rule perm_supp_eq[symmetric])
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   116
    using a **
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   117
    unfolding fresh_star_Pair
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   118
    unfolding Abs_fresh_star_iff
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   119
    unfolding fresh_star_def
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   120
    by auto
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   121
  also have "\<dots> = [q \<bullet> bs]set. (q \<bullet> x)" by simp
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   122
  also have "\<dots> = [p \<bullet> bs]set. (q \<bullet> x)" using * by simp
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   123
  finally have "[bs]set. x = [p \<bullet> bs]set. (q \<bullet> x)" .
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   124
  then show "\<exists>y. [bs]set. x = [p \<bullet> bs]set. y" by blast
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   125
qed
2586
3ebc7ecfb0dd disabled the Foo examples, because of heavy work
Christian Urban <urbanc@in.tum.de>
parents: 2585
diff changeset
   126
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   127
lemma Abs_rename_res:
2590
98dc38c250bb added abs_rename_res lemma
Christian Urban <urbanc@in.tum.de>
parents: 2589
diff changeset
   128
  fixes x::"'a::fs"
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   129
  assumes a: "(p \<bullet> bs) \<sharp>* (bs, x)" 
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   130
  and     b: "finite bs"
2590
98dc38c250bb added abs_rename_res lemma
Christian Urban <urbanc@in.tum.de>
parents: 2589
diff changeset
   131
  shows "\<exists>y. [bs]res. x = [p \<bullet> bs]res. y"
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   132
proof -
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   133
  from a b have "p \<bullet> bs \<inter> bs = {}" using at_fresh_star_inter by (simp add: fresh_star_Pair) 
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   134
  with set_renaming_perm 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   135
  obtain q where *: "q \<bullet> bs = p \<bullet> bs" and **: "supp q \<subseteq> bs \<union> (p \<bullet> bs)" using b by metis
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   136
  have "[bs]res. x =  q \<bullet> ([bs]res. x)"
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   137
    apply(rule perm_supp_eq[symmetric])
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   138
    using a **
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   139
    unfolding fresh_star_Pair
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   140
    unfolding Abs_fresh_star_iff
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   141
    unfolding fresh_star_def
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   142
    by auto
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   143
  also have "\<dots> = [q \<bullet> bs]res. (q \<bullet> x)" by simp
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   144
  also have "\<dots> = [p \<bullet> bs]res. (q \<bullet> x)" using * by simp
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   145
  finally have "[bs]res. x = [p \<bullet> bs]res. (q \<bullet> x)" .
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   146
  then show "\<exists>y. [bs]res. x = [p \<bullet> bs]res. y" by blast
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   147
qed
2590
98dc38c250bb added abs_rename_res lemma
Christian Urban <urbanc@in.tum.de>
parents: 2589
diff changeset
   148
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   149
lemma list_renaming_perm:
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   150
  fixes bs::"atom list"
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   151
  assumes a: "(p \<bullet> (set bs)) \<inter> (set bs) = {}"
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   152
  shows "\<exists>q. q \<bullet> bs = p \<bullet> bs \<and> supp q \<subseteq> (set bs) \<union> (p \<bullet> (set bs))"
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   153
using a
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   154
proof (induct bs)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   155
  case Nil
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   156
  have "0 \<bullet> [] = p \<bullet> [] \<and> supp (0::perm) \<subseteq> set [] \<union> p \<bullet> set []"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   157
    by (simp add: permute_set_eq supp_perm)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   158
  then show "\<exists>q. q \<bullet> [] = p \<bullet> [] \<and> supp q \<subseteq> set [] \<union> p \<bullet> (set [])" by blast
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   159
next
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   160
  case (Cons a bs)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   161
  then have " \<exists>q. q \<bullet> bs = p \<bullet> bs \<and> supp q \<subseteq> set bs \<union> p \<bullet> (set bs)" 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   162
    by (perm_simp) (auto)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   163
  then obtain q where *: "q \<bullet> bs = p \<bullet> bs" and **: "supp q \<subseteq> set bs \<union> p \<bullet> (set bs)" by blast 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   164
  { assume 1: "a \<in> set bs"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   165
    have "q \<bullet> a = p \<bullet> a" using * 1 by (induct bs) (auto)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   166
    then have "q \<bullet> (a # bs) = p \<bullet> (a # bs)" using * by simp 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   167
    moreover 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   168
    have "supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" using ** by (auto simp add: insert_eqvt)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   169
    ultimately 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   170
    have "\<exists>q. q \<bullet> (a # bs) = p \<bullet> (a # bs) \<and> supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" by blast
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   171
  }
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   172
  moreover
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   173
  { assume 2: "a \<notin> set bs"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   174
    def q' \<equiv> "((q \<bullet> a) \<rightleftharpoons> (p \<bullet> a)) + q"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   175
    { have "(q \<bullet> a) \<sharp> (p \<bullet> bs)" using `a \<notin> set bs` *[symmetric] 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   176
	by (simp add: fresh_permute_iff) (simp add: fresh_def supp_of_atom_list)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   177
      moreover 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   178
      have "(p \<bullet> a) \<sharp> (p \<bullet> bs)" using `a \<notin> set bs` 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   179
	by (simp add: fresh_permute_iff) (simp add: fresh_def supp_of_atom_list)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   180
      ultimately 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   181
      have "q' \<bullet> (a # bs) = p \<bullet> (a # bs)" using 2 * unfolding q'_def 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   182
        by (simp add: swap_fresh_fresh) 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   183
    }
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   184
    moreover 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   185
    { have "{q \<bullet> a, p \<bullet> a} \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   186
	using ** `a \<notin> set bs` `p \<bullet> (set (a # bs)) \<inter> set (a # bs) = {}`
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   187
	by (auto simp add: supp_perm insert_eqvt)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   188
      then have "supp (q \<bullet> a \<rightleftharpoons> p \<bullet> a) \<subseteq> set (a # bs) \<union> p \<bullet> set (a # bs)" by (simp add: supp_swap)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   189
      moreover
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   190
      have "supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   191
	using ** by (auto simp add: insert_eqvt)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   192
      ultimately 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   193
      have "supp q' \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   194
        unfolding q'_def using supp_plus_perm by blast
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   195
    }
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   196
    ultimately 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   197
    have "\<exists>q. q \<bullet> (a # bs) = p \<bullet> (a # bs) \<and> supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" by blast
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   198
  }
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   199
  ultimately show "\<exists>q. q \<bullet> (a # bs) = p \<bullet> (a # bs) \<and> supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   200
    by blast
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   201
qed
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   202
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   203
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   204
lemma Abs_rename_list:
2585
385add25dedf slightly simplified the Foo2 tests and hint at a general lemma
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   205
  fixes x::"'a::fs"
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   206
  assumes a: "(p \<bullet> (set bs)) \<sharp>* (bs, x)" 
2586
3ebc7ecfb0dd disabled the Foo examples, because of heavy work
Christian Urban <urbanc@in.tum.de>
parents: 2585
diff changeset
   207
  shows "\<exists>y. [bs]lst. x = [p \<bullet> bs]lst. y"
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   208
proof -
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   209
  from a have "p \<bullet> (set bs) \<inter> (set bs) = {}" using at_fresh_star_inter 
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   210
    by (simp add: fresh_star_Pair fresh_star_set)
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   211
  with list_renaming_perm 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   212
  obtain q where *: "q \<bullet> bs = p \<bullet> bs" and **: "supp q \<subseteq> set bs \<union> (p \<bullet> set bs)" by metis 
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   213
  have "[bs]lst. x =  q \<bullet> ([bs]lst. x)"
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   214
    apply(rule perm_supp_eq[symmetric])
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   215
    using a **
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   216
    unfolding fresh_star_Pair
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   217
    unfolding Abs_fresh_star_iff
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   218
    unfolding fresh_star_def
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   219
    by auto
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   220
  also have "\<dots> = [q \<bullet> bs]lst. (q \<bullet> x)" by simp
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   221
  also have "\<dots> = [p \<bullet> bs]lst. (q \<bullet> x)" using * by simp
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   222
  finally have "[bs]lst. x = [p \<bullet> bs]lst. (q \<bullet> x)" .
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   223
  then show "\<exists>y. [bs]lst. x = [p \<bullet> bs]lst. y" by blast
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   224
qed
2586
3ebc7ecfb0dd disabled the Foo examples, because of heavy work
Christian Urban <urbanc@in.tum.de>
parents: 2585
diff changeset
   225
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   226
(*
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   227
thm at_set_avoiding1[THEN exE]
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   228
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   229
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   230
lemma Let1_rename:
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   231
  fixes c::"'a::fs"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   232
  shows "\<exists>name' trm'. {atom name'} \<sharp>* c \<and>  Lam name trm = Lam name' trm'"
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   233
apply(simp only: foo.eq_iff)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   234
apply(rule at_set_avoiding1[where c="(c, atom name, trm)" and xs="set [atom name]", THEN exE])
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   235
apply(simp)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   236
apply(simp add: finite_supp)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   237
apply(perm_simp)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   238
apply(rule Abs_rename_list[THEN exE])
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   239
apply(simp only: set_eqvt)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   240
apply
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   241
sorry 
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   242
*)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   243
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   244
lemma test6:
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
   245
  fixes c::"'a::fs"
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   246
  assumes "\<exists>name. y = Var name \<Longrightarrow> P" 
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   247
  and     "\<exists>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
   248
  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
   249
  and     "\<exists>a1 trm1 a2 trm2.  ((set (bn a1)) \<union> (set (bn a2))) \<sharp>* c \<and> y = Let1 a1 trm1 a2 trm2 \<Longrightarrow> P"
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   250
  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
   251
  shows "P"
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   252
apply(rule_tac y="y" in foo.exhaust(1))
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   253
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
   254
apply(rule exI)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   255
apply(assumption)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   256
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
   257
apply(rule exI)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   258
apply(assumption)
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   259
(* lam case *)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   260
(*
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   261
apply(rule Let1_rename)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   262
apply(rule assms(3))
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   263
apply(assumption)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   264
apply(simp)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   265
*)
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   266
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   267
apply(subgoal_tac "\<exists>p. (p \<bullet> {atom name}) \<sharp>* (c, [atom name], trm)")
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
   268
apply(erule exE)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   269
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
   270
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
   271
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
   272
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
   273
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
   274
apply(drule meta_mp)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   275
apply(simp)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   276
apply(erule exE)
2592
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   277
apply(rule assms(3))
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
   278
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
   279
apply(erule conjE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   280
apply(assumption)
2592
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   281
apply(clarify)
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   282
apply(simp (no_asm) add: foo.eq_iff)
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   283
apply(perm_simp)
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   284
apply(assumption)
2589
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   285
apply(rule at_set_avoiding1)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   286
apply(simp)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   287
apply(simp add: finite_supp)
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   288
(* let1 case *)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   289
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
   290
apply(erule exE)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   291
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
   292
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
   293
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
   294
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
   295
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
   296
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
   297
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
   298
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
   299
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
   300
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
   301
apply(drule meta_mp)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   302
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
   303
apply(drule meta_mp)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   304
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
   305
apply(erule exE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   306
apply(rule exI)+
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   307
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
   308
apply(rule conjI)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   309
apply(simp add: fresh_star_Pair fresh_star_Un)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   310
apply(erule conjE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   311
apply(rule conjI)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   312
apply(assumption)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   313
apply(rotate_tac 4)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   314
apply(assumption)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   315
apply(rule conjI)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   316
apply(assumption)
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
   317
apply(rule conjI)
2594
515e5496171c automated alpha_perm_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2593
diff changeset
   318
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
   319
apply(rule conjI)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   320
apply(assumption)
2594
515e5496171c automated alpha_perm_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2593
diff changeset
   321
apply(rule foo.perm_bn_alpha)
2589
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   322
apply(rule at_set_avoiding1)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   323
apply(simp)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   324
apply(simp add: finite_supp)
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   325
(* let2 case *)
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   326
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
   327
apply(erule exE)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   328
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
   329
apply(simp only:)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   330
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
   331
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
   332
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
   333
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
   334
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
   335
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
   336
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
   337
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
   338
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
   339
apply(drule meta_mp)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   340
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
   341
apply(auto)[1]
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   342
apply(perm_simp)
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   343
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
   344
apply(perm_simp)
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   345
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
   346
apply(perm_simp)
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   347
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
   348
apply(drule meta_mp)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   349
apply(perm_simp)
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   350
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
   351
apply(erule exE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   352
apply(rule exI)+
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
   353
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
   354
apply(rule conjI)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   355
prefer 2
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   356
apply(rule conjI)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   357
apply(assumption)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   358
apply(assumption)
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2590
diff changeset
   359
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
   360
apply(simp add: fresh_star_def)
2589
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   361
apply(rule at_set_avoiding1)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   362
apply(simp)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   363
apply(simp add: finite_supp)
9781db0e2196 completed proofs in Foo2
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
   364
done
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
   365
2585
385add25dedf slightly simplified the Foo2 tests and hint at a general lemma
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   366
lemma test5:
385add25dedf slightly simplified the Foo2 tests and hint at a general lemma
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   367
  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
   368
  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
   369
  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
   370
  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
   371
  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
   372
  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
   373
  shows "P"
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   374
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
   375
apply(erule exE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   376
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
   377
apply(assumption)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   378
apply(erule exE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   379
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
   380
apply(assumption)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   381
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
   382
apply(auto)[2]
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   383
apply(erule exE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   384
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
   385
apply(auto)[2]
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   386
apply(erule exE)+
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   387
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
   388
apply(auto)[2]
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   389
done
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   390
2585
385add25dedf slightly simplified the Foo2 tests and hint at a general lemma
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   391
2588
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   392
lemma strong_induct:
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   393
  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
   394
  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
   395
  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
   396
  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
   397
  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
   398
  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
   399
    \<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
   400
    \<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
   401
  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
   402
    \<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
   403
  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
   404
  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
   405
  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
   406
  using assms
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   407
  apply(induction_schema)
8f5420681039 completed the strong exhausts rules for Foo2 using general lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2586
diff changeset
   408
  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
   409
  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
   410
  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
   411
  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
   412
  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
   413
  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
   414
  done
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2579
diff changeset
   415
2573
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
end
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
6c131c089ce2 added example Foo2.thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419