src/HOL/Nominal/Examples/SOS.thy
author urbanc
Fri, 04 Jan 2008 16:35:22 +0100
changeset 25832 41a014cc44c0
parent 24678 232e71c2a6d9
child 25867 c24395ea4e71
permissions -rw-r--r--
partially adapted to new inversion rules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
     1
(* "$Id$" *)
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
     2
(*                                                     *)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
     3
(* Formalisation of some typical SOS-proofs.           *)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
     4
(*                                                     *) 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
     5
(* This work arose from a challenge suggested by Adam  *)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
     6
(* Chlipala on the POPLmark mailing list.              *)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
     7
(*                                                     *) 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
     8
(* We thank Nick Benton for helping us with the        *) 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
     9
(* termination-proof for evaluation.                   *)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    10
(*                                                     *)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    11
(* The formalisation was done by Julien Narboux and    *)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    12
(* Christian Urban.                                    *)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    13
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    14
theory SOS
22534
bd4b954e85ee adapted to nominal_inductive
urbanc
parents: 22531
diff changeset
    15
  imports "../Nominal"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    16
begin
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    17
23158
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 22730
diff changeset
    18
atom_decl name
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    19
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    20
text {* types and terms *}
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    21
nominal_datatype ty = 
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    22
    TVar "nat"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    23
  | Arrow "ty" "ty" ("_\<rightarrow>_" [100,100] 100)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    24
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    25
nominal_datatype trm = 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    26
    Var "name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    27
  | Lam "\<guillemotleft>name\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    28
  | App "trm" "trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    29
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    30
lemma fresh_ty:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    31
  fixes x::"name" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    32
  and   T::"ty"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    33
  shows "x\<sharp>T"
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    34
by (induct T rule: ty.weak_induct)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    35
   (auto simp add: fresh_nat)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    36
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    37
text {* Parallel and single substitution. *}
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    38
fun
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    39
  lookup :: "(name\<times>trm) list \<Rightarrow> name \<Rightarrow> trm"   
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    40
where
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    41
  "lookup [] x        = Var x"
22502
baee64dbe8ea fixed function syntax
krauss
parents: 22472
diff changeset
    42
| "lookup ((y,e)#\<theta>) x = (if x=y then e else lookup \<theta> x)"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    43
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    44
lemma lookup_eqvt[eqvt]:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    45
  fixes pi::"name prm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    46
  and   \<theta>::"(name\<times>trm) list"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    47
  and   X::"name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    48
  shows "pi\<bullet>(lookup \<theta> X) = lookup (pi\<bullet>\<theta>) (pi\<bullet>X)"
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    49
by (induct \<theta>) (auto simp add: eqvts)
22447
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 lookup_fresh:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    52
  fixes z::"name"
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    53
  assumes a: "z\<sharp>\<theta>" and b: "z\<sharp>x"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    54
  shows "z \<sharp>lookup \<theta> x"
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    55
using a b
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    56
by (induct rule: lookup.induct) (auto simp add: fresh_list_cons)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    57
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    58
lemma lookup_fresh':
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    59
  assumes "z\<sharp>\<theta>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    60
  shows "lookup \<theta> z = Var z"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    61
using assms 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    62
by (induct rule: lookup.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    63
   (auto simp add: fresh_list_cons fresh_prod fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    64
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    65
(* parallel substitution *)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    66
consts
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    67
  psubst :: "(name\<times>trm) list \<Rightarrow> trm \<Rightarrow> trm"  ("_<_>" [95,95] 105)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    68
 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    69
nominal_primrec
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    70
  "\<theta><(Var x)> = (lookup \<theta> x)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    71
  "\<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
    72
  "x\<sharp>\<theta> \<Longrightarrow> \<theta><(Lam [x].e)> = Lam [x].(\<theta><e>)"
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    73
apply(finite_guess)+
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
    74
apply(rule TrueI)+
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
    75
apply(simp add: abs_fresh)+
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    76
apply(fresh_guess)+
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
    77
done
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    78
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    79
lemma psubst_eqvt[eqvt]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    80
  fixes pi::"name prm" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    81
  and   t::"trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    82
  shows "pi\<bullet>(\<theta><t>) = (pi\<bullet>\<theta>)<(pi\<bullet>t)>"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
    83
by (nominal_induct t avoiding: \<theta> rule: trm.induct)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
    84
   (perm_simp add: fresh_bij lookup_eqvt)+
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    85
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    86
lemma fresh_psubst: 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    87
  fixes z::"name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    88
  and   t::"trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    89
  assumes "z\<sharp>t" and "z\<sharp>\<theta>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    90
  shows "z\<sharp>(\<theta><t>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    91
using assms
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    92
by (nominal_induct t avoiding: z \<theta> t rule: trm.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    93
   (auto simp add: abs_fresh lookup_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
    94
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    95
lemma psubst_empty[simp]:
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    96
  shows "[]<t> = t"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    97
  by (nominal_induct t rule: trm.induct) 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    98
     (auto simp add: fresh_list_nil)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
    99
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   100
(* Single substitution *)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   101
abbreviation 
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   102
  subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   103
where 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   104
  "t[x::=t']  \<equiv> ([(x,t')])<t>" 
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   105
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   106
lemma subst[simp]:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   107
  shows "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   108
  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
   109
  and   "x\<sharp>(y,t') \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   110
by (simp_all add: fresh_list_cons fresh_list_nil)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   111
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   112
lemma fresh_subst:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   113
  fixes z::"name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   114
  and   t\<^isub>1::"trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   115
  and   t2::"trm"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   116
  assumes "z\<sharp>t\<^isub>1" and "z\<sharp>t\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   117
  shows "z\<sharp>t\<^isub>1[y::=t\<^isub>2]"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   118
using assms 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   119
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
   120
   (auto simp add: abs_fresh fresh_atm)
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_subst':
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   123
  assumes "x\<sharp>e'"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   124
  shows "x\<sharp>e[x::=e']"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   125
  using assms 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   126
by (nominal_induct e avoiding: x e' rule: trm.induct)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   127
   (auto simp add: fresh_atm abs_fresh fresh_nat) 
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   128
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   129
lemma forget: 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   130
  fixes x::"name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   131
  and   L::"trm"
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   132
  assumes "x\<sharp>e" 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   133
  shows "e[x::=e'] = e"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   134
  using assms
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   135
  by (nominal_induct e avoiding: x e' rule: trm.induct)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   136
     (auto simp add: fresh_atm abs_fresh)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   137
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   138
lemma psubst_subst_psubst:
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   139
  assumes h:"x\<sharp>\<theta>"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   140
  shows "\<theta><e>[x::=e'] = ((x,e')#\<theta>)<e>"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   141
  using h
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   142
apply(nominal_induct e avoiding: \<theta> x e' rule: trm.induct)
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   143
apply(auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh')
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   144
done
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   145
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   146
text {* Typing Judgements *}
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   147
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23450
diff changeset
   148
inductive
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   149
  valid :: "(name\<times>ty) list \<Rightarrow> bool"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   150
where
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   151
  v_nil[intro]:  "valid []"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   152
| v_cons[intro]: "\<lbrakk>valid \<Gamma>;x\<sharp>\<Gamma>\<rbrakk> \<Longrightarrow> valid ((x,T)#\<Gamma>)"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   153
22534
bd4b954e85ee adapted to nominal_inductive
urbanc
parents: 22531
diff changeset
   154
equivariance valid 
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   155
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   156
inductive_cases
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   157
  valid_cons_inv_auto[elim]: "valid ((x,T)#\<Gamma>)"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   158
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   159
lemma type_unicity_in_context:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   160
  assumes asm1: "(x,t\<^isub>2) \<in> set ((x,t\<^isub>1)#\<Gamma>)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   161
  and     asm2: "valid ((x,t\<^isub>1)#\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   162
  shows "t\<^isub>1=t\<^isub>2"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   163
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   164
  from asm2 have "x\<sharp>\<Gamma>" by (cases, auto)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   165
  then have "(x,t\<^isub>2) \<notin> set \<Gamma>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   166
    by (induct \<Gamma>) (auto simp add: fresh_list_cons fresh_prod fresh_atm)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   167
  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
   168
  then show "t\<^isub>1 = t\<^isub>2" by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   169
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   170
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   171
lemma case_distinction_on_context:
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   172
  fixes \<Gamma>::"(name\<times>ty) list"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   173
  assumes asm1: "valid ((m,t)#\<Gamma>)" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   174
  and     asm2: "(n,U) \<in> set ((m,T)#\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   175
  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
   176
proof -
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   177
  from asm2 have "(n,U) \<in> set [(m,T)] \<or> (n,U) \<in> set \<Gamma>" by auto
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   178
  moreover
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   179
  { assume eq: "m=n"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   180
    assume "(n,U) \<in> set \<Gamma>" 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   181
    then have "\<not> n\<sharp>\<Gamma>" 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   182
      by (induct \<Gamma>) (auto simp add: fresh_list_cons fresh_prod fresh_atm)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   183
    moreover have "m\<sharp>\<Gamma>" using asm1 by auto
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   184
    ultimately have False using eq by auto
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   185
  }
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   186
  ultimately show ?thesis by auto
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   187
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   188
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23450
diff changeset
   189
inductive
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   190
  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
   191
where
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   192
  t_Var[intro]:   "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   193
| 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"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   194
| 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
   195
22730
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
   196
equivariance typing
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
   197
22531
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22502
diff changeset
   198
nominal_inductive typing
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   199
  by (simp_all add: abs_fresh fresh_ty)
22531
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22502
diff changeset
   200
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   201
lemma typing_implies_valid:
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   202
  assumes a: "\<Gamma> \<turnstile> t : T"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   203
  shows "valid \<Gamma>"
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   204
  using a
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   205
  by (induct) (auto)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   206
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   207
lemma t_Lam_elim: 
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   208
  assumes a1:"\<Gamma> \<turnstile> Lam [x].t : T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   209
  and     a2: "x\<sharp>\<Gamma>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   210
  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"
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   211
using a1 a2
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   212
by (cases rule: typing.strong_cases [where x="x"])
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   213
   (auto simp add: abs_fresh fresh_ty alpha trm.inject)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   214
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   215
abbreviation
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   216
  "sub_context" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [55,55] 55)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   217
where
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   218
  "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2 \<equiv> \<forall>x T. (x,T)\<in>set \<Gamma>\<^isub>1 \<longrightarrow> (x,T)\<in>set \<Gamma>\<^isub>2"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   219
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   220
lemma weakening: 
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   221
  fixes \<Gamma>\<^isub>1 \<Gamma>\<^isub>2::"(name\<times>ty) list"
22650
0c5b22076fb3 tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents: 22594
diff changeset
   222
  assumes "\<Gamma>\<^isub>1 \<turnstile> e: T" and "valid \<Gamma>\<^isub>2" and "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   223
  shows "\<Gamma>\<^isub>2 \<turnstile> e: T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   224
  using assms
22534
bd4b954e85ee adapted to nominal_inductive
urbanc
parents: 22531
diff changeset
   225
proof(nominal_induct \<Gamma>\<^isub>1 e T avoiding: \<Gamma>\<^isub>2 rule: typing.strong_induct)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   226
  case (t_Lam x \<Gamma>\<^isub>1 T\<^isub>1 t T\<^isub>2 \<Gamma>\<^isub>2)
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   227
  have vc: "x\<sharp>\<Gamma>\<^isub>2" by fact
22650
0c5b22076fb3 tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents: 22594
diff changeset
   228
  have ih: "\<lbrakk>valid ((x,T\<^isub>1)#\<Gamma>\<^isub>2); (x,T\<^isub>1)#\<Gamma>\<^isub>1 \<subseteq> (x,T\<^isub>1)#\<Gamma>\<^isub>2\<rbrakk> \<Longrightarrow> (x,T\<^isub>1)#\<Gamma>\<^isub>2 \<turnstile> t : T\<^isub>2" by fact
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   229
  have "valid \<Gamma>\<^isub>2" by fact
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   230
  then have "valid ((x,T\<^isub>1)#\<Gamma>\<^isub>2)" using vc by auto
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   231
  moreover
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   232
  have "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2" by fact
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   233
  then have "(x,T\<^isub>1)#\<Gamma>\<^isub>1 \<subseteq> (x,T\<^isub>1)#\<Gamma>\<^isub>2" by simp
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   234
  ultimately have "(x,T\<^isub>1)#\<Gamma>\<^isub>2 \<turnstile> t : T\<^isub>2" using ih by simp 
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   235
  with vc show "\<Gamma>\<^isub>2 \<turnstile> Lam [x].t : T\<^isub>1\<rightarrow>T\<^isub>2" by auto
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   236
qed (auto)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   237
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   238
lemma context_exchange:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   239
  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
   240
  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
   241
proof -
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   242
  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
   243
  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
   244
    by (auto simp: fresh_list_cons fresh_atm[symmetric])
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   245
  then have "valid ((x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\<Gamma>)"
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   246
    by (auto simp: fresh_list_cons fresh_atm fresh_ty)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   247
  moreover 
22650
0c5b22076fb3 tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents: 22594
diff changeset
   248
  have "(x\<^isub>1,T\<^isub>1)#(x\<^isub>2,T\<^isub>2)#\<Gamma> \<subseteq> (x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\<Gamma>" by auto
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   249
  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
   250
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   251
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   252
lemma typing_var_unicity: 
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   253
  assumes a: "(x,T\<^isub>1)#\<Gamma> \<turnstile> Var x : T\<^isub>2"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   254
  shows "T\<^isub>1=T\<^isub>2"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   255
proof - 
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   256
  have "(x,T\<^isub>2) \<in> set ((x,T\<^isub>1)#\<Gamma>)" using a
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   257
    by (cases) (auto simp add: trm.inject)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   258
  moreover 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   259
  have "valid ((x,T\<^isub>1)#\<Gamma>)" using a by (simp add: typing_implies_valid)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   260
  ultimately show "T\<^isub>1=T\<^isub>2" by (simp only: type_unicity_in_context)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   261
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   262
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   263
lemma typing_substitution: 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   264
  fixes \<Gamma>::"(name \<times> ty) list"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   265
  assumes "(x,T')#\<Gamma> \<turnstile> e : T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   266
  and     "\<Gamma> \<turnstile> e': T'" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   267
  shows "\<Gamma> \<turnstile> e[x::=e'] : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   268
  using assms
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   269
proof (nominal_induct e avoiding: \<Gamma> e' x arbitrary: T rule: trm.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   270
  case (Var y \<Gamma> e' x T)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   271
  have h1: "(x,T')#\<Gamma> \<turnstile> Var y : T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   272
  have h2: "\<Gamma> \<turnstile> e' : T'" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   273
  show "\<Gamma> \<turnstile> (Var y)[x::=e'] : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   274
  proof (cases "x=y")
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   275
    case True
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   276
    assume as: "x=y"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   277
    then have "T=T'" using h1 typing_var_unicity by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   278
    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
   279
  next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   280
    case False
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   281
    assume as: "x\<noteq>y" 
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   282
    have "(y,T) \<in> set ((x,T')#\<Gamma>)" using h1 by (cases) (auto simp add: trm.inject)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   283
    then have "(y,T) \<in> set \<Gamma>" using as by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   284
    moreover 
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   285
    have "valid \<Gamma>" using h2 by (simp only: typing_implies_valid)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   286
    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
   287
  qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   288
next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   289
  case (Lam y t \<Gamma> e' x T)
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   290
  have vc: "y\<sharp>\<Gamma>" "y\<sharp>x" "y\<sharp>e'" by fact
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   291
  have pr1: "\<Gamma> \<turnstile> e' : T'" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   292
  have pr2: "(x,T')#\<Gamma> \<turnstile> Lam [y].t : T" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   293
  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" 
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   294
    using vc by (auto elim: t_Lam_elim simp add: fresh_list_cons fresh_ty)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   295
  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
   296
  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
   297
  have "valid \<Gamma>" using pr1 by (simp add: typing_implies_valid)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   298
  then have "valid ((y,T\<^isub>1)#\<Gamma>)" using vc by auto
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   299
  then have "(y,T\<^isub>1)#\<Gamma> \<turnstile> e' : T'" using pr1 by (simp add: weakening)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   300
  then have "(y,T\<^isub>1)#\<Gamma> \<turnstile> t[x::=e'] : T\<^isub>2" using ih pr2'' by simp
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   301
  then have "\<Gamma> \<turnstile> Lam [y].(t[x::=e']) : T\<^isub>1\<rightarrow>T\<^isub>2" using vc by auto
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   302
  then show "\<Gamma> \<turnstile> (Lam [y].t)[x::=e'] : T" using vc eq by simp
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   303
next
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   304
  case (App e1 e2 \<Gamma> e' x T)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   305
  have "(x,T')#\<Gamma> \<turnstile> App e1 e2 : T" by fact
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   306
  then obtain Tn where a1: "(x,T')#\<Gamma> \<turnstile> e1 : Tn \<rightarrow> T"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   307
                   and a2: "(x,T')#\<Gamma> \<turnstile> e2 : Tn"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   308
    by (cases) (auto simp add: trm.inject)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   309
  have a3: "\<Gamma> \<turnstile> e' : T'" by fact
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   310
  have ih1: "\<lbrakk>(x,T')#\<Gamma> \<turnstile> e1 : Tn\<rightarrow>T; \<Gamma> \<turnstile> e' : T'\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> e1[x::=e'] : Tn\<rightarrow>T" by fact
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   311
  have ih2: "\<lbrakk>(x,T')#\<Gamma> \<turnstile> e2 : Tn; \<Gamma> \<turnstile> e' : T'\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> e2[x::=e'] : Tn" by fact
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   312
  then show ?case using a1 a2 a3 ih1 ih2 by auto 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   313
qed
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   314
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   315
text {* Values *}
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   316
inductive
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   317
  val :: "trm\<Rightarrow>bool" 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   318
where
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   319
  v_Lam[intro]:   "val (Lam [x].e)"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   320
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   321
equivariance val 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   322
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   323
lemma not_val_App[simp]:
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   324
  shows 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   325
  "\<not> val (App e\<^isub>1 e\<^isub>2)" 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   326
  "\<not> val (Var x)"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   327
  by (auto elim: val.cases)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   328
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   329
text {* Big-Step Evaluation *}
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   330
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23450
diff changeset
   331
inductive
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   332
  big :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80) 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   333
where
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   334
  b_Lam[intro]:   "Lam [x].e \<Down> Lam [x].e"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   335
| 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
   336
22730
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
   337
equivariance big
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
   338
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   339
nominal_inductive big
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   340
  by (simp_all add: abs_fresh)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   341
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   342
lemma big_preserves_fresh:
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   343
  fixes x::"name"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   344
  assumes a: "e \<Down> e'" "x\<sharp>e" 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   345
  shows "x\<sharp>e'"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   346
  using a by (induct) (auto simp add: abs_fresh fresh_subst)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   347
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   348
lemma b_App_elim:
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   349
  assumes a: "App e\<^isub>1 e\<^isub>2 \<Down> e'" "x\<sharp>(e\<^isub>1,e\<^isub>2,e')"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   350
  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'"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   351
  using a
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   352
by (cases rule: big.strong_cases[where x="x" and xa="x"])
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   353
   (auto simp add: trm.inject)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   354
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   355
lemma subject_reduction:
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   356
  assumes a: "e \<Down> e'" and b: "\<Gamma> \<turnstile> e : T"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   357
  shows "\<Gamma> \<turnstile> e' : T"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   358
  using a b
22534
bd4b954e85ee adapted to nominal_inductive
urbanc
parents: 22531
diff changeset
   359
proof (nominal_induct avoiding: \<Gamma> arbitrary: T rule: big.strong_induct) 
bd4b954e85ee adapted to nominal_inductive
urbanc
parents: 22531
diff changeset
   360
  case (b_App x e\<^isub>1 e\<^isub>2 e' e e\<^isub>2' \<Gamma> T)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   361
  have vc: "x\<sharp>\<Gamma>" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   362
  have "\<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T" by fact
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   363
  then obtain T' where a1: "\<Gamma> \<turnstile> e\<^isub>1 : T'\<rightarrow>T" and a2: "\<Gamma> \<turnstile> e\<^isub>2 : T'" 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   364
    by (cases) (auto simp add: trm.inject)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   365
  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
   366
  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
   367
  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
   368
  have "\<Gamma> \<turnstile> Lam [x].e : T'\<rightarrow>T" using ih1 a1 by simp 
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   369
  then have "((x,T')#\<Gamma>) \<turnstile> e : T" using vc  
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   370
    by (auto elim: t_Lam_elim simp add: ty.inject)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   371
  moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   372
  have "\<Gamma> \<turnstile> e\<^isub>2': T'" using ih2 a2 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   373
  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
   374
  thus "\<Gamma> \<turnstile> e' : T" using ih3 by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   375
qed (blast)+
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   376
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   377
lemma unicity_of_evaluation:
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   378
  assumes a: "e \<Down> e\<^isub>1" 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   379
  and     b: "e \<Down> e\<^isub>2"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   380
  shows "e\<^isub>1 = e\<^isub>2"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   381
  using a b
22534
bd4b954e85ee adapted to nominal_inductive
urbanc
parents: 22531
diff changeset
   382
proof (nominal_induct e e\<^isub>1 avoiding: e\<^isub>2 rule: big.strong_induct)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   383
  case (b_Lam x e t\<^isub>2)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   384
  have "Lam [x].e \<Down> t\<^isub>2" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   385
  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
   386
next
22534
bd4b954e85ee adapted to nominal_inductive
urbanc
parents: 22531
diff changeset
   387
  case (b_App x e\<^isub>1 e\<^isub>2 e' e\<^isub>1' e\<^isub>2' t\<^isub>2)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   388
  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
   389
  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
   390
  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
   391
  have app: "App e\<^isub>1 e\<^isub>2 \<Down> t\<^isub>2" by fact
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   392
  have vc: "x\<sharp>e\<^isub>1" "x\<sharp>e\<^isub>2" "x\<sharp>t\<^isub>2" by fact
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   393
  then have "x\<sharp>App e\<^isub>1 e\<^isub>2" by auto
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   394
  from app 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"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   395
    by (cases rule: big.strong_cases[where x="x" and xa="x"])
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   396
       (auto simp add: trm.inject)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   397
  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
   398
  then 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   399
  have "f\<^isub>1 = e\<^isub>1'" by (auto simp add: trm.inject alpha) 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   400
  moreover 
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   401
  have "f\<^isub>2 = e\<^isub>2'" using x2 ih2 by simp
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   402
  ultimately have "e\<^isub>1'[x::=e\<^isub>2'] \<Down> t\<^isub>2" using x3 by simp
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   403
  thus "e' = t\<^isub>2" using ih3 by simp
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   404
qed
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   405
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   406
lemma reduces_evaluates_to_values:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   407
  assumes h:"t \<Down> t'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   408
  shows "val t'"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   409
  using h by (induct) (auto)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   410
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   411
lemma type_arrow_evaluates_to_lams:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   412
  assumes "\<Gamma> \<turnstile> t : \<sigma> \<rightarrow> \<tau>" and "t \<Down> t'"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   413
  obtains  x t'' where "t' = Lam [x]. t''"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   414
proof -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   415
  have "\<Gamma> \<turnstile> t' : \<sigma> \<rightarrow> \<tau>" using assms subject_reduction by simp
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   416
  moreover
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   417
  have "val t'" using reduces_evaluates_to_values assms by simp
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   418
  ultimately obtain x t'' where "t' = Lam [x]. t''" by (cases, auto simp add:ty.inject)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   419
  thus ?thesis using prems by auto
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   420
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   421
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   422
(* Valuation *)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   423
consts
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   424
  V :: "ty \<Rightarrow> trm set" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   425
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   426
nominal_primrec
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   427
  "V (TVar x) = {e. val e}"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   428
  "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}"
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   429
  by (rule TrueI)+ 
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   430
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   431
(* can go with strong inversion rules *)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   432
lemma V_eqvt:
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   433
  fixes pi::"name prm"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   434
  assumes a: "x\<in>V T"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   435
  shows "(pi\<bullet>x)\<in>V T"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   436
using a
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   437
apply(nominal_induct T arbitrary: pi x rule: ty.induct)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   438
apply(auto simp add: trm.inject perm_set_def)
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   439
apply(simp add: eqvts)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   440
apply(rule_tac x="pi\<bullet>xa" in exI)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   441
apply(rule_tac x="pi\<bullet>e" in exI)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   442
apply(simp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   443
apply(auto)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   444
apply(drule_tac x="(rev pi)\<bullet>v" in bspec)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   445
apply(force)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   446
apply(auto)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   447
apply(rule_tac x="pi\<bullet>v'" in exI)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   448
apply(auto)
22542
8279a25ad0ae - Renamed <predicate>_eqvt to <predicate>.eqvt
berghofe
parents: 22541
diff changeset
   449
apply(drule_tac pi="pi" in big.eqvt)
22541
c33b542394f3 the name for the collection of equivariance lemmas is now eqvts (changed from eqvt) in order to avoid clashes with eqvt-lemmas generated in nominal_inductive
urbanc
parents: 22534
diff changeset
   450
apply(perm_simp add: eqvts)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   451
done
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   452
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   453
lemma V_arrow_elim_weak:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   454
  assumes h:"u \<in> (V (T\<^isub>1 \<rightarrow> T\<^isub>2))"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   455
  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
   456
using h by (auto)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   457
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   458
lemma V_arrow_elim_strong:
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   459
  fixes c::"'a::fs_name"
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   460
  assumes h: "u \<in> V (T\<^isub>1 \<rightarrow> T\<^isub>2)"
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   461
  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
   462
using h
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   463
apply -
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   464
apply(erule V_arrow_elim_weak)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   465
apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(a,t,c)") (*A*)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   466
apply(erule exE)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   467
apply(drule_tac x="a'" in meta_spec)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   468
apply(drule_tac x="[(a,a')]\<bullet>t" in meta_spec)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   469
apply(drule meta_mp)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   470
apply(simp)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   471
apply(drule meta_mp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   472
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
   473
apply(perm_simp)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   474
apply(force)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   475
apply(drule meta_mp)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   476
apply(rule ballI)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   477
apply(drule_tac x="[(a,a')]\<bullet>v" in bspec)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   478
apply(simp add: V_eqvt)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   479
apply(auto)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   480
apply(rule_tac x="[(a,a')]\<bullet>v'" in exI)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   481
apply(auto)
22542
8279a25ad0ae - Renamed <predicate>_eqvt to <predicate>.eqvt
berghofe
parents: 22541
diff changeset
   482
apply(drule_tac pi="[(a,a')]" in big.eqvt)
22541
c33b542394f3 the name for the collection of equivariance lemmas is now eqvts (changed from eqvt) in order to avoid clashes with eqvt-lemmas generated in nominal_inductive
urbanc
parents: 22534
diff changeset
   483
apply(perm_simp add: eqvts calc_atm)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   484
apply(simp add: V_eqvt)
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   485
(*A*)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   486
apply(rule exists_fresh')
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   487
apply(simp add: fin_supp)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   488
done
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   489
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   490
lemma Vs_are_values:
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   491
  assumes a: "e \<in> V T"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   492
  shows "val e"
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   493
using a by (nominal_induct T arbitrary: e rule: ty.induct) (auto)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   494
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   495
lemma values_reduce_to_themselves:
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   496
  assumes a: "val v"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   497
  shows "v \<Down> v"
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   498
using a by (induct) (auto)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   499
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   500
lemma Vs_reduce_to_themselves:
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   501
  assumes a: "v \<in> V T"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   502
  shows "v \<Down> v"
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   503
using a by (simp add: values_reduce_to_themselves Vs_are_values)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   504
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   505
text {* '\<theta> maps x to e' asserts that \<theta> substitutes x with e *}
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   506
abbreviation 
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   507
  mapsto :: "(name\<times>trm) list \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> bool" ("_ maps _ to _" [55,55,55] 55) 
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   508
where
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   509
 "\<theta> maps x to e \<equiv> (lookup \<theta> x) = e"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   510
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   511
abbreviation 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   512
  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
   513
where
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   514
  "\<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)"
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   515
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   516
lemma monotonicity:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   517
  fixes m::"name"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   518
  fixes \<theta>::"(name \<times> trm) list" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   519
  assumes h1: "\<theta> Vcloses \<Gamma>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   520
  and     h2: "e \<in> V T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   521
  and     h3: "valid ((x,T)#\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   522
  shows "(x,e)#\<theta> Vcloses (x,T)#\<Gamma>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   523
proof(intro strip)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   524
  fix x' T'
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   525
  assume "(x',T') \<in> set ((x,T)#\<Gamma>)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   526
  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
   527
    by (rule_tac case_distinction_on_context)
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
  { (* first case *)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   530
    assume "(x',T') = (x,T)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   531
    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
   532
  }
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   533
  moreover
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   534
  { (* second case *)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   535
    assume "(x',T') \<in> set \<Gamma>" and neq:"x' \<noteq> x"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   536
      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
   537
      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
   538
  }
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   539
  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
   540
qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   541
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   542
lemma termination_aux:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   543
  assumes h1: "\<Gamma> \<turnstile> e : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   544
  and     h2: "\<theta> Vcloses \<Gamma>"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   545
  shows "\<exists>v. \<theta><e> \<Down> v \<and> v \<in> V T" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   546
using h2 h1
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   547
proof(nominal_induct e avoiding: \<Gamma> \<theta> arbitrary: T rule: trm.induct)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   548
  case (App e\<^isub>1 e\<^isub>2 \<Gamma> \<theta> T)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   549
  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
   550
  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
   551
  have as\<^isub>1:"\<theta> Vcloses \<Gamma>" by fact 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   552
  have as\<^isub>2: "\<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T" by fact
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   553
  then obtain T' where "\<Gamma> \<turnstile> e\<^isub>1 : T' \<rightarrow> T" and "\<Gamma> \<turnstile> e\<^isub>2 : T'" 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   554
    by (cases) (auto simp add: trm.inject)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   555
  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
   556
                      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
   557
  from "(i)" obtain x e' 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   558
            where "v\<^isub>1 = Lam[x].e'" 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   559
            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
   560
            and "(iv)":  "\<theta><e\<^isub>1> \<Down> Lam [x].e'" 
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   561
            and fr: "x\<sharp>(\<theta>,e\<^isub>1,e\<^isub>2)" by (blast elim: V_arrow_elim_strong)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   562
  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
   563
  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
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   564
  from fr\<^isub>2 "(ii)" have "x\<sharp>v\<^isub>2" by (simp add: big_preserves_fresh)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   565
  then have "x\<sharp>e'[x::=v\<^isub>2]" by (simp add: fresh_subst')
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   566
  then have fr\<^isub>3: "x\<sharp>v\<^isub>3" using "(v)" by (simp add: big_preserves_fresh)
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   567
  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
   568
  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
   569
  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
   570
next
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   571
  case (Lam x e \<Gamma> \<theta> T)
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   572
  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
   573
  have as\<^isub>1: "\<theta> Vcloses \<Gamma>" by fact
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   574
  have as\<^isub>2: "\<Gamma> \<turnstile> Lam [x].e : T" by fact
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   575
  have fs: "x\<sharp>\<Gamma>" "x\<sharp>\<theta>" by fact
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   576
  from as\<^isub>2 fs obtain T\<^isub>1 T\<^isub>2 
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   577
    where "(i)": "(x,T\<^isub>1)#\<Gamma> \<turnstile> e:T\<^isub>2" and "(ii)": "T = T\<^isub>1 \<rightarrow> T\<^isub>2" using fs
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   578
    by (cases rule: typing.strong_cases[where x="x"])
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   579
       (auto simp add: trm.inject alpha abs_fresh fresh_ty)
22472
bfd9c0fd70b1 tuned the proof
urbanc
parents: 22447
diff changeset
   580
  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
   581
  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
   582
  proof
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   583
    fix v
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   584
    assume "v \<in> (V T\<^isub>1)"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   585
    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
   586
    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
   587
    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
   588
    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
   589
  qed
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   590
  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
   591
  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
   592
  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
   593
next
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   594
  case (Var x \<Gamma> \<theta> T)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   595
  have "\<Gamma> \<turnstile> (Var x) : T" by fact
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   596
  then have "(x,T)\<in>set \<Gamma>" by (cases) (auto simp add: trm.inject)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   597
  with prems have "\<theta><Var x> \<Down> \<theta><Var x> \<and> \<theta><Var x>\<in> V T" 
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   598
    by (auto intro!: Vs_reduce_to_themselves)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   599
  then show "\<exists>v. \<theta><Var x> \<Down> v \<and> v \<in> V T" by auto
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   600
qed 
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   601
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   602
theorem termination_of_evaluation:
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   603
  assumes a: "[] \<turnstile> e : T"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   604
  shows "\<exists>v. e \<Down> v \<and> val v"
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   605
proof -
25832
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   606
  from a have "\<exists>v. []<e> \<Down> v \<and> v \<in> V T" by (rule termination_aux) (auto)
41a014cc44c0 partially adapted to new inversion rules
urbanc
parents: 24678
diff changeset
   607
  thus "\<exists>v. e \<Down> v \<and> val v" using Vs_are_values by auto
22447
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   608
qed 
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   609
013dbd8234f0 added formalisations of typical SOS-proofs
urbanc
parents:
diff changeset
   610
end