src/HOL/Nominal/Examples/SOS.thy
author urbanc
Mon, 19 Mar 2007 19:28:27 +0100
changeset 22472 bfd9c0fd70b1
parent 22447 013dbd8234f0
child 22502 baee64dbe8ea
permissions -rw-r--r--
tuned the proof
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
     1
(* "$Id$" *)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
     2
(*                                                   *)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
     3
(* Formalisation of some typical SOS-proofs          *)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
     4
(*                                                   *) 
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
     5
(* This work arose from challenge suggested by Adam  *)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
     6
(* Chlipala suggested on the POPLmark mailing list.  *)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
     7
(*                                                   *) 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
     8
(* We thank Nick Benton for helping us with the      *) 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
     9
(* termination-proof for evaluation.                 *)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    10
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    11
theory SOS
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    12
  imports "Nominal"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    13
begin
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    14
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    15
atom_decl name 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    16
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    17
nominal_datatype data = 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    18
    DNat
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    19
  | DProd "data" "data"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    20
  | DSum "data" "data"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    21
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    22
nominal_datatype ty = 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    23
    Data "data"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    24
  | Arrow "ty" "ty" ("_\<rightarrow>_" [100,100] 100)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    25
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    26
nominal_datatype trm = 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    27
    Var "name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    28
  | Lam "\<guillemotleft>name\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    29
  | App "trm" "trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    30
  | Const "nat"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    31
  | Pr "trm" "trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    32
  | Fst "trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    33
  | Snd "trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    34
  | InL "trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    35
  | InR "trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    36
  | Case "trm" "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" ("Case _ of inl _ \<rightarrow> _ | inr _ \<rightarrow> _" [100,100,100,100,100] 100)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    37
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    38
lemma in_eqvt[eqvt]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    39
  fixes pi::"name prm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    40
  and   x::"'a::pt_name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    41
  assumes "x\<in>X"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    42
  shows "pi\<bullet>x \<in> pi\<bullet>X"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    43
  using assms by (perm_simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    44
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    45
lemma perm_data[simp]: 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    46
  fixes D::"data"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    47
  and   pi::"name prm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    48
  shows "pi\<bullet>D = D"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    49
  by (induct D rule: data.induct_weak) (simp_all)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    50
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    51
lemma perm_ty[simp]: 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    52
  fixes T::"ty"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    53
  and   pi::"name prm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    54
  shows "pi\<bullet>T = T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    55
  by (induct T rule: ty.induct_weak) (simp_all)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    56
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    57
lemma fresh_ty[simp]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    58
  fixes x::"name" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    59
  and   T::"ty"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    60
  shows "x\<sharp>T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    61
  by (simp add: fresh_def supp_def)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    62
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    63
text {* substitution *}
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    64
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    65
fun
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    66
  lookup :: "(name\<times>trm) list \<Rightarrow> name \<Rightarrow> trm"   
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    67
where
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    68
  "lookup [] x        = Var x"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    69
  "lookup ((y,e)#\<theta>) x = (if x=y then e else lookup \<theta> x)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    70
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    71
lemma lookup_eqvt:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    72
  fixes pi::"name prm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    73
  and   \<theta>::"(name\<times>trm) list"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    74
  and   X::"name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    75
  shows "pi\<bullet>(lookup \<theta> X) = lookup (pi\<bullet>\<theta>) (pi\<bullet>X)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    76
by (induct \<theta>, auto simp add: perm_bij)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    77
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    78
lemma lookup_fresh:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    79
  fixes z::"name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    80
  assumes "z\<sharp>\<theta>" and "z\<sharp>x"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    81
  shows "z \<sharp>lookup \<theta> x"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    82
using assms 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    83
by (induct rule: lookup.induct) (auto simp add: fresh_list_cons)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    84
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    85
lemma lookup_fresh':
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    86
  assumes "z\<sharp>\<theta>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    87
  shows "lookup \<theta> z = Var z"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    88
using assms 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    89
by (induct rule: lookup.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    90
   (auto simp add: fresh_list_cons fresh_prod fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    91
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
    92
text {* Parallel Substitution *}
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
    93
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    94
consts
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    95
  psubst :: "(name\<times>trm) list \<Rightarrow> trm \<Rightarrow> trm"  ("_<_>" [95,95] 105)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    96
 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    97
nominal_primrec
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    98
  "\<theta><(Var x)> = (lookup \<theta> x)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    99
  "\<theta><(App e\<^isub>1 e\<^isub>2)> = App (\<theta><e\<^isub>1>) (\<theta><e\<^isub>2>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   100
  "x\<sharp>\<theta> \<Longrightarrow> \<theta><(Lam [x].e)> = Lam [x].(\<theta><e>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   101
  "\<theta><(Const n)> = Const n"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   102
  "\<theta><(Pr e\<^isub>1 e\<^isub>2)> = Pr (\<theta><e\<^isub>1>) (\<theta><e\<^isub>2>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   103
  "\<theta><(Fst e)> = Fst (\<theta><e>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   104
  "\<theta><(Snd e)> = Snd (\<theta><e>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   105
  "\<theta><(InL e)> = InL (\<theta><e>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   106
  "\<theta><(InR e)> = InR (\<theta><e>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   107
  "\<lbrakk>y\<noteq>x; x\<sharp>(e,e\<^isub>2,\<theta>); y\<sharp>(e,e\<^isub>1,\<theta>)\<rbrakk> 
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   108
   \<Longrightarrow> \<theta><Case e of inl x \<rightarrow> e\<^isub>1 | inr y \<rightarrow> e\<^isub>2> = (Case (\<theta><e>) of inl x \<rightarrow> (\<theta><e\<^isub>1>) | inr y \<rightarrow> (\<theta><e\<^isub>2>))"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   109
apply(finite_guess add: lookup_eqvt)+
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   110
apply(rule TrueI)+
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   111
apply(simp add: abs_fresh)+
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   112
apply(fresh_guess add: fs_name1 lookup_eqvt)+
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   113
done
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   114
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   115
lemma psubst_eqvt[eqvt]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   116
  fixes pi::"name prm" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   117
  and   t::"trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   118
  shows "pi\<bullet>(\<theta><t>) = (pi\<bullet>\<theta>)<(pi\<bullet>t)>"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   119
by (nominal_induct t avoiding: \<theta> rule: trm.induct)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   120
   (perm_simp add: fresh_bij lookup_eqvt)+
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   121
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   122
lemma fresh_psubst: 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   123
  fixes z::"name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   124
  and   t::"trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   125
  assumes "z\<sharp>t" and "z\<sharp>\<theta>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   126
  shows "z\<sharp>(\<theta><t>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   127
using assms
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   128
by (nominal_induct t avoiding: z \<theta> t rule: trm.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   129
   (auto simp add: abs_fresh lookup_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   130
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   131
abbreviation 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   132
 subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   133
  where "t[x::=t']  \<equiv> ([(x,t')])<t>" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   134
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   135
lemma subst[simp]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   136
  shows "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   137
  and   "(App t\<^isub>1 t\<^isub>2)[y::=t'] = App (t\<^isub>1[y::=t']) (t\<^isub>2[y::=t'])"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   138
  and   "x\<sharp>(y,t') \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   139
  and   "(Const n)[y::=t'] = Const n"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   140
  and   "(Pr e\<^isub>1 e\<^isub>2)[y::=t'] = Pr (e\<^isub>1[y::=t']) (e\<^isub>2[y::=t'])"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   141
  and   "(Fst e)[y::=t'] = Fst (e[y::=t'])"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   142
  and   "(Snd e)[y::=t'] = Snd (e[y::=t'])"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   143
  and   "(InL e)[y::=t'] = InL (e[y::=t'])"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   144
  and   "(InR e)[y::=t'] = InR (e[y::=t'])"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   145
  and   "\<lbrakk>z\<noteq>x; x\<sharp>(y,e,e\<^isub>2,t'); z\<sharp>(y,e,e\<^isub>1,t')\<rbrakk> 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   146
         \<Longrightarrow> (Case e of inl x \<rightarrow> e\<^isub>1 | inr z \<rightarrow> e\<^isub>2)[y::=t'] =
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   147
                   (Case (e[y::=t']) of inl x \<rightarrow> (e\<^isub>1[y::=t']) | inr z \<rightarrow> (e\<^isub>2[y::=t']))"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   148
by (simp_all add: fresh_list_cons fresh_list_nil)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   149
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   150
lemma subst_eqvt[eqvt]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   151
  fixes pi::"name prm" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   152
  and   t::"trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   153
  shows "pi\<bullet>(t[x::=t']) = (pi\<bullet>t)[(pi\<bullet>x)::=(pi\<bullet>t')]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   154
  by (nominal_induct t avoiding: x t' rule: trm.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   155
     (perm_simp add: fresh_bij)+
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   156
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   157
lemma fresh_subst:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   158
  fixes z::"name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   159
  and   t\<^isub>1::"trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   160
  and   t2::"trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   161
  assumes "z\<sharp>t\<^isub>1" and "z\<sharp>t\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   162
  shows "z\<sharp>t\<^isub>1[y::=t\<^isub>2]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   163
using assms 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   164
by (nominal_induct t\<^isub>1 avoiding: z y t\<^isub>2 rule: trm.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   165
   (auto simp add: abs_fresh fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   166
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   167
lemma fresh_subst':
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   168
  fixes z::"name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   169
  and   t\<^isub>1::"trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   170
  and   t2::"trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   171
  assumes "z\<sharp>[y].t\<^isub>1" and "z\<sharp>t\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   172
  shows "z\<sharp>t\<^isub>1[y::=t\<^isub>2]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   173
using assms 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   174
by (nominal_induct t\<^isub>1 avoiding: y t\<^isub>2 z  rule: trm.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   175
   (auto simp add: abs_fresh fresh_nat fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   176
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   177
lemma forget: 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   178
  fixes x::"name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   179
  and   L::"trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   180
  assumes "x\<sharp>L" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   181
  shows "L[x::=P] = L"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   182
  using assms
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   183
  by (nominal_induct L avoiding: x P rule: trm.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   184
     (auto simp add: fresh_atm abs_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   185
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   186
lemma psubst_empty[simp]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   187
  shows "[]<t> = t"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   188
  by (nominal_induct t rule: trm.induct, auto simp add:fresh_list_nil)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   189
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   190
lemma psubst_subst_psubst:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   191
assumes h:"x\<sharp>\<theta>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   192
shows "\<theta><e>[x::=e'] = ((x,e')#\<theta>)<e>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   193
using h
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   194
apply(nominal_induct e avoiding: \<theta> x e' rule: trm.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   195
apply(auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh' fresh_psubst)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   196
done
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   197
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   198
lemma fresh_subst_fresh:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   199
    assumes "a\<sharp>e"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   200
    shows "a\<sharp>t[a::=e]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   201
using assms 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   202
by (nominal_induct t avoiding: a e rule: trm.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   203
   (auto simp add: fresh_atm abs_fresh fresh_nat) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   204
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   205
text {* Typing-Judgements *}
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   206
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   207
inductive2
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   208
  valid :: "(name \<times> 'a::pt_name) list \<Rightarrow> bool"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   209
where
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   210
    v_nil[intro]:  "valid []"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   211
  | v_cons[intro]: "\<lbrakk>valid \<Gamma>;x\<sharp>\<Gamma>\<rbrakk> \<Longrightarrow> valid ((x,T)#\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   212
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   213
nominal_inductive valid
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   214
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   215
inductive_cases2  
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   216
  valid_cons_inv_auto[elim]:"valid ((x,T)#\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   217
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   218
abbreviation
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   219
  "sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ \<lless> _" [55,55] 55)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   220
where
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   221
  "\<Gamma>\<^isub>1 \<lless> \<Gamma>\<^isub>2 \<equiv> \<forall>x T. (x,T)\<in>set \<Gamma>\<^isub>1 \<longrightarrow> (x,T)\<in>set \<Gamma>\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   222
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   223
lemma type_unicity_in_context:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   224
  assumes asm1: "(x,t\<^isub>2) \<in> set ((x,t\<^isub>1)#\<Gamma>)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   225
  and     asm2: "valid ((x,t\<^isub>1)#\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   226
  shows "t\<^isub>1=t\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   227
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   228
  from asm2 have "x\<sharp>\<Gamma>" by (cases, auto)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   229
  then have "(x,t\<^isub>2) \<notin> set \<Gamma>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   230
    by (induct \<Gamma>) (auto simp add: fresh_list_cons fresh_prod fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   231
  then have "(x,t\<^isub>2) = (x,t\<^isub>1)" using asm1 by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   232
  then show "t\<^isub>1 = t\<^isub>2" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   233
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   234
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   235
lemma case_distinction_on_context:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   236
  fixes \<Gamma>::"(name \<times> ty) list"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   237
  assumes asm1: "valid ((m,t)#\<Gamma>)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   238
  and     asm2: "(n,U) \<in> set ((m,T)#\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   239
  shows "(n,U) = (m,T) \<or> ((n,U) \<in> set \<Gamma> \<and> n \<noteq> m)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   240
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   241
from asm2 have "(n,U) \<in> set [(m,T)] \<or> (n,U) \<in> set \<Gamma>" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   242
moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   243
{ assume eq: "m=n"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   244
  assume "(n,U) \<in> set \<Gamma>" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   245
  then have "\<not> n\<sharp>\<Gamma>" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   246
    by (induct \<Gamma>) (auto simp add: fresh_list_cons fresh_prod fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   247
  moreover have "m\<sharp>\<Gamma>" using asm1 by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   248
  ultimately have False using eq by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   249
}
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   250
ultimately show ?thesis by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   251
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   252
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   253
inductive2
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   254
  typing :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   255
where
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   256
  t_Var[intro]:   "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   257
| t_App[intro]:   "\<lbrakk>\<Gamma> \<turnstile> e\<^isub>1 : T\<^isub>1\<rightarrow>T\<^isub>2; \<Gamma> \<turnstile> e\<^isub>2 : T\<^isub>1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   258
| t_Lam[intro]:   "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^isub>1)#\<Gamma> \<turnstile> e : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].e : T\<^isub>1\<rightarrow>T\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   259
| t_Const[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Const n : Data(DNat)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   260
| t_Pr[intro]:    "\<lbrakk>\<Gamma> \<turnstile> e\<^isub>1 : Data(S\<^isub>1); \<Gamma> \<turnstile> e\<^isub>2 : Data(S\<^isub>2)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Pr e\<^isub>1 e\<^isub>2 : Data (DProd S\<^isub>1 S\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   261
| t_Fst[intro]:   "\<lbrakk>\<Gamma> \<turnstile> e : Data(DProd S\<^isub>1 S\<^isub>2)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Fst e : Data(S\<^isub>1)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   262
| t_Snd[intro]:   "\<lbrakk>\<Gamma> \<turnstile> e : Data(DProd S\<^isub>1 S\<^isub>2)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Snd e : Data(S\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   263
| t_InL[intro]:   "\<lbrakk>\<Gamma> \<turnstile> e : Data(S\<^isub>1)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> InL e : Data(DSum S\<^isub>1 S\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   264
| t_InR[intro]:   "\<lbrakk>\<Gamma> \<turnstile> e : Data(S\<^isub>2)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> InR e : Data(DSum S\<^isub>1 S\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   265
| t_Case[intro]:  "\<lbrakk>x\<^isub>1\<sharp>(\<Gamma>,e,e\<^isub>2,x\<^isub>2); x\<^isub>2\<sharp>(\<Gamma>,e,e\<^isub>1,x\<^isub>1); \<Gamma> \<turnstile> e: Data(DSum S\<^isub>1 S\<^isub>2); 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   266
                   (x\<^isub>1,Data(S\<^isub>1))#\<Gamma> \<turnstile> e\<^isub>1 : T; (x\<^isub>2,Data(S\<^isub>2))#\<Gamma> \<turnstile> e\<^isub>2 : T\<rbrakk> 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   267
                   \<Longrightarrow> \<Gamma> \<turnstile> (Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2) : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   268
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   269
lemma typing_implies_valid:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   270
  assumes "\<Gamma> \<turnstile> t : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   271
  shows "valid \<Gamma>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   272
  using assms
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   273
  by (induct) (auto)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   274
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   275
lemma typing_eqvt:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   276
  fixes pi::"name prm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   277
  assumes a: "\<Gamma> \<turnstile> t : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   278
  shows "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>t) : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   279
using a
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   280
apply(induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   281
apply(auto simp add: fresh_bij set_eqvt valid_eqvt) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   282
apply(rule t_Var)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   283
apply(drule valid_eqvt)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   284
apply(assumption)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   285
apply(drule in_eqvt)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   286
apply(simp add: set_eqvt)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   287
done
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   288
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   289
declare trm.inject [simp add]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   290
declare ty.inject  [simp add]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   291
declare data.inject [simp add]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   292
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   293
inductive_cases2 t_Lam_inv_auto[elim]: "\<Gamma> \<turnstile> Lam [x].t : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   294
inductive_cases2 t_Var_inv_auto[elim]: "\<Gamma> \<turnstile> Var x : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   295
inductive_cases2 t_App_inv_auto[elim]: "\<Gamma> \<turnstile> App x y : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   296
inductive_cases2 t_Const_inv_auto[elim]: "\<Gamma> \<turnstile> Const n : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   297
inductive_cases2 t_Fst_inv_auto[elim]: "\<Gamma> \<turnstile> Fst x : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   298
inductive_cases2 t_Snd_inv_auto[elim]: "\<Gamma> \<turnstile> Snd x : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   299
inductive_cases2 t_InL_inv_auto[elim]: "\<Gamma> \<turnstile> InL x : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   300
inductive_cases2 t_InL_inv_auto'[elim]: "\<Gamma> \<turnstile> InL x : Data (DSum T\<^isub>1 T2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   301
inductive_cases2 t_InR_inv_auto[elim]: "\<Gamma> \<turnstile> InR x : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   302
inductive_cases2 t_InR_inv_auto'[elim]: "\<Gamma> \<turnstile> InR x : Data (DSum T\<^isub>1 T2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   303
inductive_cases2 t_Pr_inv_auto[elim]:  "\<Gamma> \<turnstile> Pr x y : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   304
inductive_cases2 t_Pr_inv_auto'[elim]: "\<Gamma> \<turnstile> Pr e\<^isub>1 e\<^isub>2 : Data (DProd \<sigma>1 \<sigma>\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   305
inductive_cases2 t_Case_inv_auto[elim]: "\<Gamma> \<turnstile> Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2 : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   306
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   307
declare trm.inject [simp del]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   308
declare ty.inject [simp del]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   309
declare data.inject [simp del]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   310
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   311
lemma typing_induct_strong
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   312
  [consumes 1, case_names t_Var t_App t_Lam t_Const t_Pr t_Fst t_Snd t_InL t_InR t_Case]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   313
    fixes  P::"'a::fs_name \<Rightarrow> (name\<times>ty) list \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow>bool"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   314
    and    x :: "'a::fs_name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   315
    assumes a: "\<Gamma> \<turnstile> e : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   316
    and a1: "\<And>\<Gamma> x T c. \<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> P c \<Gamma> (Var x) T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   317
    and a2: "\<And>\<Gamma> e\<^isub>1 T\<^isub>1 T\<^isub>2 e\<^isub>2 c. \<lbrakk>\<And>c. P c \<Gamma> e\<^isub>1 (T\<^isub>1\<rightarrow>T\<^isub>2); \<And>c. P c \<Gamma> e\<^isub>2 T\<^isub>1\<rbrakk> \<Longrightarrow> P c \<Gamma> (App e\<^isub>1 e\<^isub>2) T\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   318
    and a3: "\<And>x \<Gamma> T\<^isub>1 t T\<^isub>2 c. \<lbrakk>x\<sharp>(\<Gamma>,c); \<And>c. P c ((x,T\<^isub>1)#\<Gamma>) t T\<^isub>2\<rbrakk> \<Longrightarrow> P c \<Gamma> (Lam [x].t) (T\<^isub>1\<rightarrow>T\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   319
    and a4: "\<And>\<Gamma> n c. valid \<Gamma> \<Longrightarrow> P c \<Gamma> (Const n) (Data DNat)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   320
    and a5: "\<And>\<Gamma> e\<^isub>1 S\<^isub>1 e\<^isub>2 S\<^isub>2 c. \<lbrakk>\<And>c. P c \<Gamma> e\<^isub>1 (Data S\<^isub>1); \<And>c. P c \<Gamma> e\<^isub>2 (Data S\<^isub>2)\<rbrakk>
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   321
            \<Longrightarrow> P c \<Gamma> (Pr e\<^isub>1 e\<^isub>2) (Data (DProd S\<^isub>1 S\<^isub>2))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   322
    and a6: "\<And>\<Gamma> e S\<^isub>1 S\<^isub>2 c. \<lbrakk>\<And>c. P c \<Gamma> e (Data (DProd S\<^isub>1 S\<^isub>2))\<rbrakk> \<Longrightarrow> P c \<Gamma> (Fst e) (Data S\<^isub>1)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   323
    and a7: "\<And>\<Gamma> e S\<^isub>1 S\<^isub>2 c. \<lbrakk>\<And>c. P c \<Gamma> e (Data (DProd S\<^isub>1 S\<^isub>2))\<rbrakk> \<Longrightarrow> P c \<Gamma> (Snd e) (Data S\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   324
    and a8: "\<And>\<Gamma> e S\<^isub>1 S\<^isub>2 c. \<lbrakk>\<And>c. P c \<Gamma> e (Data S\<^isub>1)\<rbrakk> \<Longrightarrow> P c \<Gamma> (InL e) (Data (DSum S\<^isub>1 S\<^isub>2))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   325
    and a9: "\<And>\<Gamma> e S\<^isub>2 S\<^isub>1 c. \<lbrakk>\<And>c. P c \<Gamma> e (Data S\<^isub>2)\<rbrakk> \<Longrightarrow> P c \<Gamma> (InR e) (Data (DSum S\<^isub>1 S\<^isub>2))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   326
    and a10:"\<And>x\<^isub>1 \<Gamma> e e\<^isub>2 x\<^isub>2 e\<^isub>1 S\<^isub>1 S\<^isub>2 T c.
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   327
             \<lbrakk>x\<^isub>1\<sharp>(\<Gamma>,e,e\<^isub>2,x\<^isub>2,c); x\<^isub>2\<sharp>(\<Gamma>,e,e\<^isub>1,x\<^isub>1,c); 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   328
             \<And>c. P c \<Gamma> e (Data (DSum S\<^isub>1 S\<^isub>2)); 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   329
             \<And>c. P c ((x\<^isub>1,Data S\<^isub>1)#\<Gamma>) e\<^isub>1 T; 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   330
             \<And>c. P c ((x\<^isub>2,Data S\<^isub>2)#\<Gamma>) e\<^isub>2 T\<rbrakk>
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   331
             \<Longrightarrow> P c \<Gamma> (Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2) T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   332
    shows "P c \<Gamma> e T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   333
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   334
  from a have "\<And>(pi::name prm) c. P c (pi\<bullet>\<Gamma>) (pi\<bullet>e) T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   335
  proof (induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   336
    case (t_Var \<Gamma> x T pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   337
    have "valid \<Gamma>" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   338
    then have "valid (pi\<bullet>\<Gamma>)" by (simp only: eqvt)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   339
    moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   340
    have "(x,T)\<in>set \<Gamma>" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   341
    then have "pi\<bullet>(x,T)\<in>pi\<bullet>(set \<Gamma>)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst])  
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   342
    then have "(pi\<bullet>x,T)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: eqvt)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   343
    ultimately show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Var x)) T" using a1 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   344
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   345
    case (t_App \<Gamma> e\<^isub>1 T\<^isub>1 T\<^isub>2 e\<^isub>2 pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   346
    thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(App e\<^isub>1 e\<^isub>2)) T\<^isub>2" using a2 by (simp, blast)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   347
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   348
    case (t_Lam x \<Gamma> T\<^isub>1 t T\<^isub>2 pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   349
    obtain y::"name" where fs: "y\<sharp>(pi\<bullet>x,pi\<bullet>\<Gamma>,pi\<bullet>t,c)" by (erule exists_fresh[OF fs_name1])
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   350
    let ?sw = "[(pi\<bullet>x,y)]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   351
    let ?pi' = "?sw@pi"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   352
    have f0: "x\<sharp>\<Gamma>" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   353
    have f1: "(pi\<bullet>x)\<sharp>(pi\<bullet>\<Gamma>)" using f0 by (simp add: fresh_bij)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   354
    have f2: "y\<sharp>?pi'\<bullet>\<Gamma>" by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   355
    have ih1: "\<And>c. P c (?pi'\<bullet>((x,T\<^isub>1)#\<Gamma>)) (?pi'\<bullet>t) T\<^isub>2" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   356
    then have "P c (?pi'\<bullet>\<Gamma>) (Lam [y].(?pi'\<bullet>t)) (T\<^isub>1\<rightarrow>T\<^isub>2)" using fs f2 a3 by (simp add: calc_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   357
    then have "P c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>t))) (T\<^isub>1\<rightarrow>T\<^isub>2)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   358
      by (simp del: append_Cons add: calc_atm pt_name2)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   359
    moreover have "(?sw\<bullet>pi\<bullet>\<Gamma>) = (pi\<bullet>\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   360
      by (rule perm_fresh_fresh) (simp_all add: fs f1)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   361
    moreover have "(?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>t))) = Lam [(pi\<bullet>x)].(pi\<bullet>t)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   362
      by (rule perm_fresh_fresh) (simp_all add: fs f1 fresh_atm abs_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   363
    ultimately show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Lam [x].t)) (T\<^isub>1\<rightarrow>T\<^isub>2)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   364
      by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   365
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   366
    case (t_Const \<Gamma> n pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   367
    thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Const n)) (Data DNat)" using a4 by (simp, blast intro: eqvt)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   368
  next 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   369
    case (t_Pr \<Gamma> e\<^isub>1 S\<^isub>1 e\<^isub>2 S\<^isub>2 pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   370
    thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Pr e\<^isub>1 e\<^isub>2)) (Data (DProd S\<^isub>1 S\<^isub>2))" using a5 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   371
      by (simp)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   372
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   373
    case (t_Fst \<Gamma> e S\<^isub>1 S\<^isub>2 pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   374
    thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Fst e)) (Data S\<^isub>1)" using a6 by (simp, blast)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   375
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   376
    case (t_Snd \<Gamma> e S\<^isub>1 S\<^isub>2 pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   377
    thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Snd e)) (Data S\<^isub>2)" using a7 by (simp, blast)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   378
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   379
    case (t_InL \<Gamma> e S\<^isub>1 S\<^isub>2 pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   380
    thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(InL e)) (Data (DSum S\<^isub>1 S\<^isub>2))" using a8 by (simp)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   381
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   382
    case (t_InR \<Gamma> e S\<^isub>2 S\<^isub>1 pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   383
    thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(InR e)) (Data (DSum S\<^isub>1 S\<^isub>2))" using a9 by (simp)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   384
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   385
    case (t_Case x\<^isub>1 \<Gamma> e e\<^isub>2 x\<^isub>2 e\<^isub>1 S\<^isub>1 S\<^isub>2 T pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   386
    obtain y\<^isub>1::"name" where fs1: "y\<^isub>1\<sharp>(pi\<bullet>x\<^isub>1,pi\<bullet>x\<^isub>2,pi\<bullet>e,pi\<bullet>e\<^isub>1,pi\<bullet>e\<^isub>2,pi\<bullet>\<Gamma>,c)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   387
      by (erule exists_fresh[OF fs_name1])
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   388
    obtain y\<^isub>2::"name" where fs2: "y\<^isub>2\<sharp>(pi\<bullet>x\<^isub>1,pi\<bullet>x\<^isub>2,pi\<bullet>e,pi\<bullet>e\<^isub>1,pi\<bullet>e\<^isub>2,pi\<bullet>\<Gamma>,c,y\<^isub>1)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   389
      by (erule exists_fresh[OF fs_name1])
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   390
    let ?sw1 = "[(pi\<bullet>x\<^isub>1,y\<^isub>1)]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   391
    let ?sw2 = "[(pi\<bullet>x\<^isub>2,y\<^isub>2)]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   392
    let ?pi' = "?sw2@?sw1@pi"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   393
    have f01: "x\<^isub>1\<sharp>(\<Gamma>,e,e\<^isub>2,x\<^isub>2)" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   394
    have f11: "(pi\<bullet>x\<^isub>1)\<sharp>(pi\<bullet>\<Gamma>,pi\<bullet>e,pi\<bullet>e\<^isub>2,pi\<bullet>x\<^isub>2)" using f01 by (simp add: fresh_bij)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   395
    have f21: "y\<^isub>1\<sharp>(?pi'\<bullet>\<Gamma>,?pi'\<bullet>e,?pi'\<bullet>e\<^isub>2)" using f01 fs1 fs2
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   396
      by (simp add: fresh_atm fresh_prod fresh_left calc_atm pt_name2 perm_pi_simp)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   397
    have f02: "x\<^isub>2\<sharp>(\<Gamma>,e,e\<^isub>1,x\<^isub>1)" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   398
    have f12: "(pi\<bullet>x\<^isub>2)\<sharp>(pi\<bullet>\<Gamma>,pi\<bullet>e,pi\<bullet>e\<^isub>1,pi\<bullet>x\<^isub>1)" using f02 by (simp add: fresh_bij)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   399
    have f22: "y\<^isub>2\<sharp>(?pi'\<bullet>\<Gamma>,?pi'\<bullet>e,?pi'\<bullet>e\<^isub>1)" using f02 fs1 fs2
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   400
      by (auto simp add: fresh_atm fresh_prod fresh_left calc_atm pt_name2 perm_pi_simp)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   401
    have ih1: "\<And>c. P c (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>e) (Data (DSum S\<^isub>1 S\<^isub>2))" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   402
    moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   403
    have ih2: "\<And>c. P c (?pi'\<bullet>((x\<^isub>1,Data S\<^isub>1)#\<Gamma>)) (?pi'\<bullet>e\<^isub>1) T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   404
    then have "\<And>c. P c ((y\<^isub>1,Data S\<^isub>1)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>e\<^isub>1) T" using fs1 fs2
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   405
      by (auto simp add: calc_atm fresh_prod fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   406
    moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   407
    have ih3: "\<And>c. P c (?pi'\<bullet>((x\<^isub>2,Data S\<^isub>2)#\<Gamma>)) (?pi'\<bullet>e\<^isub>2) T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   408
    then have "\<And>c. P c ((y\<^isub>2,Data S\<^isub>2)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>e\<^isub>2) T" using fs1 fs2 f11 f12
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   409
      by (simp add: calc_atm fresh_prod fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   410
    ultimately have "P c (?pi'\<bullet>\<Gamma>) (Case (?pi'\<bullet>e) of inl y\<^isub>1 \<rightarrow> (?pi'\<bullet>e\<^isub>1) | inr y\<^isub>2 \<rightarrow> (?pi'\<bullet>e\<^isub>2)) T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   411
      using f21 f22 fs1 fs2 a10 by (force simp add: fresh_atm fresh_prod)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   412
    then have "P c (?sw2\<bullet>?sw1\<bullet>pi\<bullet>\<Gamma>) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   413
                   (?sw2\<bullet>?sw1\<bullet>(Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))) T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   414
      using fs1 fs2 f01 f02 f11 f12 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   415
      by (auto simp del: append_Cons simp add: pt_name2 fresh_atm fresh_prod calc_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   416
    moreover have "(?sw1\<bullet>pi\<bullet>\<Gamma>) = (pi\<bullet>\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   417
      by (rule perm_fresh_fresh) (simp_all add: fs1 f11)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   418
    moreover have "(?sw2\<bullet>pi\<bullet>\<Gamma>) = (pi\<bullet>\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   419
      by (rule perm_fresh_fresh) (simp_all add: fs2 f12)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   420
    moreover have "?sw1\<bullet>(Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   421
                    = (Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   422
      by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12 abs_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   423
    moreover have "?sw2\<bullet>(Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   424
                    = (Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   425
      by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12 abs_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   426
    ultimately show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2)) T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   427
      by (simp only:, simp) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   428
  qed 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   429
  then have "P c (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>e) T" by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   430
  then show "P c \<Gamma> e T" by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   431
qed 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   432
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   433
lemma t_Lam_elim [elim] : 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   434
  assumes a1:"\<Gamma> \<turnstile> Lam [x].t : T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   435
  and     a2: "x\<sharp>\<Gamma>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   436
  obtains T\<^isub>1 and T\<^isub>2 where "(x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2" and "T=T\<^isub>1\<rightarrow>T\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   437
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   438
  from a1 obtain x' t' T\<^isub>1 T\<^isub>2 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   439
    where b1: "x'\<sharp>\<Gamma>" and b2: "(x',T\<^isub>1)#\<Gamma> \<turnstile> t' : T\<^isub>2" and b3: "[x'].t' = [x].t" and b4: "T=T\<^isub>1\<rightarrow>T\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   440
    by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   441
  obtain c::"name" where "c\<sharp>(\<Gamma>,x,x',t,t')" by (erule exists_fresh[OF fs_name1])
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   442
  then have fs: "c\<sharp>\<Gamma>" "c\<noteq>x" "c\<noteq>x'" "c\<sharp>t" "c\<sharp>t'" by (simp_all add: fresh_atm[symmetric]) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   443
  then have b5: "[(x',c)]\<bullet>t'=[(x,c)]\<bullet>t" using b3 fs by (simp add: alpha')
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   444
  have "([(x,c)]\<bullet>[(x',c)]\<bullet>((x',T\<^isub>1)#\<Gamma>)) \<turnstile> ([(x,c)]\<bullet>[(x',c)]\<bullet>t') : T\<^isub>2" using b2
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   445
    by (simp only: typing_eqvt[simplified perm_ty])
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   446
  then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2" using fs b1 a2 b5 by (perm_simp add: calc_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   447
  then show ?thesis using prems b4 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   448
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   449
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   450
lemma t_Case_elim[elim] : 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   451
  assumes "\<Gamma> \<turnstile> Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2 : T" and "x\<^isub>1\<sharp>\<Gamma>" and "x\<^isub>2\<sharp>\<Gamma>" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   452
  obtains \<sigma>\<^isub>1 \<sigma>\<^isub>2 where "\<Gamma> \<turnstile> e : Data (DSum \<sigma>\<^isub>1 \<sigma>\<^isub>2)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   453
                  and "(x\<^isub>1, Data \<sigma>\<^isub>1)#\<Gamma> \<turnstile> e\<^isub>1 : T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   454
                  and "(x\<^isub>2, Data \<sigma>\<^isub>2)#\<Gamma> \<turnstile> e\<^isub>2 : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   455
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   456
  have f:"x\<^isub>1\<sharp>\<Gamma>" "x\<^isub>2\<sharp>\<Gamma>" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   457
  have "\<Gamma> \<turnstile> Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2 : T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   458
  then obtain \<sigma>\<^isub>1 \<sigma>\<^isub>2 x\<^isub>1' x\<^isub>2' e\<^isub>1' e\<^isub>2' where 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   459
    h:"\<Gamma> \<turnstile> e : Data (DSum \<sigma>\<^isub>1 \<sigma>\<^isub>2)" and 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   460
    h1:"(x\<^isub>1',Data \<sigma>\<^isub>1)#\<Gamma> \<turnstile> e\<^isub>1' : T" and 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   461
    h2:"(x\<^isub>2',Data \<sigma>\<^isub>2)#\<Gamma> \<turnstile> e\<^isub>2' : T" and
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   462
    e1:"[x\<^isub>1].e\<^isub>1=[x\<^isub>1'].e\<^isub>1'" and e2:"[x\<^isub>2].e\<^isub>2=[x\<^isub>2'].e\<^isub>2'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   463
    by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   464
  obtain c::name where f':"c \<sharp> (x\<^isub>1,x\<^isub>1',e\<^isub>1,e\<^isub>1',\<Gamma>)" by (erule exists_fresh[OF fs_name1])
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   465
  have e1':"[(x\<^isub>1,c)]\<bullet>e\<^isub>1 = [(x\<^isub>1',c)]\<bullet>e\<^isub>1'" using e1 f' by (auto simp add: alpha' fresh_prod fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   466
  have "[(x\<^isub>1',c)]\<bullet>((x\<^isub>1',Data \<sigma>\<^isub>1)# \<Gamma>) \<turnstile> [(x\<^isub>1',c)]\<bullet>e\<^isub>1' : T" using h1 typing_eqvt by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   467
  then have x:"(c,Data \<sigma>\<^isub>1)#( [(x\<^isub>1',c)]\<bullet>\<Gamma>) \<turnstile> [(x\<^isub>1',c)]\<bullet>e\<^isub>1': T" using f' 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   468
    by (auto simp add: fresh_atm calc_atm)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   469
  have "x\<^isub>1' \<sharp> \<Gamma>" using h1 typing_implies_valid by auto
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   470
  then have "(c,Data \<sigma>\<^isub>1)#\<Gamma> \<turnstile> [(x\<^isub>1 ,c)]\<bullet>e\<^isub>1 : T" using f' x e1' by (auto simp add: perm_fresh_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   471
  then have "[(x\<^isub>1,c)]\<bullet>((c,Data \<sigma>\<^isub>1)#\<Gamma>) \<turnstile> [(x\<^isub>1,c)]\<bullet>[(x\<^isub>1 ,c)]\<bullet>e\<^isub>1 : T" using typing_eqvt by blast 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   472
  then have "([(x\<^isub>1,c)]\<bullet>(c,Data \<sigma>\<^isub>1)) #\<Gamma> \<turnstile> [(x\<^isub>1,c)]\<bullet>[(x\<^isub>1 ,c)]\<bullet>e\<^isub>1 : T" using f f' 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   473
    by (auto simp add: perm_fresh_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   474
  then have "([(x\<^isub>1,c)]\<bullet>(c,Data \<sigma>\<^isub>1)) #\<Gamma> \<turnstile> e\<^isub>1 : T" by perm_simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   475
  then have g1:"(x\<^isub>1, Data \<sigma>\<^isub>1)#\<Gamma> \<turnstile> e\<^isub>1 : T"  using f' by (auto simp add: fresh_atm calc_atm fresh_prod)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   476
    (* The second part of the proof is the same *)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   477
  obtain c::name where f':"c \<sharp> (x\<^isub>2,x\<^isub>2',e\<^isub>2,e\<^isub>2',\<Gamma>)" by (erule exists_fresh[OF fs_name1])
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   478
  have e2':"[(x\<^isub>2,c)]\<bullet>e\<^isub>2 = [(x\<^isub>2',c)]\<bullet>e\<^isub>2'" using e2 f' by (auto simp add: alpha' fresh_prod fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   479
  have "[(x\<^isub>2',c)]\<bullet>((x\<^isub>2',Data \<sigma>\<^isub>2)# \<Gamma>) \<turnstile> [(x\<^isub>2',c)]\<bullet>e\<^isub>2' : T" using h2 typing_eqvt by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   480
  then have x:"(c,Data \<sigma>\<^isub>2)#([(x\<^isub>2',c)]\<bullet>\<Gamma>) \<turnstile> [(x\<^isub>2',c)]\<bullet>e\<^isub>2': T" using f' 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   481
    by (auto simp add: fresh_atm calc_atm)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   482
  have "x\<^isub>2' \<sharp> \<Gamma>" using h2 typing_implies_valid by auto
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   483
  then have "(c,Data \<sigma>\<^isub>2)#\<Gamma> \<turnstile> [(x\<^isub>2 ,c)]\<bullet>e\<^isub>2 : T" using f' x e2' by (auto simp add: perm_fresh_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   484
  then have "[(x\<^isub>2,c)]\<bullet>((c,Data \<sigma>\<^isub>2)#\<Gamma>) \<turnstile> [(x\<^isub>2,c)]\<bullet>[(x\<^isub>2 ,c)]\<bullet>e\<^isub>2 : T" using typing_eqvt by blast 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   485
  then have "([(x\<^isub>2,c)]\<bullet>(c,Data \<sigma>\<^isub>2))#\<Gamma> \<turnstile> [(x\<^isub>2,c)]\<bullet>[(x\<^isub>2 ,c)]\<bullet>e\<^isub>2 : T" using f f' 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   486
    by (auto simp add: perm_fresh_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   487
  then have "([(x\<^isub>2,c)]\<bullet>(c,Data \<sigma>\<^isub>2)) #\<Gamma> \<turnstile> e\<^isub>2 : T" by perm_simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   488
  then have g2:"(x\<^isub>2,Data \<sigma>\<^isub>2)#\<Gamma> \<turnstile> e\<^isub>2 : T"  using f' by (auto simp add: fresh_atm calc_atm fresh_prod)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   489
  show ?thesis using g1 g2 prems by auto 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   490
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   491
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   492
lemma weakening: 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   493
  assumes "\<Gamma>\<^isub>1 \<turnstile> e: T" and "valid \<Gamma>\<^isub>2" and "\<Gamma>\<^isub>1 \<lless> \<Gamma>\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   494
  shows "\<Gamma>\<^isub>2 \<turnstile> e: T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   495
  using assms
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   496
proof(nominal_induct \<Gamma>\<^isub>1 e T avoiding: \<Gamma>\<^isub>2 rule: typing_induct_strong)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   497
  case (t_Lam x \<Gamma>\<^isub>1 T\<^isub>1 t T\<^isub>2 \<Gamma>\<^isub>2)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   498
  have ih: "\<lbrakk>valid ((x,T\<^isub>1)#\<Gamma>\<^isub>2); (x,T\<^isub>1)#\<Gamma>\<^isub>1 \<lless> (x,T\<^isub>1)#\<Gamma>\<^isub>2\<rbrakk> \<Longrightarrow> (x,T\<^isub>1)#\<Gamma>\<^isub>2 \<turnstile> t : T\<^isub>2" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   499
  have H1: "valid \<Gamma>\<^isub>2" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   500
  have H2: "\<Gamma>\<^isub>1 \<lless> \<Gamma>\<^isub>2" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   501
  have fs: "x\<sharp>\<Gamma>\<^isub>2" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   502
  then have "valid ((x,T\<^isub>1)#\<Gamma>\<^isub>2)" using H1 by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   503
  moreover have "(x,T\<^isub>1)#\<Gamma>\<^isub>1 \<lless> (x,T\<^isub>1)#\<Gamma>\<^isub>2" using H2 by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   504
  ultimately have "(x,T\<^isub>1)#\<Gamma>\<^isub>2 \<turnstile> t : T\<^isub>2" using ih by simp 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   505
  thus "\<Gamma>\<^isub>2 \<turnstile> Lam [x].t : T\<^isub>1\<rightarrow>T\<^isub>2" using fs by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   506
next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   507
  case (t_Case x\<^isub>1 \<Gamma>\<^isub>1 e e\<^isub>2 x\<^isub>2 e\<^isub>1 S\<^isub>1 S\<^isub>2 T \<Gamma>\<^isub>2)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   508
  then have ih\<^isub>1: "valid ((x\<^isub>1,Data S\<^isub>1)#\<Gamma>\<^isub>2) \<Longrightarrow> (x\<^isub>1,Data S\<^isub>1)#\<Gamma>\<^isub>2 \<turnstile> e\<^isub>1 : T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   509
       and  ih\<^isub>2: "valid ((x\<^isub>2,Data S\<^isub>2)#\<Gamma>\<^isub>2) \<Longrightarrow> (x\<^isub>2,Data S\<^isub>2)#\<Gamma>\<^isub>2 \<turnstile> e\<^isub>2 : T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   510
       and  ih\<^isub>3: "\<Gamma>\<^isub>2 \<turnstile> e : Data (DSum S\<^isub>1 S\<^isub>2)" by auto 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   511
  have fs\<^isub>1: "x\<^isub>1\<sharp>\<Gamma>\<^isub>2" "x\<^isub>1\<sharp>e" "x\<^isub>1\<sharp>e\<^isub>2" "x\<^isub>1\<sharp>x\<^isub>2" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   512
  have fs\<^isub>2: "x\<^isub>2\<sharp>\<Gamma>\<^isub>2" "x\<^isub>2\<sharp>e" "x\<^isub>2\<sharp>e\<^isub>1" "x\<^isub>2\<sharp>x\<^isub>1" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   513
  have "valid \<Gamma>\<^isub>2" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   514
  then have "valid ((x\<^isub>1,Data S\<^isub>1)#\<Gamma>\<^isub>2)" and "valid ((x\<^isub>2,Data S\<^isub>2)#\<Gamma>\<^isub>2)" using fs\<^isub>1 fs\<^isub>2 by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   515
  then have "(x\<^isub>1, Data S\<^isub>1)#\<Gamma>\<^isub>2 \<turnstile> e\<^isub>1 : T" and "(x\<^isub>2, Data S\<^isub>2)#\<Gamma>\<^isub>2 \<turnstile> e\<^isub>2 : T" using ih\<^isub>1 ih\<^isub>2 by simp_all
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   516
  with ih\<^isub>3 show "\<Gamma>\<^isub>2 \<turnstile> Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2 : T" using fs\<^isub>1 fs\<^isub>2 by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   517
qed (auto)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   518
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   519
lemma context_exchange:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   520
  assumes a: "(x\<^isub>1,T\<^isub>1)#(x\<^isub>2,T\<^isub>2)#\<Gamma> \<turnstile> e : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   521
  shows "(x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\<Gamma> \<turnstile> e : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   522
proof -
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   523
  from  a have "valid ((x\<^isub>1,T\<^isub>1)#(x\<^isub>2,T\<^isub>2)#\<Gamma>)" by (simp add: typing_implies_valid)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   524
  then have "x\<^isub>1\<noteq>x\<^isub>2" "x\<^isub>1\<sharp>\<Gamma>" "x\<^isub>2\<sharp>\<Gamma>" "valid \<Gamma>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   525
    by (auto simp: fresh_list_cons fresh_atm[symmetric])
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   526
  then have "valid ((x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   527
    by (auto simp: fresh_list_cons fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   528
  moreover 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   529
  have "(x\<^isub>1,T\<^isub>1)#(x\<^isub>2,T\<^isub>2)#\<Gamma> \<lless> (x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\<Gamma>" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   530
  ultimately show "(x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\<Gamma> \<turnstile> e : T" using a by (auto intro: weakening)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   531
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   532
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   533
lemma typing_var_unicity: 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   534
  assumes "(x,t\<^isub>1)#\<Gamma> \<turnstile> Var x : t\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   535
  shows "t\<^isub>1=t\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   536
proof - 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   537
  have "(x,t\<^isub>2) \<in> set ((x,t\<^isub>1)#\<Gamma>)" and "valid ((x,t\<^isub>1)#\<Gamma>)" using assms by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   538
  thus "t\<^isub>1=t\<^isub>2" by (simp only: type_unicity_in_context)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   539
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   540
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   541
lemma typing_substitution: 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   542
  fixes \<Gamma>::"(name \<times> ty) list"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   543
  assumes "(x,T')#\<Gamma> \<turnstile> e : T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   544
  and     "\<Gamma> \<turnstile> e': T'" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   545
  shows "\<Gamma> \<turnstile> e[x::=e'] : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   546
  using assms
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   547
proof (nominal_induct e avoiding: \<Gamma> e' x arbitrary: T rule: trm.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   548
  case (Var y \<Gamma> e' x T)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   549
  have h1: "(x,T')#\<Gamma> \<turnstile> Var y : T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   550
  have h2: "\<Gamma> \<turnstile> e' : T'" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   551
  show "\<Gamma> \<turnstile> (Var y)[x::=e'] : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   552
  proof (cases "x=y")
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   553
    case True
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   554
    assume as: "x=y"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   555
    then have "T=T'" using h1 typing_var_unicity by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   556
    then show "\<Gamma> \<turnstile> (Var y)[x::=e'] : T" using as h2 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   557
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   558
    case False
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   559
    assume as: "x\<noteq>y" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   560
    have "(y,T) \<in> set ((x,T')#\<Gamma>)" using h1 by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   561
    then have "(y,T) \<in> set \<Gamma>" using as by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   562
    moreover 
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   563
    have "valid \<Gamma>" using h2 by (simp only: typing_implies_valid)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   564
    ultimately show "\<Gamma> \<turnstile> (Var y)[x::=e'] : T" using as by (simp add: t_Var)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   565
  qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   566
next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   567
  case (Lam y t \<Gamma> e' x T)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   568
  have vc: "y\<sharp>\<Gamma>" "y\<sharp>x" "y\<sharp>e'"  by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   569
  have pr1: "\<Gamma> \<turnstile> e' : T'" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   570
  have pr2: "(x,T')#\<Gamma> \<turnstile> Lam [y].t : T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   571
  then obtain T\<^isub>1 T\<^isub>2 where pr2': "(y,T\<^isub>1)#(x,T')#\<Gamma> \<turnstile> t : T\<^isub>2" and eq: "T = T\<^isub>1\<rightarrow>T\<^isub>2" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   572
    using vc by (auto simp add: fresh_list_cons)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   573
  then have pr2'':"(x,T')#(y,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2" by (simp add: context_exchange)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   574
  have ih: "\<lbrakk>(x,T')#(y,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2; (y,T\<^isub>1)#\<Gamma> \<turnstile> e' : T'\<rbrakk> \<Longrightarrow> (y,T\<^isub>1)#\<Gamma> \<turnstile> t[x::=e'] : T\<^isub>2" by fact
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   575
  have "valid \<Gamma>" using pr1 by (simp add: typing_implies_valid)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   576
  then have "valid ((y,T\<^isub>1)#\<Gamma>)" using vc by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   577
  then have "(y,T\<^isub>1)#\<Gamma> \<turnstile> e' : T'" using pr1 by (auto intro: weakening)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   578
  then have "(y,T\<^isub>1)#\<Gamma> \<turnstile> t[x::=e'] : T\<^isub>2" using ih pr2'' by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   579
  then have "\<Gamma> \<turnstile> Lam [y].(t[x::=e']) : T\<^isub>1\<rightarrow>T\<^isub>2" using vc by (auto intro: t_Lam)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   580
  thus "\<Gamma> \<turnstile> (Lam [y].t)[x::=e'] : T" using vc eq by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   581
next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   582
  case (Case t\<^isub>1 x\<^isub>1 t\<^isub>2 x\<^isub>2 t3 \<Gamma> e' x T)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   583
  have vc: "x\<^isub>1\<sharp>\<Gamma>" "x\<^isub>1\<sharp>e'" "x\<^isub>1\<sharp>x""x\<^isub>1\<sharp>t\<^isub>1" "x\<^isub>1\<sharp>t3" "x\<^isub>2\<sharp>\<Gamma>" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   584
           "x\<^isub>2\<sharp>e'" "x\<^isub>2\<sharp>x"  "x\<^isub>2\<sharp>t\<^isub>1" "x\<^isub>2\<sharp>t\<^isub>2" "x\<^isub>2\<noteq>x\<^isub>1" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   585
  have as1: "\<Gamma> \<turnstile> e' : T'" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   586
  have as2: "(x,T')#\<Gamma> \<turnstile> Case t\<^isub>1 of inl x\<^isub>1 \<rightarrow> t\<^isub>2 | inr x\<^isub>2 \<rightarrow> t3 : T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   587
  then obtain S\<^isub>1 S\<^isub>2 where 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   588
    h1:"(x,T')#\<Gamma> \<turnstile> t\<^isub>1 : Data (DSum S\<^isub>1 S\<^isub>2)" and
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   589
    h2:"(x\<^isub>1,Data S\<^isub>1)#(x,T')#\<Gamma> \<turnstile> t\<^isub>2 : T" and
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   590
    h3:"(x\<^isub>2,Data S\<^isub>2)#(x,T')#\<Gamma> \<turnstile> t3 : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   591
    using vc by (auto simp add: fresh_list_cons)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   592
  have ih1: "\<lbrakk>(x,T')#\<Gamma> \<turnstile> t\<^isub>1 : T; \<Gamma> \<turnstile> e' : T'\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1[x::=e'] : T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   593
  and  ih2: "\<lbrakk>(x,T')#(x\<^isub>1,Data S\<^isub>1)#\<Gamma> \<turnstile> t\<^isub>2:T; (x\<^isub>1,Data S\<^isub>1)#\<Gamma> \<turnstile> e':T'\<rbrakk> \<Longrightarrow> (x\<^isub>1,Data S\<^isub>1)#\<Gamma> \<turnstile> t\<^isub>2[x::=e']:T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   594
  and  ih3: "\<lbrakk>(x,T')#(x\<^isub>2,Data S\<^isub>2)#\<Gamma> \<turnstile> t3:T; (x\<^isub>2,Data S\<^isub>2)#\<Gamma> \<turnstile> e':T'\<rbrakk> \<Longrightarrow> (x\<^isub>2,Data S\<^isub>2)#\<Gamma> \<turnstile> t3[x::=e']:T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   595
    by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   596
  from h2 have h2': "(x,T')#(x\<^isub>1,Data S\<^isub>1)#\<Gamma> \<turnstile> t\<^isub>2 : T" by (rule context_exchange)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   597
  from h3 have h3': "(x,T')#(x\<^isub>2,Data S\<^isub>2)#\<Gamma> \<turnstile> t3 : T" by (rule context_exchange)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   598
  have "\<Gamma> \<turnstile> t\<^isub>1[x::=e'] : Data (DSum S\<^isub>1 S\<^isub>2)" using h1 ih1 as1 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   599
  moreover
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   600
  have "valid ((x\<^isub>1,Data S\<^isub>1)#\<Gamma>)" using h2' by (auto dest: typing_implies_valid)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   601
  then have "(x\<^isub>1,Data S\<^isub>1)#\<Gamma> \<turnstile> e' : T'" using as1 by (auto simp add: weakening)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   602
  then have "(x\<^isub>1,Data S\<^isub>1)#\<Gamma> \<turnstile> t\<^isub>2[x::=e'] : T" using ih2 h2' by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   603
  moreover 
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   604
  have "valid ((x\<^isub>2,Data S\<^isub>2)#\<Gamma>)" using h3' by (auto dest: typing_implies_valid)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   605
  then have "(x\<^isub>2,Data S\<^isub>2)#\<Gamma> \<turnstile> e' : T'" using as1 by (auto simp add: weakening)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   606
  then have "(x\<^isub>2,Data S\<^isub>2)#\<Gamma> \<turnstile> t3[x::=e'] : T" using ih3 h3' by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   607
  ultimately have "\<Gamma> \<turnstile> Case (t\<^isub>1[x::=e']) of inl x\<^isub>1 \<rightarrow> (t\<^isub>2[x::=e']) | inr x\<^isub>2 \<rightarrow> (t3[x::=e']) : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   608
    using vc by (auto simp add: fresh_atm fresh_subst)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   609
  thus "\<Gamma> \<turnstile> (Case t\<^isub>1 of inl x\<^isub>1 \<rightarrow> t\<^isub>2 | inr x\<^isub>2 \<rightarrow> t3)[x::=e'] : T" using vc by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   610
qed (simp, fast)+
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   611
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   612
text {* Big-Step Evaluation *}
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   613
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   614
inductive2
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   615
  big :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   616
where
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   617
  b_Lam[intro]:   "Lam [x].e \<Down> Lam [x].e"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   618
| b_App[intro]:   "\<lbrakk>x\<sharp>(e\<^isub>1,e\<^isub>2,e'); e\<^isub>1\<Down>Lam [x].e; e\<^isub>2\<Down>e\<^isub>2'; e[x::=e\<^isub>2']\<Down>e'\<rbrakk> \<Longrightarrow> App e\<^isub>1 e\<^isub>2 \<Down> e'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   619
| b_Const[intro]: "Const n \<Down> Const n"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   620
| b_Pr[intro]:    "\<lbrakk>e\<^isub>1\<Down>e\<^isub>1'; e\<^isub>2\<Down>e\<^isub>2'\<rbrakk> \<Longrightarrow> Pr e\<^isub>1 e\<^isub>2 \<Down> Pr e\<^isub>1' e\<^isub>2'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   621
| b_Fst[intro]:   "e\<Down>Pr e\<^isub>1 e\<^isub>2 \<Longrightarrow> Fst e\<Down>e\<^isub>1"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   622
| b_Snd[intro]:   "e\<Down>Pr e\<^isub>1 e\<^isub>2 \<Longrightarrow> Snd e\<Down>e\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   623
| b_InL[intro]:   "e\<Down>e' \<Longrightarrow> InL e \<Down> InL e'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   624
| b_InR[intro]:   "e\<Down>e' \<Longrightarrow> InR e \<Down> InR e'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   625
| b_CaseL[intro]: "\<lbrakk>x\<^isub>1\<sharp>(e,e\<^isub>2,e'',x\<^isub>2); x\<^isub>2\<sharp>(e,e\<^isub>1,e'',x\<^isub>1) ; e\<Down>InL e'; e\<^isub>1[x\<^isub>1::=e']\<Down>e''\<rbrakk> 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   626
                   \<Longrightarrow> Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2 \<Down> e''"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   627
| b_CaseR[intro]: "\<lbrakk>x\<^isub>1\<sharp>(e,e\<^isub>2,e'',x\<^isub>2); x\<^isub>2\<sharp>(e,e\<^isub>1,e'',x\<^isub>1) ; e\<Down>InR e'; e\<^isub>2[x\<^isub>2::=e']\<Down>e''\<rbrakk> 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   628
                   \<Longrightarrow> Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2 \<Down> e''"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   629
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   630
nominal_inductive big
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   631
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   632
lemma big_eqvt[eqvt]:
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   633
  fixes pi::"name prm"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   634
  assumes a: "t \<Down> t'"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   635
  shows "(pi\<bullet>t) \<Down> (pi\<bullet>t')"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   636
  using a
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   637
  apply(induct)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   638
  apply(auto simp add: eqvt)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   639
  apply(rule_tac x="pi\<bullet>x" in b_App)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   640
  apply(auto simp add: eqvt fresh_bij fresh_prod)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   641
  done
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   642
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   643
lemma big_eqvt':
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   644
  fixes pi::"name prm"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   645
  assumes a: "(pi\<bullet>t) \<Down> (pi\<bullet>t')"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   646
  shows "t \<Down> t'"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   647
using a
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   648
apply -
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   649
apply(drule_tac pi="rev pi" in big_eqvt)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   650
apply(perm_simp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   651
done
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   652
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   653
lemma fresh_preserved:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   654
  fixes x::name
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   655
  fixes t::trm
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   656
  fixes t'::trm
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   657
  assumes "e \<Down> e'" and "x\<sharp>e" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   658
  shows "x\<sharp>e'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   659
  using assms by (induct) (auto simp add:fresh_subst')
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   660
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   661
declare trm.inject  [simp add]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   662
declare ty.inject  [simp add]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   663
declare data.inject [simp add]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   664
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   665
inductive_cases2 b_App_inv_auto[elim]: "App e\<^isub>1 e\<^isub>2 \<Down> t" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   666
inductive_cases2 b_Case_inv_auto[elim]: "Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2 \<Down> t"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   667
inductive_cases2 b_Lam_inv_auto[elim]: "Lam[x].t \<Down> t"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   668
inductive_cases2 b_Const_inv_auto[elim]: "Const n \<Down> t"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   669
inductive_cases2 b_Fst_inv_auto[elim]: "Fst e \<Down> t"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   670
inductive_cases2 b_Snd_inv_auto[elim]: "Snd e \<Down> t"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   671
inductive_cases2 b_InL_inv_auto[elim]: "InL e \<Down> t"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   672
inductive_cases2 b_InR_inv_auto[elim]: "InR e \<Down> t"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   673
inductive_cases2 b_Pr_inv_auto[elim]: "Pr e\<^isub>1 e\<^isub>2 \<Down> t"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   674
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   675
declare trm.inject  [simp del]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   676
declare ty.inject  [simp del]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   677
declare data.inject [simp del]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   678
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   679
lemma big_induct_strong
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   680
  [consumes 1, case_names b_Lam b_App b_Const b_Pr b_Fst b_Snd b_InL b_InR b_CaseL b_CaseR]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   681
    fixes  P::"'a::fs_name \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow>bool"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   682
    and    x :: "'a::fs_name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   683
    assumes a: "t \<Down> t'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   684
    and a1: "\<And>x e c. P c (Lam [x].e) (Lam [x].e)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   685
    and a2: "\<And>x e\<^isub>1 e\<^isub>2 e\<^isub>2' e' e\<^isub>1' c. 
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   686
             \<lbrakk>x\<sharp>(e\<^isub>1,e\<^isub>2,e',c); \<And>c. P c e\<^isub>1 (Lam [x].e\<^isub>1'); \<And>c. P c e\<^isub>2 e\<^isub>2'; \<And>c. P c (e\<^isub>1'[x::=e\<^isub>2']) e'\<rbrakk> 
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   687
             \<Longrightarrow> P c (App e\<^isub>1 e\<^isub>2) e'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   688
    and a3: "\<And>n c. P c (Const n) (Const n)"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   689
    and a4: "\<And>e\<^isub>1 e\<^isub>1' e\<^isub>2 e\<^isub>2' c. \<lbrakk>\<And>c. P c e\<^isub>1 e\<^isub>1'; \<And>c. P c e\<^isub>2 e\<^isub>2'\<rbrakk>
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   690
             \<Longrightarrow> P c (Pr e\<^isub>1 e\<^isub>2) (Pr e\<^isub>1' e\<^isub>2')"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   691
    and a5: "\<And>e e\<^isub>1 e\<^isub>2 c. \<lbrakk>\<And>c. P c e (Pr e\<^isub>1 e\<^isub>2)\<rbrakk> \<Longrightarrow> P c (Fst e) e\<^isub>1"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   692
    and a6: "\<And>e e\<^isub>1 e\<^isub>2 c. \<lbrakk>\<And>c. P c e (Pr e\<^isub>1 e\<^isub>2)\<rbrakk> \<Longrightarrow> P c (Snd e) e\<^isub>2"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   693
    and a7: "\<And>e e' c. \<lbrakk>\<And>c. P c e e'\<rbrakk> \<Longrightarrow> P c (InL e) (InL e')"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   694
    and a8: "\<And>e e' c. \<lbrakk>\<And>c. P c e e'\<rbrakk> \<Longrightarrow> P c (InR e) (InR e')"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   695
    and a9: "\<And>x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' c.
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   696
             \<lbrakk>x\<^isub>1\<sharp>(e,e\<^isub>2,e'',x\<^isub>2,c); x\<^isub>2\<sharp>(e,e\<^isub>1,e'',x\<^isub>1,c);  \<And>c. P c e (InL e'); \<And>c. P c (e\<^isub>1[x\<^isub>1::=e']) e''\<rbrakk>
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   697
             \<Longrightarrow> P c (Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2) e''"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   698
    and a10:"\<And>x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' c.
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   699
             \<lbrakk>x\<^isub>1\<sharp>(e,e\<^isub>2,e'',x\<^isub>2,c); x\<^isub>2\<sharp>(e,e\<^isub>1,e'',x\<^isub>1,c); \<And>c. P c e (InR e'); \<And>c. P c (e\<^isub>2[x\<^isub>2::=e']) e''\<rbrakk>
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   700
             \<Longrightarrow> P c (Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2) e''"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   701
    shows "P c t t'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   702
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   703
  from a have "\<And>(pi::name prm) c. P c (pi\<bullet>t) (pi\<bullet>t')"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   704
  proof (induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   705
    case (b_Lam x e pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   706
    show "P c (pi\<bullet>(Lam [x].e)) (pi\<bullet>(Lam [x].e))" using a1 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   707
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   708
    case (b_App x e\<^isub>1 e\<^isub>2 e' e\<^isub>1' e\<^isub>2' pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   709
    obtain y::"name" where fs: "y\<sharp>(pi\<bullet>x,pi\<bullet>e\<^isub>1,pi\<bullet>e\<^isub>2,pi\<bullet>e\<^isub>2',pi\<bullet>e',pi\<bullet>e\<^isub>1,c)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   710
      by (erule exists_fresh[OF fs_name1])
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   711
    let ?sw = "[(pi\<bullet>x,y)]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   712
    let ?pi' = "?sw@pi"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   713
    have f0: "x\<sharp>(e\<^isub>1,e\<^isub>2,e')" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   714
    have f1: "(pi\<bullet>x)\<sharp>(pi\<bullet>e\<^isub>1,pi\<bullet>e\<^isub>2,pi\<bullet>e')" using f0 by (simp add: fresh_bij)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   715
    have f2: "y\<sharp>(?pi'\<bullet>e\<^isub>1,?pi'\<bullet>e\<^isub>2,?pi'\<bullet>e')" using f0 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   716
      by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp fresh_prod)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   717
    have ih1: "\<And>c. P c (?pi'\<bullet>e\<^isub>1) (?pi'\<bullet>(Lam [x].e\<^isub>1'))" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   718
    then have "\<And>c. P c (?pi'\<bullet>e\<^isub>1) (Lam [y].(?pi'\<bullet>e\<^isub>1'))" by (simp add: calc_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   719
    moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   720
    have ih2: "\<And>c. P c (?pi'\<bullet>e\<^isub>2) (?pi'\<bullet>e\<^isub>2')" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   721
    moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   722
    have ih3: "\<And>c. P c (?pi'\<bullet>(e\<^isub>1'[x::=e\<^isub>2'])) (?pi'\<bullet>e')" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   723
    then have "\<And>c. P c ((?pi'\<bullet>e\<^isub>1')[y::=(?pi'\<bullet>e\<^isub>2')]) (?pi'\<bullet>e')" by (simp add: calc_atm subst_eqvt)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   724
    ultimately have "P c (App (?pi'\<bullet>e\<^isub>1) (?pi'\<bullet>e\<^isub>2)) (?pi'\<bullet>e')" using fs f2
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   725
      by (auto intro!: a2 simp add: calc_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   726
    then have "P c (?sw\<bullet>(App (pi\<bullet>e\<^isub>1) (pi\<bullet>e\<^isub>2))) (?sw\<bullet>(pi\<bullet>e'))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   727
      by (simp del: append_Cons add: pt_name2)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   728
    moreover have "(?sw\<bullet>(App (pi\<bullet>e\<^isub>1) (pi\<bullet>e\<^isub>2))) = App (pi\<bullet>e\<^isub>1) (pi\<bullet>e\<^isub>2)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   729
      by (rule perm_fresh_fresh) (simp_all add: fs f1)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   730
    moreover have "(?sw\<bullet>(pi\<bullet>e')) = pi\<bullet>e'" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   731
      by (rule perm_fresh_fresh) (simp_all add: fs f1)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   732
    ultimately show "P c (pi\<bullet>(App e\<^isub>1 e\<^isub>2)) (pi\<bullet>e')" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   733
      by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   734
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   735
    case (b_Const n pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   736
    show "P c (pi\<bullet>(Const n)) (pi\<bullet>(Const n))" using a3 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   737
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   738
    case (b_Pr e\<^isub>1 e\<^isub>1' e\<^isub>2 e\<^isub>2' pi c)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   739
    then show "P c (pi\<bullet>(Pr e\<^isub>1 e\<^isub>2)) (pi\<bullet>(Pr e\<^isub>1' e\<^isub>2'))" using a4 by simp
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   740
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   741
    case (b_Fst e e\<^isub>1 e\<^isub>2 pi c)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   742
    then show "P c (pi\<bullet>(Fst e)) (pi\<bullet>e\<^isub>1)" using a5 by (simp, blast)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   743
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   744
    case (b_Snd e e\<^isub>1 e\<^isub>2 pi c)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   745
    then show "P c (pi\<bullet>(Snd e)) (pi\<bullet>e\<^isub>2)" using a6 by (simp, blast)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   746
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   747
    case (b_InL e e' pi c)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   748
    then show "P c (pi\<bullet>(InL e)) (pi\<bullet>(InL e'))" using a7 by (simp)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   749
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   750
    case (b_InR e e' pi c)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   751
    then show "P c (pi\<bullet>(InR e)) (pi\<bullet>(InR e'))" using a8 by (simp)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   752
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   753
    case (b_CaseL x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   754
    obtain y\<^isub>1::"name" where fs1: "y\<^isub>1\<sharp>(pi\<bullet>x\<^isub>1,pi\<bullet>e,pi\<bullet>e\<^isub>1,pi\<bullet>e\<^isub>2,pi\<bullet>e'',pi\<bullet>x\<^isub>2,c)"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   755
      by (rule exists_fresh[OF fin_supp])
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   756
    obtain y\<^isub>2::"name" where fs2: "y\<^isub>2\<sharp>(pi\<bullet>x\<^isub>2,pi\<bullet>e,pi\<bullet>e\<^isub>1,pi\<bullet>e\<^isub>2,pi\<bullet>e'',pi\<bullet>x\<^isub>1,c,y\<^isub>1)"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   757
      by (rule exists_fresh[OF fin_supp])
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   758
    let ?sw1 = "[(pi\<bullet>x\<^isub>1,y\<^isub>1)]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   759
    let ?sw2 = "[(pi\<bullet>x\<^isub>2,y\<^isub>2)]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   760
    let ?pi' = "?sw2@?sw1@pi"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   761
    have f01: "x\<^isub>1\<sharp>(e,e\<^isub>2,e'',x\<^isub>2)" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   762
    have f11: "(pi\<bullet>x\<^isub>1)\<sharp>(pi\<bullet>e,pi\<bullet>e\<^isub>2,pi\<bullet>e'',pi\<bullet>x\<^isub>2)" using f01 by (simp add: fresh_bij)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   763
    have f21: "y\<^isub>1\<sharp>(?pi'\<bullet>e,?pi'\<bullet>e\<^isub>2,?pi'\<bullet>e'')" using f01 fs1 fs2 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   764
      by (simp add: fresh_atm fresh_prod fresh_left calc_atm pt_name2 perm_pi_simp)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   765
    have f02: "x\<^isub>2\<sharp>(e,e\<^isub>1,e'',x\<^isub>1)" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   766
    have f12: "(pi\<bullet>x\<^isub>2)\<sharp>(pi\<bullet>e,pi\<bullet>e\<^isub>1,pi\<bullet>e'',pi\<bullet>x\<^isub>1)" using f02 by (simp add: fresh_bij)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   767
    have f22: "y\<^isub>2\<sharp>(?pi'\<bullet>e,?pi'\<bullet>e\<^isub>1,?pi'\<bullet>e'')" using f02 fs1 fs2 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   768
      by (auto simp add: fresh_atm fresh_prod fresh_left calc_atm pt_name2 perm_pi_simp)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   769
    have ih1: "\<And>c. P c (?pi'\<bullet>e) (?pi'\<bullet>(InL e'))" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   770
    moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   771
    have ih2: "\<And>c. P c (?pi'\<bullet>(e\<^isub>1[x\<^isub>1::=e'])) (?pi'\<bullet>e'')" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   772
    then have "\<And>c. P c ((?pi'\<bullet>e\<^isub>1)[y\<^isub>1::=(?pi'\<bullet>e')]) (?pi'\<bullet>e'')" using fs1 fs2 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   773
      by (auto simp add: calc_atm subst_eqvt fresh_prod fresh_atm del: append_Cons)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   774
    ultimately have "P c (Case (?pi'\<bullet>e) of inl y\<^isub>1 \<rightarrow> (?pi'\<bullet>e\<^isub>1) | inr y\<^isub>2 \<rightarrow> (?pi'\<bullet>e\<^isub>2)) (?pi'\<bullet>e'')"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   775
      using f21 f22 fs1 fs2  by (auto intro!: a9 simp add: fresh_atm fresh_prod)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   776
    then have "P c (?sw2\<bullet>?sw1\<bullet>(Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   777
                   (?sw2\<bullet>?sw1\<bullet>(pi\<bullet>e''))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   778
      using fs1 fs2 f01 f02 f11 f12 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   779
      by (auto simp del: append_Cons simp add: pt_name2 fresh_atm fresh_prod calc_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   780
    moreover have "?sw1\<bullet>(Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   781
                    = (Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   782
      by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12 abs_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   783
    moreover have "?sw2\<bullet>(Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   784
                    = (Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   785
      by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12 abs_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   786
    moreover have "(?sw1\<bullet>(pi\<bullet>e'')) = (pi\<bullet>e'')"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   787
      by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   788
    moreover have "(?sw2\<bullet>(pi\<bullet>e'')) = (pi\<bullet>e'')"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   789
      by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   790
    ultimately show "P c (pi\<bullet>(Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2)) (pi\<bullet>e'')" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   791
      by (simp only:, simp)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   792
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   793
    case (b_CaseR x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' pi c)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   794
    obtain y\<^isub>1::"name" where fs1: "y\<^isub>1\<sharp>(pi\<bullet>x\<^isub>1,pi\<bullet>e,pi\<bullet>e\<^isub>1,pi\<bullet>e\<^isub>2,pi\<bullet>e'',pi\<bullet>x\<^isub>2,c)"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   795
      by (rule exists_fresh[OF fin_supp])
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   796
    obtain y\<^isub>2::"name" where fs2: "y\<^isub>2\<sharp>(pi\<bullet>x\<^isub>2,pi\<bullet>e,pi\<bullet>e\<^isub>1,pi\<bullet>e\<^isub>2,pi\<bullet>e'',pi\<bullet>x\<^isub>1,c,y\<^isub>1)"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   797
      by (rule exists_fresh[OF fin_supp])
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   798
    let ?sw1 = "[(pi\<bullet>x\<^isub>1,y\<^isub>1)]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   799
    let ?sw2 = "[(pi\<bullet>x\<^isub>2,y\<^isub>2)]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   800
    let ?pi' = "?sw2@?sw1@pi"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   801
    have f01: "x\<^isub>1\<sharp>(e,e\<^isub>2,e'',x\<^isub>2)" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   802
    have f11: "(pi\<bullet>x\<^isub>1)\<sharp>(pi\<bullet>e,pi\<bullet>e\<^isub>2,pi\<bullet>e'',pi\<bullet>x\<^isub>2)" using f01 by (simp add: fresh_bij)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   803
    have f21: "y\<^isub>1\<sharp>(?pi'\<bullet>e,?pi'\<bullet>e\<^isub>2,?pi'\<bullet>e'')" using f01 fs1 fs2 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   804
      by (simp add: fresh_atm fresh_prod fresh_left calc_atm pt_name2 perm_pi_simp)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   805
    have f02: "x\<^isub>2\<sharp>(e,e\<^isub>1,e'',x\<^isub>1)" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   806
    have f12: "(pi\<bullet>x\<^isub>2)\<sharp>(pi\<bullet>e,pi\<bullet>e\<^isub>1,pi\<bullet>e'',pi\<bullet>x\<^isub>1)" using f02 by (simp add: fresh_bij)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   807
    have f22: "y\<^isub>2\<sharp>(?pi'\<bullet>e,?pi'\<bullet>e\<^isub>1,?pi'\<bullet>e'')" using f02 fs1 fs2 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   808
      by (auto simp add: fresh_atm fresh_prod fresh_left calc_atm pt_name2 perm_pi_simp)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   809
    have ih1: "\<And>c. P c (?pi'\<bullet>e) (?pi'\<bullet>(InR e'))" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   810
    moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   811
    have ih2: "\<And>c. P c (?pi'\<bullet>(e\<^isub>2[x\<^isub>2::=e'])) (?pi'\<bullet>e'')" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   812
    then have "\<And>c. P c ((?pi'\<bullet>e\<^isub>2)[y\<^isub>2::=(?pi'\<bullet>e')]) (?pi'\<bullet>e'')" using fs1 fs2 f11 f12 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   813
      by (auto simp add: calc_atm subst_eqvt fresh_prod fresh_atm del: append_Cons)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   814
    ultimately have "P c (Case (?pi'\<bullet>e) of inl y\<^isub>1 \<rightarrow> (?pi'\<bullet>e\<^isub>1) | inr y\<^isub>2 \<rightarrow> (?pi'\<bullet>e\<^isub>2)) (?pi'\<bullet>e'')"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   815
      using f21 f22 fs1 fs2  by (auto intro!: a10 simp add: fresh_atm fresh_prod)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   816
    then have "P c (?sw2\<bullet>?sw1\<bullet>(Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   817
                   (?sw2\<bullet>?sw1\<bullet>(pi\<bullet>e''))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   818
      using fs1 fs2 f01 f02 f11 f12 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   819
      by (auto simp del: append_Cons simp add: pt_name2 fresh_atm fresh_prod calc_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   820
    moreover have "?sw1\<bullet>(Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   821
                    = (Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   822
      by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12 abs_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   823
    moreover have "?sw2\<bullet>(Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   824
                    = (Case (pi\<bullet>e) of inl (pi\<bullet>x\<^isub>1) \<rightarrow> (pi\<bullet>e\<^isub>1) | inr (pi\<bullet>x\<^isub>2) \<rightarrow> (pi\<bullet>e\<^isub>2))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   825
      by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12 abs_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   826
    moreover have "(?sw1\<bullet>(pi\<bullet>e'')) = (pi\<bullet>e'')"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   827
      by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   828
    moreover have "(?sw2\<bullet>(pi\<bullet>e'')) = (pi\<bullet>e'')"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   829
      by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   830
    ultimately show "P c (pi\<bullet>(Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2)) (pi\<bullet>e'')" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   831
      by (simp only:, simp)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   832
  qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   833
  then have "P c (([]::name prm)\<bullet>t) (([]::name prm)\<bullet>t')" by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   834
  then show "P c t t'" by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   835
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   836
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   837
lemma b_App_elim[elim]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   838
  assumes "App e\<^isub>1 e\<^isub>2 \<Down> e'" and "x\<sharp>(e\<^isub>1,e\<^isub>2,e')"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   839
  obtains f\<^isub>1 and f\<^isub>2 where "e\<^isub>1 \<Down> Lam [x]. f\<^isub>1" "e\<^isub>2 \<Down> f\<^isub>2" "f\<^isub>1[x::=f\<^isub>2] \<Down> e'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   840
  using assms
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   841
  apply -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   842
  apply(erule b_App_inv_auto)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   843
  apply(drule_tac pi="[(xa,x)]" in big_eqvt)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   844
  apply(drule_tac pi="[(xa,x)]" in big_eqvt)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   845
  apply(drule_tac pi="[(xa,x)]" in big_eqvt)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   846
  apply(perm_simp add: calc_atm eqvt)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   847
  done
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   848
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   849
lemma  b_CaseL_elim[elim]: 
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   850
  assumes "Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2 \<Down> e''" 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   851
  and     "\<And> t. \<not>  e \<Down> InR t"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   852
  and     "x\<^isub>1\<sharp>e''" "x\<^isub>1\<sharp>e" "x\<^isub>2\<sharp>e''" "x\<^isub>1\<sharp>e"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   853
  obtains e' where "e \<Down> InL e'" and "e\<^isub>1[x\<^isub>1::=e'] \<Down> e''"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   854
  using assms 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   855
  apply -
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   856
  apply(rule b_Case_inv_auto)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   857
  apply(auto)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   858
  apply(simp add: alpha)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   859
  apply(auto)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   860
  apply(drule_tac x="[(x\<^isub>1,x\<^isub>1')]\<bullet>e'" in meta_spec)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   861
  apply(drule meta_mp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   862
  apply(rule_tac pi="[(x\<^isub>1,x\<^isub>1')]" in big_eqvt')
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   863
  apply(perm_simp add: fresh_prod)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   864
  apply(drule meta_mp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   865
  apply(rule_tac pi="[(x\<^isub>1,x\<^isub>1')]" in big_eqvt')
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   866
  apply(perm_simp add: eqvt calc_atm)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   867
  apply(assumption)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   868
  apply(drule_tac x="[(x\<^isub>1,x\<^isub>1')]\<bullet>e'" in meta_spec)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   869
  apply(drule meta_mp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   870
  apply(rule_tac pi="[(x\<^isub>1,x\<^isub>1')]" in big_eqvt')
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   871
  apply(perm_simp add: fresh_prod)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   872
  apply(drule meta_mp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   873
  apply(rule_tac pi="[(x\<^isub>1,x\<^isub>1')]" in big_eqvt')
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   874
  apply(perm_simp add: eqvt calc_atm)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   875
  apply(assumption)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   876
done
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   877
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   878
lemma b_CaseR_elim[elim]: 
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   879
  assumes "Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2 \<Down> e''" 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   880
  and     "\<And> t. \<not> e \<Down> InL t"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   881
  and     "x\<^isub>1\<sharp>e''" "x\<^isub>1\<sharp>e" "x\<^isub>2\<sharp>e''" "x\<^isub>2\<sharp>e"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   882
  obtains e' where "e \<Down> InR e'" and "e\<^isub>2[x\<^isub>2::=e'] \<Down> e''"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   883
  using assms 
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   884
  apply -
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   885
  apply(rule b_Case_inv_auto)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   886
  apply(auto)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   887
  apply(simp add: alpha)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   888
  apply(auto)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   889
  apply(drule_tac x="[(x\<^isub>2,x\<^isub>2')]\<bullet>e'" in meta_spec)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   890
  apply(drule meta_mp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   891
  apply(rule_tac pi="[(x\<^isub>2,x\<^isub>2')]" in big_eqvt')
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   892
  apply(perm_simp add: fresh_prod)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   893
  apply(drule meta_mp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   894
  apply(rule_tac pi="[(x\<^isub>2,x\<^isub>2')]" in big_eqvt')
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   895
  apply(perm_simp add: eqvt calc_atm)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   896
  apply(assumption)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   897
  apply(drule_tac x="[(x\<^isub>2,x\<^isub>2')]\<bullet>e'" in meta_spec)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   898
  apply(drule meta_mp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   899
  apply(rule_tac pi="[(x\<^isub>2,x\<^isub>2')]" in big_eqvt')
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   900
  apply(perm_simp add: fresh_prod)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   901
  apply(drule meta_mp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   902
  apply(rule_tac pi="[(x\<^isub>2,x\<^isub>2')]" in big_eqvt')
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   903
  apply(perm_simp add: eqvt calc_atm)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   904
  apply(assumption)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   905
done
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   906
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   907
inductive2
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   908
  val :: "trm\<Rightarrow>bool" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   909
where
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   910
  v_Lam[intro]:   "val (Lam [x].e)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   911
| v_Const[intro]: "val (Const n)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   912
| v_Pr[intro]:    "\<lbrakk>val e\<^isub>1; val e\<^isub>2\<rbrakk> \<Longrightarrow> val (Pr e\<^isub>1 e\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   913
| v_InL[intro]:   "val e \<Longrightarrow> val (InL e)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   914
| v_InR[intro]:   "val e \<Longrightarrow> val (InR e)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   915
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   916
declare trm.inject  [simp add]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   917
declare ty.inject  [simp add]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   918
declare data.inject [simp add]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   919
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   920
inductive_cases2 v_Const_inv_auto[elim]: "val (Const n)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   921
inductive_cases2 v_Pr_inv_auto[elim]: "val (Pr e\<^isub>1 e\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   922
inductive_cases2 v_InL_inv_auto[elim]: "val (InL e)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   923
inductive_cases2 v_InR_inv_auto[elim]: "val (InR e)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   924
inductive_cases2 v_Fst_inv_auto[elim]: "val (Fst e)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   925
inductive_cases2 v_Snd_inv_auto[elim]: "val (Snd e)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   926
inductive_cases2 v_Case_inv_auto[elim]: "val (Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   927
inductive_cases2 v_Var_inv_auto[elim]: "val (Var x)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   928
inductive_cases2 v_Lam_inv_auto[elim]: "val (Lam [x].e)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   929
inductive_cases2 v_App_inv_auto[elim]: "val (App e\<^isub>1 e\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   930
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   931
declare trm.inject  [simp del]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   932
declare ty.inject  [simp del]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   933
declare data.inject [simp del]
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   934
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   935
lemma subject_reduction:
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   936
  assumes a: "e \<Down> e'" 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   937
  and     b: "\<Gamma> \<turnstile> e : T"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   938
  shows "\<Gamma> \<turnstile> e' : T"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   939
  using a b
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   940
proof (nominal_induct avoiding: \<Gamma> arbitrary: T rule: big_induct_strong) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   941
  case (b_App x e\<^isub>1 e\<^isub>2 e\<^isub>2' e' e \<Gamma> T)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   942
  have vc: "x\<sharp>\<Gamma>" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   943
  have "\<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T" by fact
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   944
  then obtain T' where 
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   945
    a1: "\<Gamma> \<turnstile> e\<^isub>1 : T'\<rightarrow>T" and  
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   946
    a2: "\<Gamma> \<turnstile> e\<^isub>2 : T'" by auto
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   947
  have ih1: "\<Gamma> \<turnstile> e\<^isub>1 : T' \<rightarrow> T \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].e : T' \<rightarrow> T" by fact
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   948
  have ih2: "\<Gamma> \<turnstile> e\<^isub>2 : T' \<Longrightarrow> \<Gamma> \<turnstile> e\<^isub>2' : T'" by fact 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   949
  have ih3: "\<Gamma> \<turnstile> e[x::=e\<^isub>2'] : T \<Longrightarrow> \<Gamma> \<turnstile> e' : T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   950
  have "\<Gamma> \<turnstile> Lam [x].e : T'\<rightarrow>T" using ih1 a1 by simp 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   951
  then have "((x,T')#\<Gamma>) \<turnstile> e : T" using vc by (auto simp add: ty.inject)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   952
  moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   953
  have "\<Gamma> \<turnstile> e\<^isub>2': T'" using ih2 a2 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   954
  ultimately have "\<Gamma> \<turnstile> e[x::=e\<^isub>2'] : T" by (simp add: typing_substitution)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   955
  thus "\<Gamma> \<turnstile> e' : T" using ih3 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   956
next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   957
  case (b_CaseL x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' \<Gamma>)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   958
  have vc: "x\<^isub>1\<sharp>\<Gamma>" "x\<^isub>2\<sharp>\<Gamma>" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   959
  have "\<Gamma> \<turnstile> Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2 : T" by fact 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   960
  then obtain S\<^isub>1 S\<^isub>2 e\<^isub>1' e\<^isub>2' where 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   961
    a1: "\<Gamma> \<turnstile> e : Data (DSum S\<^isub>1 S\<^isub>2)" and 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   962
    a2: "((x\<^isub>1,Data S\<^isub>1)#\<Gamma>) \<turnstile> e\<^isub>1 : T" using vc by auto 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   963
  have ih1:"\<Gamma> \<turnstile> e : Data (DSum S\<^isub>1 S\<^isub>2) \<Longrightarrow> \<Gamma> \<turnstile> InL e' : Data (DSum S\<^isub>1 S\<^isub>2)" by fact 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   964
  have ih2:"\<Gamma> \<turnstile> e\<^isub>1[x\<^isub>1::=e'] : T \<Longrightarrow> \<Gamma> \<turnstile> e'' : T " by fact 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   965
  have "\<Gamma> \<turnstile> InL e' : Data (DSum S\<^isub>1 S\<^isub>2)" using ih1 a1 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   966
  then have "\<Gamma> \<turnstile> e' : Data S\<^isub>1" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   967
  then have "\<Gamma> \<turnstile> e\<^isub>1[x\<^isub>1::=e'] : T" using a2 by (simp add: typing_substitution)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   968
  then show "\<Gamma> \<turnstile> e'' : T" using ih2 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   969
next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   970
 case (b_CaseR x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' \<Gamma> T)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   971
 then show "\<Gamma> \<turnstile> e'' : T" by (blast intro: typing_substitution)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   972
qed (blast)+
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   973
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   974
lemma unicity_of_evaluation:
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   975
  assumes a: "e \<Down> e\<^isub>1" 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   976
  and     b: "e \<Down> e\<^isub>2"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   977
  shows "e\<^isub>1 = e\<^isub>2"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   978
  using a b
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   979
proof (nominal_induct e e\<^isub>1 avoiding: e\<^isub>2 rule: big_induct_strong)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   980
  case (b_Lam x e t\<^isub>2)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   981
  have "Lam [x].e \<Down> t\<^isub>2" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   982
  thus "Lam [x].e = t\<^isub>2" by (cases, simp_all add: trm.inject)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   983
next
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   984
  case (b_App x e\<^isub>1 e\<^isub>2 e\<^isub>2' e' e\<^isub>1' t\<^isub>2)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   985
  have ih1: "\<And>t. e\<^isub>1 \<Down> t \<Longrightarrow> Lam [x].e\<^isub>1' = t" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   986
  have ih2:"\<And>t. e\<^isub>2 \<Down> t \<Longrightarrow> e\<^isub>2' = t" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   987
  have ih3: "\<And>t. e\<^isub>1'[x::=e\<^isub>2'] \<Down> t \<Longrightarrow> e' = t" by fact
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   988
  have app: "App e\<^isub>1 e\<^isub>2 \<Down> t\<^isub>2" by fact
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   989
  have vc: "x\<sharp>e\<^isub>1" "x\<sharp>e\<^isub>2" by fact
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   990
  then have "x \<sharp> App e\<^isub>1 e\<^isub>2" by auto
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   991
  then have vc': "x\<sharp>t\<^isub>2" using fresh_preserved app by blast
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   992
  from vc vc' obtain f\<^isub>1 f\<^isub>2 where x1: "e\<^isub>1 \<Down> Lam [x]. f\<^isub>1" and x2: "e\<^isub>2 \<Down> f\<^isub>2" and x3: "f\<^isub>1[x::=f\<^isub>2] \<Down> t\<^isub>2"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   993
    using app by (auto simp add: fresh_prod)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   994
  then have "Lam [x]. f\<^isub>1 = Lam [x]. e\<^isub>1'" using ih1 by simp
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   995
  then 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   996
  have "f\<^isub>1 = e\<^isub>1'" by (auto simp add: trm.inject alpha) 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   997
  moreover 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   998
  have "f\<^isub>2 = e\<^isub>2'" using x2 ih2 by simp
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   999
  ultimately have "e\<^isub>1'[x::=e\<^isub>2'] \<Down> t\<^isub>2" using x3 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1000
  thus ?case using ih3 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1001
next
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1002
  case (b_CaseL  x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' t\<^isub>2)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1003
  have fs: "x\<^isub>1\<sharp>e" "x\<^isub>1\<sharp>t\<^isub>2" "x\<^isub>2\<sharp>e" "x\<^isub>2\<sharp>t\<^isub>2" by fact 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1004
  have ih1:"\<And>t. e \<Down> t \<Longrightarrow> InL e' = t" by fact 
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1005
  have ih2:"\<And>t. e\<^isub>1[x\<^isub>1::=e'] \<Down> t \<Longrightarrow> e'' = t" by fact
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1006
  have ha: "\<not>(\<exists>t. e \<Down> InR t)" using ih1 by force
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1007
  have "Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2 \<Down> t\<^isub>2" by fact
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1008
  then obtain f' where "e \<Down> InL f'" and h: "e\<^isub>1[x\<^isub>1::=f']\<Down>t\<^isub>2" using ha fs by auto
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1009
  then have "InL f' = InL e'" using ih1 by simp
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1010
  then have "f' = e'" by (simp add: trm.inject)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1011
  then have "e\<^isub>1[x\<^isub>1::=e'] \<Down> t\<^isub>2" using h by simp
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1012
  then show "e'' = t\<^isub>2" using ih2 by simp
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1013
next 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1014
  case (b_CaseR x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' t\<^isub>2 )
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1015
  have fs: "x\<^isub>1\<sharp>e" "x\<^isub>1\<sharp>t\<^isub>2" "x\<^isub>2\<sharp>e" "x\<^isub>2\<sharp>t\<^isub>2" by fact 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1016
  have ih1: "\<And>t. e \<Down> t \<Longrightarrow> InR e' = t" by fact
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1017
  have ih2: "\<And>t. e\<^isub>2[x\<^isub>2::=e'] \<Down> t \<Longrightarrow> e'' = t" by fact
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1018
  have ha: "\<not>(\<exists>t. e \<Down> InL t)" using ih1 by force
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1019
  have "Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2 \<Down> t\<^isub>2" by fact
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1020
  then obtain f' where "e \<Down> InR f'" and h: "e\<^isub>2[x\<^isub>2::=f']\<Down>t\<^isub>2"  using ha fs by auto
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1021
  then have "InR f' = InR e'" using ih1 by simp
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1022
  then have "e\<^isub>2[x\<^isub>2::=e'] \<Down> t\<^isub>2" using h by (simp add: trm.inject)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1023
  thus "e'' = t\<^isub>2" using ih2 by simp
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1024
qed (force simp add: trm.inject)+
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1025
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1026
lemma not_val_App[simp]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1027
  shows 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1028
  "\<not> val (App e\<^isub>1 e\<^isub>2)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1029
  "\<not> val (Fst e)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1030
  "\<not> val (Snd e)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1031
  "\<not> val (Var x)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1032
  "\<not> val (Case e of inl x\<^isub>1 \<rightarrow> e\<^isub>1 | inr x\<^isub>2 \<rightarrow> e\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1033
by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1034
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1035
lemma reduces_evaluates_to_values:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1036
  assumes h:"t \<Down> t'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1037
  shows "val t'"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1038
  using h by (induct) (auto)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1039
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1040
lemma type_prod_evaluates_to_pairs:
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1041
  assumes a: "\<Gamma> \<turnstile> t : Data (DProd S\<^isub>1 S\<^isub>2)" 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1042
  and     b: "t \<Down> t'"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1043
  obtains t\<^isub>1 t\<^isub>2 where "t' = Pr t\<^isub>1 t\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1044
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1045
   have "\<Gamma> \<turnstile> t' : Data (DProd S\<^isub>1 S\<^isub>2)" using assms subject_reduction by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1046
   moreover
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1047
   have "val t'" using reduces_evaluates_to_values assms by simp
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1048
   ultimately obtain t\<^isub>1 t\<^isub>2 where "t' = Pr t\<^isub>1 t\<^isub>2" by (cases, auto simp add:ty.inject data.inject)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1049
   thus ?thesis using prems by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1050
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1051
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1052
lemma type_sum_evaluates_to_ins:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1053
  assumes "\<Gamma> \<turnstile> t : Data (DSum \<sigma>\<^isub>1 \<sigma>\<^isub>2)" and "t \<Down> t'"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1054
  shows "(\<exists>t''. t' = InL t'') \<or> (\<exists>t''. t' = InR t'')"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1055
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1056
  have "\<Gamma> \<turnstile> t' : Data (DSum \<sigma>\<^isub>1 \<sigma>\<^isub>2)" using assms subject_reduction by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1057
  moreover
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1058
  have "val t'" using reduces_evaluates_to_values assms by simp
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1059
  ultimately obtain t'' where "t' = InL t'' \<or>  t' = InR t''"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1060
    by (cases, auto simp add:ty.inject data.inject)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1061
  thus ?thesis by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1062
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1063
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1064
lemma type_arrow_evaluates_to_lams:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1065
  assumes "\<Gamma> \<turnstile> t : \<sigma> \<rightarrow> \<tau>" and "t \<Down> t'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1066
  obtains  x t'' where "t' = Lam [x]. t''"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1067
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1068
  have "\<Gamma> \<turnstile> t' : \<sigma> \<rightarrow> \<tau>" using assms subject_reduction by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1069
  moreover
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1070
  have "val t'" using reduces_evaluates_to_values assms by simp
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1071
  ultimately obtain x t'' where "t' = Lam [x]. t''" by (cases, auto simp add:ty.inject data.inject)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1072
  thus ?thesis using prems by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1073
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1074
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1075
lemma type_nat_evaluates_to_consts:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1076
  assumes "\<Gamma> \<turnstile> t : Data DNat" and "t \<Down> t'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1077
  obtains n where "t' = Const n"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1078
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1079
  have "\<Gamma> \<turnstile> t' : Data DNat " using assms subject_reduction by simp
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1080
  moreover have "val t'" using reduces_evaluates_to_values assms by simp
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1081
  ultimately obtain n where "t' = Const n" by (cases, auto simp add:ty.inject data.inject)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1082
  thus ?thesis using prems by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1083
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1084
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1085
consts
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1086
  V' :: "data \<Rightarrow> trm set" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1087
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1088
nominal_primrec    
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1089
  "V' (DNat) = {Const n | n. n \<in> (UNIV::nat set)}"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1090
  "V' (DProd S\<^isub>1 S\<^isub>2) = {Pr x y | x y. x \<in> V' S\<^isub>1 \<and> y \<in> V' S\<^isub>2}"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1091
  "V' (DSum S\<^isub>1 S\<^isub>2) = {InL x | x. x \<in> V' S\<^isub>1} \<union> {InR y | y. y \<in> V' S\<^isub>2}"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1092
apply(rule TrueI)+
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1093
done
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1094
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1095
lemma Vprimes_are_values :
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1096
  fixes S::"data"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1097
  assumes h: "e \<in> V' S"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1098
  shows "val e"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1099
using h 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1100
by (nominal_induct S arbitrary: e rule:data.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1101
   (auto)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1102
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1103
lemma V'_eqvt:
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1104
  fixes pi::"name prm"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1105
  assumes a: "v \<in> V' S"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1106
  shows "(pi\<bullet>v) \<in> V' S"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1107
using a
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1108
by (nominal_induct S arbitrary: v rule: data.induct)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1109
   (auto simp add: trm.inject)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1110
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1111
consts
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1112
  V :: "ty \<Rightarrow> trm set" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1113
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1114
nominal_primrec
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1115
  "V (Data S) = V' S"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1116
  "V (T\<^isub>1 \<rightarrow> T\<^isub>2) = {Lam [x].e | x e. \<forall> v \<in> (V T\<^isub>1). \<exists> v'. e[x::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2}"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1117
apply(rule TrueI)+ 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1118
done
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1119
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1120
lemma V_eqvt:
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1121
  fixes pi::"name prm"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1122
  assumes a: "x\<in>V T"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1123
  shows "(pi\<bullet>x)\<in>V T"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1124
using a
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1125
apply(nominal_induct T arbitrary: pi x rule: ty.induct)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1126
apply(auto simp add: trm.inject perm_set_def)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1127
apply(perm_simp add: V'_eqvt)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1128
apply(rule_tac x="pi\<bullet>xa" in exI)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1129
apply(rule_tac x="pi\<bullet>e" in exI)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1130
apply(simp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1131
apply(auto)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1132
apply(drule_tac x="(rev pi)\<bullet>v" in bspec)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1133
apply(force)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1134
apply(auto)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1135
apply(rule_tac x="pi\<bullet>v'" in exI)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1136
apply(auto)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1137
apply(drule_tac pi="pi" in big_eqvt)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1138
apply(perm_simp add: eqvt)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1139
done
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1140
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1141
lemma V_arrow_elim_weak[elim] :
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1142
  assumes h:"u \<in> (V (T\<^isub>1 \<rightarrow> T\<^isub>2))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1143
  obtains a t where "u = Lam[a].t" and "\<forall> v \<in> (V T\<^isub>1). \<exists> v'. t[a::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1144
using h by (auto)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1145
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1146
lemma V_arrow_elim_strong[elim]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1147
  fixes c::"'a::fs_name"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1148
  assumes h: "u \<in> V (T\<^isub>1 \<rightarrow> T\<^isub>2)"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1149
  obtains a t where "a\<sharp>c" "u = Lam[a].t" "\<forall>v \<in> (V T\<^isub>1). \<exists> v'. t[a::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1150
using h
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1151
apply -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1152
apply(erule V_arrow_elim_weak)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1153
apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(a,t,c)") (*A*)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1154
apply(erule exE)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1155
apply(drule_tac x="a'" in meta_spec)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1156
apply(drule_tac x="[(a,a')]\<bullet>t" in meta_spec)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1157
apply(drule meta_mp)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1158
apply(simp)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1159
apply(drule meta_mp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1160
apply(simp add: trm.inject alpha fresh_left fresh_prod calc_atm fresh_atm)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1161
apply(perm_simp)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1162
apply(force)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1163
apply(drule meta_mp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1164
apply(rule ballI)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1165
apply(drule_tac x="[(a,a')]\<bullet>v" in bspec)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1166
apply(simp add: V_eqvt)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1167
apply(auto)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1168
apply(rule_tac x="[(a,a')]\<bullet>v'" in exI)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1169
apply(auto)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1170
apply(drule_tac pi="[(a,a')]" in big_eqvt)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1171
apply(perm_simp add: eqvt calc_atm)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1172
apply(simp add: V_eqvt)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1173
(*A*)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1174
apply(rule exists_fresh')
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1175
apply(simp add: fin_supp)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1176
done
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1177
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1178
lemma V_are_values :
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1179
  fixes T::"ty"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1180
  assumes h:"e \<in> V T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1181
  shows "val e"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1182
using h by (nominal_induct T arbitrary: e rule:ty.induct, auto simp add: Vprimes_are_values)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1183
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1184
lemma values_reduce_to_themselves:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1185
  assumes h:"val v"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1186
  shows "v \<Down> v"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1187
using h by (induct,auto)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1188
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1189
lemma Vs_reduce_to_themselves[simp]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1190
  assumes h:"v \<in> V T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1191
  shows "v \<Down> v"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1192
using h by (simp add: values_reduce_to_themselves V_are_values)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1193
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1194
lemma V_sum:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1195
  assumes h:"x \<in> V (Data (DSum S\<^isub>1 S\<^isub>2))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1196
  shows "(\<exists> y. x= InL y \<and>  y \<in> V' S\<^isub>1) \<or> (\<exists> y. x= InR y \<and> y \<in> V' S\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1197
using h by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1198
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1199
abbreviation 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1200
 mapsto :: "(name\<times>trm) list \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> bool" ("_ maps _ to _" [55,55,55] 55) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1201
where
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1202
 "\<theta> maps x to e\<equiv> (lookup \<theta> x) = e"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1203
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1204
abbreviation 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1205
  v_closes :: "(name\<times>trm) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ Vcloses _" [55,55] 55) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1206
where
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1207
  "\<theta> Vcloses \<Gamma> \<equiv> \<forall>x T. ((x,T) \<in> set \<Gamma> \<longrightarrow> (\<exists>v. \<theta> maps x to v \<and> v \<in> (V T)))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1208
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1209
lemma monotonicity:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1210
  fixes m::"name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1211
  fixes \<theta>::"(name \<times> trm) list" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1212
  assumes h1: "\<theta> Vcloses \<Gamma>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1213
  and     h2: "e \<in> V T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1214
  and     h3: "valid ((x,T)#\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1215
  shows "(x,e)#\<theta> Vcloses (x,T)#\<Gamma>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1216
proof(intro strip)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1217
  fix x' T'
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1218
  assume "(x',T') \<in> set ((x,T)#\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1219
  then have "((x',T')=(x,T)) \<or> ((x',T')\<in>set \<Gamma> \<and> x'\<noteq>x)" using h3 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1220
    by (rule_tac case_distinction_on_context)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1221
  moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1222
  { (* first case *)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1223
    assume "(x',T') = (x,T)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1224
    then have "\<exists>e'. ((x,e)#\<theta>) maps x to e' \<and> e' \<in> V T'" using h2 by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1225
  }
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1226
  moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1227
  { (* second case *)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1228
    assume "(x',T') \<in> set \<Gamma>" and neq:"x' \<noteq> x"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1229
      then have "\<exists>e'. \<theta> maps x' to e' \<and> e' \<in> V T'" using h1 by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1230
      then have "\<exists>e'. ((x,e)#\<theta>) maps x' to e' \<and> e' \<in> V T'" using neq by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1231
  }
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1232
  ultimately show "\<exists>e'.  ((x,e)#\<theta>) maps x' to e'  \<and> e' \<in> V T'" by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1233
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1234
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1235
lemma termination_aux:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1236
  fixes T :: "ty"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1237
  fixes \<Gamma> :: "(name \<times> ty) list"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1238
  fixes \<theta> :: "(name \<times> trm) list"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1239
  fixes e :: "trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1240
  assumes h1: "\<Gamma> \<turnstile> e : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1241
  and     h2: "\<theta> Vcloses \<Gamma>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1242
  shows "\<exists>v. \<theta><e> \<Down> v \<and> v \<in> V T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1243
using h2 h1
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1244
proof(nominal_induct e avoiding: \<Gamma> \<theta> arbitrary: T rule: trm.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1245
  case (App e\<^isub>1 e\<^isub>2 \<Gamma> \<theta> T)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1246
  have ih\<^isub>1:"\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e\<^isub>1 : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e\<^isub>1> \<Down> v \<and> v \<in> V T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1247
  have ih\<^isub>2:"\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e\<^isub>2 : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e\<^isub>2> \<Down> v \<and> v \<in> V T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1248
  have as\<^isub>1:"\<theta> Vcloses \<Gamma>" by fact 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1249
  have as\<^isub>2: "\<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1250
  from as\<^isub>2 obtain T' where "\<Gamma> \<turnstile> e\<^isub>1 : T' \<rightarrow> T" and "\<Gamma> \<turnstile> e\<^isub>2 : T'" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1251
  then obtain v\<^isub>1 v\<^isub>2 where "(i)": "\<theta><e\<^isub>1> \<Down> v\<^isub>1" "v\<^isub>1 \<in> V (T' \<rightarrow> T)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1252
                      and "(ii)":"\<theta><e\<^isub>2> \<Down> v\<^isub>2" "v\<^isub>2 \<in> V T'" using ih\<^isub>1 ih\<^isub>2 as\<^isub>1 by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1253
  from "(i)" obtain x e' 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1254
            where "v\<^isub>1 = Lam[x].e'" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1255
            and "(iii)": "(\<forall>v \<in> (V T').\<exists> v'. e'[x::=v] \<Down> v' \<and> v' \<in> V T)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1256
            and "(iv)":  "\<theta><e\<^isub>1> \<Down> Lam [x].e'" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1257
            and fr: "x\<sharp>(\<theta>,e\<^isub>1,e\<^isub>2)" by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1258
  from fr have fr\<^isub>1: "x\<sharp>\<theta><e\<^isub>1>" and fr\<^isub>2: "x\<sharp>\<theta><e\<^isub>2>" by (simp_all add: fresh_psubst)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1259
  from "(ii)" "(iii)" obtain v\<^isub>3 where "(v)": "e'[x::=v\<^isub>2] \<Down> v\<^isub>3" "v\<^isub>3 \<in> V T" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1260
  from fr\<^isub>2 "(ii)" have "x\<sharp>v\<^isub>2" by (simp add: fresh_preserved)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1261
  then have "x\<sharp>e'[x::=v\<^isub>2]" by (simp add: fresh_subst_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1262
  then have fr\<^isub>3: "x\<sharp>v\<^isub>3" using "(v)" by (simp add: fresh_preserved)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1263
  from fr\<^isub>1 fr\<^isub>2 fr\<^isub>3 have "x\<sharp>(\<theta><e\<^isub>1>,\<theta><e\<^isub>2>,v\<^isub>3)" by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1264
  with "(iv)" "(ii)" "(v)" have "App (\<theta><e\<^isub>1>) (\<theta><e\<^isub>2>) \<Down> v\<^isub>3" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1265
  then show "\<exists>v. \<theta><App e\<^isub>1 e\<^isub>2> \<Down> v \<and> v \<in> V T" using "(v)" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1266
next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1267
  case (Pr t\<^isub>1 t\<^isub>2 \<Gamma> \<theta> T)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1268
  have "\<Gamma> \<turnstile> Pr t\<^isub>1 t\<^isub>2 : T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1269
  then obtain T\<^isub>a T\<^isub>b where ta:"\<Gamma> \<turnstile> t\<^isub>1 : Data T\<^isub>a" and "\<Gamma> \<turnstile> t\<^isub>2 : Data T\<^isub>b" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1270
                      and eq:"T=Data (DProd T\<^isub>a T\<^isub>b)" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1271
  have h:"\<theta> Vcloses \<Gamma>" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1272
  then obtain v\<^isub>1 v\<^isub>2 where "\<theta><t\<^isub>1> \<Down> v\<^isub>1 \<and> v\<^isub>1 \<in> V (Data T\<^isub>a)" "\<theta><t\<^isub>2> \<Down> v\<^isub>2 \<and> v\<^isub>2 \<in> V (Data T\<^isub>b)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1273
    using prems by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1274
  thus "\<exists>v. \<theta><Pr t\<^isub>1 t\<^isub>2> \<Down> v \<and> v \<in> V T" using eq by auto
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1275
next 
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1276
  case (Lam x e \<Gamma> \<theta> T)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1277
  have ih:"\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e> \<Down> v \<and> v \<in> V T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1278
  have as\<^isub>1: "\<theta> Vcloses \<Gamma>" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1279
  have as\<^isub>2: "\<Gamma> \<turnstile> Lam [x].e : T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1280
  have fs: "x\<sharp>\<Gamma>" "x\<sharp>\<theta>" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1281
  from as\<^isub>2 fs obtain T\<^isub>1 T\<^isub>2 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1282
    where "(i)": "(x,T\<^isub>1)#\<Gamma> \<turnstile> e:T\<^isub>2" and "(ii)": "T = T\<^isub>1 \<rightarrow> T\<^isub>2" by auto
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1283
  from "(i)" have "(iii)": "valid ((x,T\<^isub>1)#\<Gamma>)" by (simp add: typing_implies_valid)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1284
  have "\<forall>v \<in> (V T\<^isub>1). \<exists>v'. (\<theta><e>)[x::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1285
  proof
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1286
    fix v
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1287
    assume "v \<in> (V T\<^isub>1)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1288
    with "(iii)" as\<^isub>1 have "(x,v)#\<theta> Vcloses (x,T\<^isub>1)#\<Gamma>" using monotonicity by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1289
    with ih "(i)" obtain v' where "((x,v)#\<theta>)<e> \<Down> v' \<and> v' \<in> V T\<^isub>2" by blast
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1290
    then have "\<theta><e>[x::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2" using fs by (simp add: psubst_subst_psubst)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1291
    then show "\<exists>v'. \<theta><e>[x::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1292
  qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1293
  then have "Lam[x].\<theta><e> \<in> V (T\<^isub>1 \<rightarrow> T\<^isub>2)" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1294
  then have "\<theta><Lam [x].e> \<Down> Lam[x].\<theta><e> \<and> Lam[x].\<theta><e> \<in> V (T\<^isub>1\<rightarrow>T\<^isub>2)" using fs by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1295
  thus "\<exists>v. \<theta><Lam [x].e> \<Down> v \<and> v \<in> V T" using "(ii)" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1296
next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1297
  case (Case t' n\<^isub>1 t\<^isub>1 n\<^isub>2 t\<^isub>2 \<Gamma> \<theta> T)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1298
  have f: "n\<^isub>1\<sharp>\<Gamma>" "n\<^isub>1\<sharp>\<theta>" "n\<^isub>2\<sharp>\<Gamma>" "n\<^isub>2\<sharp>\<theta>" "n\<^isub>2\<noteq>n\<^isub>1" "n\<^isub>1\<sharp>t'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1299
  "n\<^isub>1\<sharp>t\<^isub>2" "n\<^isub>2\<sharp>t'" "n\<^isub>2\<sharp>t\<^isub>1" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1300
  have h:"\<theta> Vcloses \<Gamma>" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1301
  have th:"\<Gamma> \<turnstile> Case t' of inl n\<^isub>1 \<rightarrow> t\<^isub>1 | inr n\<^isub>2 \<rightarrow> t\<^isub>2 : T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1302
  then obtain S\<^isub>1 S\<^isub>2 where 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1303
    hm:"\<Gamma> \<turnstile> t' : Data (DSum S\<^isub>1 S\<^isub>2)" and
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1304
    hl:"(n\<^isub>1,Data S\<^isub>1)#\<Gamma> \<turnstile> t\<^isub>1 : T" and
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1305
    hr:"(n\<^isub>2,Data S\<^isub>2)#\<Gamma> \<turnstile> t\<^isub>2 : T" using f by auto
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1306
  then obtain v\<^isub>0 where ht':"\<theta><t'> \<Down> v\<^isub>0" and hS:"v\<^isub>0 \<in> V (Data (DSum S\<^isub>1 S\<^isub>2))" using prems h by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1307
  (* We distinguish between the cases InL and InR *)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1308
  { fix v\<^isub>0'
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1309
    assume eqc:"v\<^isub>0 = InL v\<^isub>0'" and "v\<^isub>0' \<in> V' S\<^isub>1"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1310
    then have inc: "v\<^isub>0' \<in> V (Data S\<^isub>1)" by auto
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1311
    have "valid \<Gamma>" using th typing_implies_valid by auto
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1312
    then moreover have "valid ((n\<^isub>1,Data S\<^isub>1)#\<Gamma>)" using f by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1313
    then moreover have "(n\<^isub>1,v\<^isub>0')#\<theta> Vcloses (n\<^isub>1,Data S\<^isub>1)#\<Gamma>" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1314
      using inc h monotonicity by blast
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1315
    moreover 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1316
    have ih:"\<And>\<Gamma> \<theta> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> t\<^isub>1 : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><t\<^isub>1> \<Down> v \<and> v \<in> V T" by fact
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1317
    ultimately obtain v\<^isub>1 where ho: "((n\<^isub>1,v\<^isub>0')#\<theta>)<t\<^isub>1> \<Down> v\<^isub>1 \<and> v\<^isub>1 \<in> V T" using hl by blast
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1318
    then have r:"\<theta><t\<^isub>1>[n\<^isub>1::=v\<^isub>0'] \<Down> v\<^isub>1 \<and> v\<^isub>1 \<in> V T" using psubst_subst_psubst f by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1319
    then moreover have "n\<^isub>1\<sharp>(\<theta><t'>,\<theta><t\<^isub>2>,v\<^isub>1,n\<^isub>2)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1320
      proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1321
	have "n\<^isub>1\<sharp>v\<^isub>0" using ht' fresh_preserved fresh_psubst f by auto 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1322
	then have "n\<^isub>1\<sharp>v\<^isub>0'" using eqc by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1323
	then have "n\<^isub>1\<sharp>v\<^isub>1" using f r fresh_preserved fresh_subst_fresh by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1324
	thus "n\<^isub>1\<sharp>(\<theta><t'>,\<theta><t\<^isub>2>,v\<^isub>1,n\<^isub>2)" using f by (simp add: fresh_atm fresh_psubst)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1325
      qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1326
    moreover have "n\<^isub>2\<sharp>(\<theta><t'>,\<theta><t\<^isub>1>,v\<^isub>1,n\<^isub>1)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1327
      proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1328
	have "n\<^isub>2\<sharp>v\<^isub>0" using ht' fresh_preserved fresh_psubst f by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1329
	then have "n\<^isub>2\<sharp>v\<^isub>0'" using eqc by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1330
	then have "n\<^isub>2\<sharp>((n\<^isub>1,v\<^isub>0')#\<theta>)" using f fresh_list_cons fresh_atm by force 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1331
	then have "n\<^isub>2\<sharp>((n\<^isub>1,v\<^isub>0')#\<theta>)<t\<^isub>1>" using f fresh_psubst by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1332
	moreover then have "n\<^isub>2 \<sharp> v\<^isub>1" using fresh_preserved ho by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1333
	ultimately show  "n\<^isub>2\<sharp>(\<theta><t'>,\<theta><t\<^isub>1>,v\<^isub>1,n\<^isub>1)" using f by (simp add: fresh_psubst fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1334
      qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1335
    ultimately have "Case \<theta><t'> of inl n\<^isub>1 \<rightarrow> \<theta><t\<^isub>1> | inr n\<^isub>2 \<rightarrow> \<theta><t\<^isub>2> \<Down> v\<^isub>1 \<and> v\<^isub>1 \<in> V T" using ht' eqc by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1336
    moreover 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1337
      have "Case \<theta><t'> of inl n\<^isub>1 \<rightarrow> \<theta><t\<^isub>1> | inr n\<^isub>2 \<rightarrow> \<theta><t\<^isub>2> = \<theta><Case t' of inl n\<^isub>1 \<rightarrow> t\<^isub>1 | inr n\<^isub>2 \<rightarrow> t\<^isub>2>" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1338
      using f by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1339
    ultimately have "\<exists>v. \<theta><Case t' of inl n\<^isub>1 \<rightarrow> t\<^isub>1 | inr n\<^isub>2 \<rightarrow> t\<^isub>2> \<Down> v \<and> v \<in> V T" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1340
  }
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1341
  moreover 
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1342
  { fix v\<^isub>0'
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1343
    assume eqc:"v\<^isub>0 = InR v\<^isub>0'" and "v\<^isub>0' \<in> V' S\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1344
    then have inc:"v\<^isub>0' \<in> V (Data S\<^isub>2)" by auto
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1345
    have "valid \<Gamma>" using th typing_implies_valid by auto
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1346
    then moreover have "valid ((n\<^isub>2,Data S\<^isub>2)#\<Gamma>)" using f by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1347
    then moreover have "(n\<^isub>2,v\<^isub>0')#\<theta> Vcloses (n\<^isub>2,Data S\<^isub>2)#\<Gamma>" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1348
      using inc h monotonicity by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1349
    moreover have ih:"\<And>\<Gamma> \<theta> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> t\<^isub>2 : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><t\<^isub>2> \<Down> v \<and> v \<in> V T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1350
    ultimately obtain v\<^isub>2 where ho:"((n\<^isub>2,v\<^isub>0')#\<theta>)<t\<^isub>2> \<Down> v\<^isub>2 \<and> v\<^isub>2 \<in> V T" using hr by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1351
    then have r:"\<theta><t\<^isub>2>[n\<^isub>2::=v\<^isub>0'] \<Down> v\<^isub>2 \<and> v\<^isub>2 \<in> V T" using psubst_subst_psubst f by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1352
    moreover have "n\<^isub>1\<sharp>(\<theta><t'>,\<theta><t\<^isub>2>,v\<^isub>2,n\<^isub>2)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1353
    proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1354
      have "n\<^isub>1\<sharp>\<theta><t'>" using fresh_psubst f by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1355
      then have "n\<^isub>1\<sharp>v\<^isub>0" using ht' fresh_preserved by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1356
      then have "n\<^isub>1\<sharp>v\<^isub>0'" using eqc by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1357
      then have "n\<^isub>1\<sharp>((n\<^isub>2,v\<^isub>0')#\<theta>)" using f fresh_list_cons fresh_atm by force 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1358
      then have "n\<^isub>1\<sharp>((n\<^isub>2,v\<^isub>0')#\<theta>)<t\<^isub>2>" using f fresh_psubst by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1359
      moreover then have "n\<^isub>1\<sharp>v\<^isub>2" using fresh_preserved ho by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1360
      ultimately show  "n\<^isub>1 \<sharp> (\<theta><t'>,\<theta><t\<^isub>2>,v\<^isub>2,n\<^isub>2)" using f by (simp add: fresh_psubst fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1361
    qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1362
    moreover have "n\<^isub>2 \<sharp> (\<theta><t'>,\<theta><t\<^isub>1>,v\<^isub>2,n\<^isub>1)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1363
      proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1364
	have "n\<^isub>2\<sharp>\<theta><t'>" using fresh_psubst f by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1365
	then have "n\<^isub>2\<sharp>v\<^isub>0" using ht' fresh_preserved by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1366
	then have "n\<^isub>2\<sharp>v\<^isub>0'" using eqc by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1367
	then have "n\<^isub>2\<sharp>\<theta><t\<^isub>2>[n\<^isub>2::=v\<^isub>0']" using f fresh_subst_fresh by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1368
	then have "n\<^isub>2\<sharp>v\<^isub>2" using f fresh_preserved r by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1369
	then show "n\<^isub>2\<sharp>(\<theta><t'>,\<theta><t\<^isub>1>,v\<^isub>2,n\<^isub>1)" using f by (simp add: fresh_atm fresh_psubst)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1370
      qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1371
    ultimately have "Case \<theta><t'> of inl n\<^isub>1 \<rightarrow> \<theta><t\<^isub>1> | inr n\<^isub>2 \<rightarrow> \<theta><t\<^isub>2> \<Down> v\<^isub>2 \<and> v\<^isub>2 \<in> V T" using ht' eqc by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1372
    then have "\<exists>v. \<theta><Case t' of inl n\<^isub>1 \<rightarrow> t\<^isub>1 | inr n\<^isub>2 \<rightarrow> t\<^isub>2> \<Down> v \<and> v \<in> V T" using f by auto
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
  1373
  }
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1374
  ultimately show "\<exists>v. \<theta><Case t' of inl n\<^isub>1 \<rightarrow> t\<^isub>1 | inr n\<^isub>2 \<rightarrow> t\<^isub>2> \<Down> v \<and> v \<in> V T" using hS V_sum by blast
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1375
qed (force)+
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1376
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1377
theorem termination_of_evaluation:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1378
  assumes a: "[] \<turnstile> e : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1379
  shows "\<exists>v. e \<Down> v \<and> val v"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1380
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1381
  from a have "\<exists>v. (([]::(name \<times> trm) list)<e>) \<Down> v \<and> v \<in> V T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1382
    by (rule termination_aux) (auto)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1383
  thus "\<exists>v. e \<Down> v \<and> val v" using V_are_values by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1384
qed 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1385
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
  1386
end