Nominal/Ex/Foo1.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 13 Nov 2010 22:23:26 +0000
changeset 2562 e8ec504dddf2
parent 2561 7926f1cb45eb
child 2564 5be8e34c2c0e
permissions -rw-r--r--
lifted permute_bn constants
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Foo1
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports "../Nominal2" 
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
(* 
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
  Contrived example that has more than one
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
  binding function for a datatype
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
*)
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
atom_decl name
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
nominal_datatype foo: trm =
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
  Var "name"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
| App "trm" "trm"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
| Lam x::"name" t::"trm"  bind x in t
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
| Let1 a::"assg" t::"trm"  bind "bn1 a" in t
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
| Let2 a::"assg" t::"trm"  bind "bn2 a" in t
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
| Let3 a::"assg" t::"trm"  bind "bn3 a" in t
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
and assg =
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  As "name" "name" "trm"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
binder
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  bn1::"assg \<Rightarrow> atom list" and
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  bn2::"assg \<Rightarrow> atom list" and
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  bn3::"assg \<Rightarrow> atom list"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
where
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
  "bn1 (As x y t) = [atom x]"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
| "bn2 (As x y t) = [atom y]"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
| "bn3 (As x y t) = [atom x, atom y]"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
thm foo.distinct
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
thm foo.induct
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
thm foo.inducts
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
thm foo.exhaust
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
thm foo.fv_defs
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
thm foo.bn_defs
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
thm foo.perm_simps
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
thm foo.eq_iff
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
thm foo.fv_bn_eqvt
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
thm foo.size_eqvt
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
thm foo.supports
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
thm foo.fsupp
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
thm foo.supp
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
thm foo.fresh
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
2500
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    45
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    46
lemmas permute_bn1 = permute_bn1_raw.simps[quot_lifted]
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    47
lemmas permute_bn2 = permute_bn2_raw.simps[quot_lifted]
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    48
lemmas permute_bn3 = permute_bn3_raw.simps[quot_lifted]
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    49
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    50
lemma uu1:
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    51
  shows "alpha_bn1 as (permute_bn1 p as)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    52
apply(induct as rule: foo.inducts(2))
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    53
apply(auto)[6]
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    54
apply(simp add: permute_bn1)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    55
apply(simp add: foo.eq_iff)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    56
done
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    57
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    58
lemma uu2:
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    59
  shows "alpha_bn2 as (permute_bn2 p as)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    60
apply(induct as rule: foo.inducts(2))
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    61
apply(auto)[6]
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    62
apply(simp add: permute_bn2)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    63
apply(simp add: foo.eq_iff)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    64
done
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    65
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    66
lemma uu3:
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    67
  shows "alpha_bn3 as (permute_bn3 p as)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    68
apply(induct as rule: foo.inducts(2))
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    69
apply(auto)[6]
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    70
apply(simp add: permute_bn3)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    71
apply(simp add: foo.eq_iff)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    72
done
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    73
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    74
lemma tt1:
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    75
  shows "(p \<bullet> bn1 as) = bn1 (permute_bn1 p as)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    76
apply(induct as rule: foo.inducts(2))
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    77
apply(auto)[6]
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    78
apply(simp add: permute_bn1 foo.bn_defs)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    79
apply(simp add: atom_eqvt)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    80
done
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    81
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    82
lemma tt2:
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    83
  shows "(p \<bullet> bn2 as) = bn2 (permute_bn2 p as)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    84
apply(induct as rule: foo.inducts(2))
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    85
apply(auto)[6]
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    86
apply(simp add: permute_bn2 foo.bn_defs)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    87
apply(simp add: atom_eqvt)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    88
done
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    89
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    90
lemma tt3:
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    91
  shows "(p \<bullet> bn3 as) = bn3 (permute_bn3 p as)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    92
apply(induct as rule: foo.inducts(2))
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    93
apply(auto)[6]
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    94
apply(simp add: permute_bn3 foo.bn_defs)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    95
apply(simp add: atom_eqvt)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    96
done
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    97
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    98
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
    99
lemma strong_exhaust1:
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   100
  fixes c::"'a::fs"
2503
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   101
  assumes "\<And>name. y = Var name \<Longrightarrow> P" 
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   102
  and     "\<And>trm1 trm2. y = App trm1 trm2 \<Longrightarrow> P"
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   103
  and     "\<And>name trm. \<lbrakk>{atom name} \<sharp>* c; y = Lam name trm\<rbrakk> \<Longrightarrow> P" 
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   104
  and     "\<And>assn trm. \<lbrakk>set (bn1 assn) \<sharp>* c; y = Let1 assn trm\<rbrakk> \<Longrightarrow> P"
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   105
  and     "\<And>assn trm. \<lbrakk>set (bn2 assn) \<sharp>* c; y = Let2 assn trm\<rbrakk> \<Longrightarrow> P"
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   106
  and     "\<And>assn trm. \<lbrakk>set (bn3 assn) \<sharp>* c; y = Let3 assn trm\<rbrakk> \<Longrightarrow> P"
2500
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   107
  shows "P"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   108
apply(rule_tac y="y" in foo.exhaust(1))
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   109
apply(rule assms(1))
2503
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   110
apply(assumption)
2500
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   111
apply(rule assms(2))
2503
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   112
apply(assumption)
2500
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   113
apply(subgoal_tac "\<exists>q. (q \<bullet> {atom name}) \<sharp>* c \<and> supp (Lam name trm) \<sharp>* q")
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   114
apply(erule exE)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   115
apply(erule conjE)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   116
apply(rule assms(3))
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   117
apply(perm_simp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   118
apply(assumption)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   119
apply(drule supp_perm_eq[symmetric])
2503
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   120
apply(perm_simp)
2500
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   121
apply(simp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   122
apply(rule at_set_avoiding2)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   123
apply(simp add: finite_supp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   124
apply(simp add: finite_supp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   125
apply(simp add: finite_supp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   126
apply(simp add: foo.fresh fresh_star_def)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   127
apply(subgoal_tac "\<exists>q. (q \<bullet> (set (bn1 assg))) \<sharp>* c \<and> supp ([bn1 assg]lst.trm) \<sharp>* q")
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   128
apply(erule exE)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   129
apply(erule conjE)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   130
apply(rule assms(4))
2503
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   131
apply(perm_simp add: tt1)
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   132
apply(assumption)
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   133
apply(drule supp_perm_eq[symmetric])
2500
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   134
apply(simp add: foo.eq_iff)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   135
apply(simp add: tt1 uu1)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   136
apply(rule at_set_avoiding2)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   137
apply(simp add: finite_supp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   138
apply(simp add: finite_supp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   139
apply(simp add: finite_supp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   140
apply(simp add: Abs_fresh_star)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   141
apply(subgoal_tac "\<exists>q. (q \<bullet> (set (bn2 assg))) \<sharp>* c \<and> supp ([bn2 assg]lst.trm) \<sharp>* q")
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   142
apply(erule exE)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   143
apply(erule conjE)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   144
apply(rule assms(5))
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   145
apply(simp add: set_eqvt)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   146
apply(simp add: tt2)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   147
apply(simp add: foo.eq_iff)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   148
apply(drule supp_perm_eq[symmetric])
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   149
apply(simp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   150
apply(simp add: tt2 uu2)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   151
apply(rule at_set_avoiding2)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   152
apply(simp add: finite_supp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   153
apply(simp add: finite_supp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   154
apply(simp add: finite_supp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   155
apply(simp add: Abs_fresh_star)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   156
apply(subgoal_tac "\<exists>q. (q \<bullet> (set (bn3 assg))) \<sharp>* c \<and> supp ([bn3 assg]lst.trm) \<sharp>* q")
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   157
apply(erule exE)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   158
apply(erule conjE)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   159
apply(rule assms(6))
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   160
apply(simp add: set_eqvt)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   161
apply(simp add: tt3)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   162
apply(simp add: foo.eq_iff)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   163
apply(drule supp_perm_eq[symmetric])
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   164
apply(simp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   165
apply(simp add: tt3 uu3)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   166
apply(rule at_set_avoiding2)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   167
apply(simp add: finite_supp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   168
apply(simp add: finite_supp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   169
apply(simp add: finite_supp)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   170
apply(simp add: Abs_fresh_star)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   171
done
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   172
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   173
lemma strong_exhaust2:
2503
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   174
  assumes "\<And>x y t. as = As x y t \<Longrightarrow> P" 
2500
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   175
  shows "P"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   176
apply(rule_tac y="as" in foo.exhaust(2))
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   177
apply(rule assms(1))
2503
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   178
apply(assumption)
2500
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   179
done
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   180
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   181
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   182
lemma 
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   183
  fixes t::trm
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   184
  and   as::assg
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   185
  and   c::"'a::fs"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   186
  assumes a1: "\<And>x c. P1 c (Var x)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   187
  and     a2: "\<And>t1 t2 c. \<lbrakk>\<And>d. P1 d t1; \<And>d. P1 d t2\<rbrakk> \<Longrightarrow> P1 c (App t1 t2)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   188
  and     a3: "\<And>x t c. \<lbrakk>{atom x} \<sharp>* c; \<And>d. P1 d t\<rbrakk> \<Longrightarrow> P1 c (Lam x t)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   189
  and     a4: "\<And>as t c. \<lbrakk>set (bn1 as) \<sharp>* c; \<And>d. P2 d as; \<And>d. P1 d t\<rbrakk> \<Longrightarrow> P1 c (Let1 as t)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   190
  and     a5: "\<And>as t c. \<lbrakk>set (bn2 as) \<sharp>* c; \<And>d. P2 d as; \<And>d. P1 d t\<rbrakk> \<Longrightarrow> P1 c (Let2 as t)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   191
  and     a6: "\<And>as t c. \<lbrakk>set (bn3 as) \<sharp>* c; \<And>d. P2 d as; \<And>d. P1 d t\<rbrakk> \<Longrightarrow> P1 c (Let3 as t)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   192
  and     a7: "\<And>x y t c. \<And>d. P1 d t \<Longrightarrow> P2 c (As x y t)"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   193
  shows "P1 c t" "P2 c as"
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   194
using assms
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   195
apply(induction_schema)
2503
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   196
apply(rule_tac y="t" and c="c" in strong_exhaust1)
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   197
apply(simp_all)[6]
2500
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   198
apply(rule_tac as="as" in strong_exhaust2)
2503
cc5d23547341 simplified exhaust proofs
Christian Urban <urbanc@in.tum.de>
parents: 2500
diff changeset
   199
apply(simp)
2500
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   200
apply(relation "measure (sum_case (\<lambda>y. size (snd y)) (\<lambda>z. size (snd z)))")
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   201
apply(simp_all add: foo.size)
3b6a70e73006 worked example Foo1 with induct_schema
Christian Urban <urbanc@in.tum.de>
parents: 2494
diff changeset
   202
done
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
end
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207