src/HOL/Nominal/Examples/Crary.thy
author wenzelm
Fri, 19 Oct 2007 23:21:08 +0200
changeset 25107 dbf09ca6a80e
parent 24231 85fb973a8207
child 26647 147c920ed5f7
permissions -rw-r--r--
tuned proofs: avoid implicit prems;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
     1
(* "$Id$" *)
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
     2
(*                                                    *)
22082
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
     3
(* Formalisation of the chapter on Logical Relations  *)
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
     4
(* and a Case Study in Equivalence Checking           *)
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
     5
(* by Karl Crary from the book on Advanced Topics in  *)
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
     6
(* Types and Programming Languages, MIT Press 2005    *)
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
     7
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
     8
(* The formalisation was done by Julien Narboux and   *)
22594
33a690455f88 add a few details in the Fst and Snd cases of unicity proof
narboux
parents: 22542
diff changeset
     9
(* Christian Urban.                                   *)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    10
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    11
theory Crary
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    12
  imports "../Nominal"
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    13
begin
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    14
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    15
atom_decl name 
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    16
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22730
diff changeset
    17
nominal_datatype ty = 
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22730
diff changeset
    18
    TBase 
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22730
diff changeset
    19
  | TUnit 
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22730
diff changeset
    20
  | Arrow "ty" "ty" ("_\<rightarrow>_" [100,100] 100)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    21
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22730
diff changeset
    22
nominal_datatype trm = 
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22730
diff changeset
    23
    Unit
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
    24
  | Var "name" ("Var _" [100] 100)
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22730
diff changeset
    25
  | Lam "\<guillemotleft>name\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
    26
  | App "trm" "trm" ("App _ _" [110,110] 100)
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22730
diff changeset
    27
  | Const "nat"
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    28
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    29
types Ctxt  = "(name\<times>ty) list"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    30
types Subst = "(name\<times>trm) list" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    31
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    32
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    33
lemma perm_ty[simp]: 
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    34
  fixes T::"ty"
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    35
  and   pi::"name prm"
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    36
  shows "pi\<bullet>T = T"
22542
8279a25ad0ae - Renamed <predicate>_eqvt to <predicate>.eqvt
berghofe
parents: 22538
diff changeset
    37
  by (induct T rule: ty.weak_induct) (simp_all)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    38
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    39
lemma fresh_ty[simp]:
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    40
  fixes x::"name" 
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    41
  and   T::"ty"
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    42
  shows "x\<sharp>T"
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    43
  by (simp add: fresh_def supp_def)
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    44
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    45
lemma ty_cases:
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    46
  fixes T::ty
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    47
  shows "(\<exists> T\<^isub>1 T\<^isub>2. T=T\<^isub>1\<rightarrow>T\<^isub>2) \<or> T=TUnit \<or> T=TBase"
22542
8279a25ad0ae - Renamed <predicate>_eqvt to <predicate>.eqvt
berghofe
parents: 22538
diff changeset
    48
by (induct T rule:ty.weak_induct) (auto)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    49
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    50
instance ty :: size ..
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    51
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    52
nominal_primrec
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    53
  "size (TBase) = 1"
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    54
  "size (TUnit) = 1"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    55
  "size (T\<^isub>1\<rightarrow>T\<^isub>2) = size T\<^isub>1 + size T\<^isub>2"
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    56
by (rule TrueI)+
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    57
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    58
lemma ty_size_greater_zero[simp]:
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    59
  fixes T::"ty"
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    60
  shows "size T > 0"
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    61
by (nominal_induct rule:ty.induct) (simp_all)
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
    62
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    63
section {* Substitutions *}
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    64
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    65
fun
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    66
  lookup :: "Subst \<Rightarrow> name \<Rightarrow> trm"   
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    67
where
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    68
  "lookup [] x        = Var x"
22501
cb41c397a699 fixed function syntax
krauss
parents: 22494
diff changeset
    69
| "lookup ((y,T)#\<theta>) x = (if x=y then T else lookup \<theta> x)"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    70
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    71
lemma lookup_eqvt[eqvt]:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    72
  fixes pi::"name prm"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    73
  shows "pi\<bullet>(lookup \<theta> x) = lookup (pi\<bullet>\<theta>) (pi\<bullet>x)"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    74
by (induct \<theta>) (auto simp add: perm_bij)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    75
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    76
lemma lookup_fresh:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    77
  fixes z::"name"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    78
  assumes a: "z\<sharp>\<theta>" "z\<sharp>x"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    79
  shows "z\<sharp> lookup \<theta> x"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    80
using a
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    81
by (induct rule: lookup.induct) 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    82
   (auto simp add: fresh_list_cons)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    83
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    84
lemma lookup_fresh':
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    85
  assumes a: "z\<sharp>\<theta>"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    86
  shows "lookup \<theta> z = Var z"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    87
using a
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    88
by (induct rule: lookup.induct)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    89
   (auto simp add: fresh_list_cons fresh_prod fresh_atm)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    90
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    91
consts
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
    92
  psubst :: "Subst \<Rightarrow> trm \<Rightarrow> trm"  ("_<_>" [100,100] 130)
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    93
 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    94
nominal_primrec
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    95
  "\<theta><(Var x)> = (lookup \<theta> x)"
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
    96
  "\<theta><(App t\<^isub>1 t\<^isub>2)> = App \<theta><t\<^isub>1> \<theta><t\<^isub>2>"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    97
  "x\<sharp>\<theta> \<Longrightarrow> \<theta><(Lam [x].t)> = Lam [x].(\<theta><t>)"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    98
  "\<theta><(Const n)> = Const n"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
    99
  "\<theta><(Unit)> = Unit"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   100
apply(finite_guess)+
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   101
apply(rule TrueI)+
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   102
apply(simp add: abs_fresh)+
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   103
apply(fresh_guess)+
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   104
done
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   105
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   106
abbreviation 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   107
 subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   108
where
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   109
  "t[x::=t']  \<equiv> ([(x,t')])<t>" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   110
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   111
lemma subst[simp]:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   112
  shows "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   113
  and   "(App t\<^isub>1 t\<^isub>2)[y::=t'] = App (t\<^isub>1[y::=t']) (t\<^isub>2[y::=t'])"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   114
  and   "x\<sharp>(y,t') \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   115
  and   "Const n[y::=t'] = Const n"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   116
  and   "Unit [y::=t'] = Unit"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   117
  by (simp_all add: fresh_list_cons fresh_list_nil)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   118
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   119
lemma subst_eqvt[eqvt]:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   120
  fixes pi::"name prm" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   121
  shows "pi\<bullet>(t[x::=t']) = (pi\<bullet>t)[(pi\<bullet>x)::=(pi\<bullet>t')]"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   122
  by (nominal_induct t avoiding: x t' rule: trm.induct)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   123
     (perm_simp add: fresh_bij)+
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   124
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   125
lemma subst_rename: 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   126
  fixes c::"name"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   127
  assumes a: "c\<sharp>t\<^isub>1"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   128
  shows "t\<^isub>1[a::=t\<^isub>2] = ([(c,a)]\<bullet>t\<^isub>1)[c::=t\<^isub>2]"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   129
using a
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   130
apply(nominal_induct t\<^isub>1 avoiding: a c t\<^isub>2 rule: trm.induct)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   131
apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def)+
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   132
done
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   133
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   134
lemma fresh_psubst: 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   135
  fixes z::"name"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   136
  assumes a: "z\<sharp>t" "z\<sharp>\<theta>"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   137
  shows "z\<sharp>(\<theta><t>)"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   138
using a
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   139
by (nominal_induct t avoiding: z \<theta> t rule: trm.induct)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   140
   (auto simp add: abs_fresh lookup_fresh)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   141
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   142
lemma fresh_subst'':
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   143
  fixes z::"name"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   144
  assumes "z\<sharp>t\<^isub>2"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   145
  shows "z\<sharp>t\<^isub>1[z::=t\<^isub>2]"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   146
using assms 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   147
by (nominal_induct t\<^isub>1 avoiding: t\<^isub>2 z rule: trm.induct)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   148
   (auto simp add: abs_fresh fresh_nat fresh_atm)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   149
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   150
lemma fresh_subst':
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   151
  fixes z::"name"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   152
  assumes "z\<sharp>[y].t\<^isub>1" "z\<sharp>t\<^isub>2"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   153
  shows "z\<sharp>t\<^isub>1[y::=t\<^isub>2]"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   154
using assms 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   155
by (nominal_induct t\<^isub>1 avoiding: y t\<^isub>2 z rule: trm.induct)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   156
   (auto simp add: abs_fresh fresh_nat fresh_atm)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   157
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   158
lemma fresh_subst:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   159
  fixes z::"name"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   160
  assumes a: "z\<sharp>t\<^isub>1" "z\<sharp>t\<^isub>2"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   161
  shows "z\<sharp>t\<^isub>1[y::=t\<^isub>2]"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   162
using a 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   163
by (auto simp add: fresh_subst' abs_fresh) 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   164
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   165
lemma fresh_psubst_simp:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   166
  assumes "x\<sharp>t"
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   167
  shows "((x,u)#\<theta>)<t> = \<theta><t>" 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   168
using assms
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   169
proof (nominal_induct t avoiding: x u \<theta> rule: trm.induct)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   170
  case (Lam y t x u)
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
   171
  have fs: "y\<sharp>\<theta>" "y\<sharp>x" "y\<sharp>u" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   172
  moreover have "x\<sharp> Lam [y].t" by fact 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   173
  ultimately have "x\<sharp>t" by (simp add: abs_fresh fresh_atm)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   174
  moreover have ih:"\<And>n T. n\<sharp>t \<Longrightarrow> ((n,T)#\<theta>)<t> = \<theta><t>" by fact
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   175
  ultimately have "((x,u)#\<theta>)<t> = \<theta><t>" by auto
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   176
  moreover have "((x,u)#\<theta>)<Lam [y].t> = Lam [y].(((x,u)#\<theta>)<t>)" using fs 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   177
    by (simp add: fresh_list_cons fresh_prod)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   178
  moreover have " \<theta><Lam [y].t> = Lam [y]. (\<theta><t>)" using fs by simp
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   179
  ultimately show "((x,u)#\<theta>)<Lam [y].t> = \<theta><Lam [y].t>" by auto
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   180
qed (auto simp add: fresh_atm abs_fresh)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   181
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   182
lemma forget: 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   183
  fixes x::"name"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   184
  assumes a: "x\<sharp>t" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   185
  shows "t[x::=t'] = t"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   186
  using a
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   187
by (nominal_induct t avoiding: x t' rule: trm.induct)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   188
   (auto simp add: fresh_atm abs_fresh)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   189
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   190
lemma subst_fun_eq:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   191
  fixes u::trm
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   192
  assumes h:"[x].t\<^isub>1 = [y].t\<^isub>2"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   193
  shows "t\<^isub>1[x::=u] = t\<^isub>2[y::=u]"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   194
proof -
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   195
  { 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   196
    assume "x=y" and "t\<^isub>1=t\<^isub>2"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   197
    then have ?thesis using h by simp
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   198
  }
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   199
  moreover 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   200
  {
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   201
    assume h1:"x \<noteq> y" and h2:"t\<^isub>1=[(x,y)] \<bullet> t\<^isub>2" and h3:"x \<sharp> t\<^isub>2"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   202
    then have "([(x,y)] \<bullet> t\<^isub>2)[x::=u] = t\<^isub>2[y::=u]" by (simp add: subst_rename)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   203
    then have ?thesis using h2 by simp 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   204
  }
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   205
  ultimately show ?thesis using alpha h by blast
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   206
qed
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   207
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   208
lemma psubst_empty[simp]:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   209
  shows "[]<t> = t"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   210
by (nominal_induct t rule: trm.induct) 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   211
   (auto simp add: fresh_list_nil)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   212
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   213
lemma psubst_subst_psubst:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   214
  assumes h:"c\<sharp>\<theta>"
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   215
  shows "\<theta><t>[c::=s] = ((c,s)#\<theta>)<t>"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   216
  using h
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   217
by (nominal_induct t avoiding: \<theta> c s rule: trm.induct)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   218
   (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh' fresh_psubst)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   219
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   220
lemma subst_fresh_simp:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   221
  assumes a: "x\<sharp>\<theta>"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   222
  shows "\<theta><Var x> = Var x"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   223
using a
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   224
by (induct \<theta> arbitrary: x, auto simp add:fresh_list_cons fresh_prod fresh_atm)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   225
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   226
lemma psubst_subst_propagate: 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   227
  assumes "x\<sharp>\<theta>" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   228
  shows "\<theta><t[x::=u]> = \<theta><t>[x::=\<theta><u>]"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   229
using assms
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   230
proof (nominal_induct t avoiding: x u \<theta> rule: trm.induct)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   231
  case (Var n x u \<theta>)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   232
  { assume "x=n"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   233
    moreover have "x\<sharp>\<theta>" by fact 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   234
    ultimately have "\<theta><Var n[x::=u]> = \<theta><Var n>[x::=\<theta><u>]" using subst_fresh_simp by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   235
  }
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   236
  moreover 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   237
  { assume h:"x\<noteq>n"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   238
    then have "x\<sharp>Var n" by (auto simp add: fresh_atm) 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   239
    moreover have "x\<sharp>\<theta>" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   240
    ultimately have "x\<sharp>\<theta><Var n>" using fresh_psubst by blast
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   241
    then have " \<theta><Var n>[x::=\<theta><u>] =  \<theta><Var n>" using forget by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   242
    then have "\<theta><Var n[x::=u]> = \<theta><Var n>[x::=\<theta><u>]" using h by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   243
  }
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   244
  ultimately show ?case by auto  
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   245
next
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   246
  case (Lam n t x u \<theta>)
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
   247
  have fs:"n\<sharp>x" "n\<sharp>u" "n\<sharp>\<theta>" "x\<sharp>\<theta>" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   248
  have ih:"\<And> y s \<theta>. y\<sharp>\<theta> \<Longrightarrow> ((\<theta><(t[y::=s])>) = ((\<theta><t>)[y::=(\<theta><s>)]))" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   249
  have "\<theta> <(Lam [n].t)[x::=u]> = \<theta><Lam [n]. (t [x::=u])>" using fs by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   250
  then have "\<theta> <(Lam [n].t)[x::=u]> = Lam [n]. \<theta><t [x::=u]>" using fs by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   251
  moreover have "\<theta><t[x::=u]> = \<theta><t>[x::=\<theta><u>]" using ih fs by blast
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   252
  ultimately have "\<theta> <(Lam [n].t)[x::=u]> = Lam [n].(\<theta><t>[x::=\<theta><u>])" by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   253
  moreover have "Lam [n].(\<theta><t>[x::=\<theta><u>]) = (Lam [n].\<theta><t>)[x::=\<theta><u>]" using fs fresh_psubst by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   254
  ultimately have "\<theta><(Lam [n].t)[x::=u]> = (Lam [n].\<theta><t>)[x::=\<theta><u>]" using fs by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   255
  then show "\<theta><(Lam [n].t)[x::=u]> = \<theta><Lam [n].t>[x::=\<theta><u>]" using fs by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   256
qed (auto)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   257
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   258
section {* Typing *}
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   259
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23450
diff changeset
   260
inductive
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   261
  valid :: "Ctxt \<Rightarrow> bool"
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   262
where
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   263
  v_nil[intro]:  "valid []"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   264
| v_cons[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk> \<Longrightarrow> valid ((a,T)#\<Gamma>)"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   265
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   266
equivariance valid 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   267
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   268
inductive_cases
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   269
  valid_cons_elim_auto[elim]:"valid ((x,T)#\<Gamma>)"
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   270
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   271
abbreviation
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: 22609
diff changeset
   272
  "sub_context" :: "Ctxt \<Rightarrow> Ctxt \<Rightarrow> bool" (" _ \<subseteq> _ " [55,55] 55)
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   273
where
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: 22609
diff changeset
   274
  "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2 \<equiv> \<forall>a T. (a,T)\<in>set \<Gamma>\<^isub>1 \<longrightarrow> (a,T)\<in>set \<Gamma>\<^isub>2"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   275
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   276
lemma valid_monotonicity[elim]:
24231
85fb973a8207 added type constraints to resolve syntax ambiguities;
wenzelm
parents: 24088
diff changeset
   277
 fixes \<Gamma> \<Gamma>' :: Ctxt
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: 22609
diff changeset
   278
 assumes a: "\<Gamma> \<subseteq> \<Gamma>'" 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   279
 and     b: "x\<sharp>\<Gamma>'"
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: 22609
diff changeset
   280
 shows "(x,T\<^isub>1)#\<Gamma> \<subseteq> (x,T\<^isub>1)#\<Gamma>'"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   281
using a b by auto
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   282
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   283
lemma fresh_context: 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   284
  fixes  \<Gamma> :: "Ctxt"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   285
  and    a :: "name"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   286
  assumes "a\<sharp>\<Gamma>"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   287
  shows "\<not>(\<exists>\<tau>::ty. (a,\<tau>)\<in>set \<Gamma>)"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   288
using assms 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   289
by (induct \<Gamma>)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   290
   (auto simp add: fresh_prod fresh_list_cons fresh_atm)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   291
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   292
lemma type_unicity_in_context:
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   293
  assumes a: "valid \<Gamma>" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   294
  and     b: "(x,T\<^isub>1) \<in> set \<Gamma>" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   295
  and     c: "(x,T\<^isub>2) \<in> set \<Gamma>"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   296
  shows "T\<^isub>1=T\<^isub>2"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   297
using a b c
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   298
by (induct \<Gamma>)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   299
   (auto dest!: fresh_context)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   300
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23450
diff changeset
   301
inductive
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   302
  typing :: "Ctxt\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [60,60,60] 60) 
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   303
where
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   304
  T_Var[intro]:   "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   305
| 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"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   306
| T_Lam[intro]:   "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T\<^isub>1\<rightarrow>T\<^isub>2"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   307
| T_Const[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Const n : TBase"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   308
| T_Unit[intro]:  "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Unit : TUnit"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   309
22730
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
   310
equivariance typing
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
   311
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   312
nominal_inductive typing
22531
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22501
diff changeset
   313
  by (simp_all add: abs_fresh)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   314
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   315
lemma typing_implies_valid:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   316
  assumes a: "\<Gamma> \<turnstile> t : T"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   317
  shows "valid \<Gamma>"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   318
  using a by (induct) (auto)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   319
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   320
declare trm.inject [simp add]
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   321
declare ty.inject  [simp add]
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   322
24088
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   323
inductive_cases typing_inv_auto[elim]:
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   324
  "\<Gamma> \<turnstile> Lam [x].t : T"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   325
  "\<Gamma> \<turnstile> Var x : T"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   326
  "\<Gamma> \<turnstile> App x y : T"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   327
  "\<Gamma> \<turnstile> Const n : T"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   328
  "\<Gamma> \<turnstile> Unit : TUnit"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   329
  "\<Gamma> \<turnstile> s : TUnit"
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   330
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   331
declare trm.inject [simp del]
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   332
declare ty.inject [simp del]
24088
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   333
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   334
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   335
section {* Definitional Equivalence *}
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   336
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23450
diff changeset
   337
inductive
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   338
  def_equiv :: "Ctxt\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<equiv> _ : _" [60,60] 60) 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   339
where
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   340
  Q_Refl[intro]:  "\<Gamma> \<turnstile> t : T \<Longrightarrow> \<Gamma> \<turnstile> t \<equiv> t : T"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   341
| Q_Symm[intro]:  "\<Gamma> \<turnstile> t \<equiv> s : T \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   342
| Q_Trans[intro]: "\<lbrakk>\<Gamma> \<turnstile> s \<equiv> t : T; \<Gamma> \<turnstile> t \<equiv> u : T\<rbrakk> \<Longrightarrow>  \<Gamma> \<turnstile> s \<equiv> u : T"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   343
| Q_Abs[intro]:   "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^isub>1)#\<Gamma> \<turnstile> s\<^isub>2 \<equiv> t\<^isub>2 : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x]. s\<^isub>2 \<equiv>  Lam [x]. t\<^isub>2 : T\<^isub>1 \<rightarrow> T\<^isub>2"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   344
| Q_App[intro]:   "\<lbrakk>\<Gamma> \<turnstile> s\<^isub>1 \<equiv> t\<^isub>1 : T\<^isub>1 \<rightarrow> T\<^isub>2 ; \<Gamma> \<turnstile> s\<^isub>2 \<equiv> t\<^isub>2 : T\<^isub>1\<rbrakk> \<Longrightarrow>  \<Gamma> \<turnstile> App s\<^isub>1 s\<^isub>2 \<equiv> App t\<^isub>1 t\<^isub>2 : T\<^isub>2"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   345
| Q_Beta[intro]:  "\<lbrakk>x\<sharp>(\<Gamma>,s\<^isub>2,t\<^isub>2); (x,T\<^isub>1)#\<Gamma> \<turnstile> s\<^isub>1 \<equiv> t\<^isub>1 : T\<^isub>2 ; \<Gamma> \<turnstile> s\<^isub>2 \<equiv> t\<^isub>2 : T\<^isub>1\<rbrakk> 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   346
                   \<Longrightarrow>  \<Gamma> \<turnstile> App (Lam [x]. s\<^isub>1) s\<^isub>2 \<equiv> t\<^isub>1[x::=t\<^isub>2] : T\<^isub>2"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   347
| Q_Ext[intro]:   "\<lbrakk>x\<sharp>(\<Gamma>,s,t); (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<equiv> App t (Var x) : T\<^isub>2\<rbrakk> 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   348
                   \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T\<^isub>1 \<rightarrow> T\<^isub>2"
23370
513a8ee192f1 added the Q_Unit rule (was missing) and adjusted the proof accordingly
urbanc
parents: 22829
diff changeset
   349
| Q_Unit[intro]:  "\<lbrakk>\<Gamma> \<turnstile> s : TUnit; \<Gamma> \<turnstile> t: TUnit\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : TUnit"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   350
22730
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
   351
equivariance def_equiv
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
   352
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   353
nominal_inductive def_equiv
22531
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22501
diff changeset
   354
  by (simp_all add: abs_fresh fresh_subst'')
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   355
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   356
lemma def_equiv_implies_valid:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   357
  assumes a: "\<Gamma> \<turnstile> t \<equiv> s : T"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   358
  shows "valid \<Gamma>"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   359
using a by (induct) (auto elim: typing_implies_valid)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   360
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   361
lemma test30:
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   362
  fixes x::"name"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   363
  assumes a: "(x,T) \<in> set \<Gamma>"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   364
  shows "x\<in>supp \<Gamma>"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   365
using a
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   366
apply(induct \<Gamma>)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   367
apply(auto simp add: supp_list_cons supp_prod supp_atm)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   368
done
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   369
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   370
lemma supp_ty[simp]: 
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   371
  fixes T::"ty"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   372
  shows "(supp T) = ({}::name set)"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   373
apply(simp add: supp_def)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   374
done
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   375
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   376
lemma test3a:
24231
85fb973a8207 added type constraints to resolve syntax ambiguities;
wenzelm
parents: 24088
diff changeset
   377
  fixes \<Gamma> :: Ctxt and t :: trm
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   378
  assumes a: "\<Gamma> \<turnstile> t : T"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   379
  shows "(supp t) \<subseteq> ((supp \<Gamma>)::name set)"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   380
using a
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   381
apply(induct)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   382
apply(auto simp add: trm.supp supp_atm supp_list_cons abs_supp supp_prod)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   383
apply(simp add: test30)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   384
apply(simp add: supp_def perm_nat_def)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   385
done
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   386
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   387
lemma test3b:
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   388
  shows "supp (t\<^isub>1[x::=t\<^isub>2]) \<subseteq> ((supp ([x].t\<^isub>1,t\<^isub>2))::name set)"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   389
apply(nominal_induct t\<^isub>1 avoiding: x t\<^isub>2 rule: trm.induct)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   390
apply(auto simp add: trm.supp supp_prod abs_supp supp_atm)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   391
apply(blast)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   392
apply(blast)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   393
apply(simp add: supp_def perm_nat_def)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   394
done
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   395
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   396
lemma test3:
24231
85fb973a8207 added type constraints to resolve syntax ambiguities;
wenzelm
parents: 24088
diff changeset
   397
  fixes \<Gamma> :: Ctxt and s t :: trm
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   398
  assumes a: "\<Gamma> \<turnstile> s \<equiv> t : T"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   399
  shows "(supp (s,t)) \<subseteq> ((supp \<Gamma>)::name set)"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   400
using a
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   401
apply(induct)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   402
apply(auto simp add: supp_prod  supp_list_cons trm.supp abs_supp supp_atm)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   403
apply(drule test3a)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   404
apply(blast)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   405
apply(subgoal_tac "supp (t\<^isub>1[x::=t\<^isub>2]) \<subseteq> ((supp ([x].t\<^isub>1,t\<^isub>2))::name set)")
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   406
apply(auto simp add: supp_prod abs_supp)[1]
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   407
apply(rule test3b)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   408
apply(case_tac "x=xa")
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   409
apply(simp add: fresh_def supp_prod)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   410
apply(blast)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   411
apply(case_tac "x=xa")
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   412
apply(simp add: fresh_def supp_prod)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   413
apply(blast)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   414
apply(drule test3a)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   415
apply(blast)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   416
apply(drule test3a)+
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   417
apply(blast)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   418
done
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   419
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   420
lemma test0:
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   421
  assumes a: "(x,T\<^isub>1)#\<Gamma> \<turnstile> s\<^isub>1 \<equiv> t\<^isub>1 : T\<^isub>2" 
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   422
  and     b: "\<Gamma> \<turnstile> s\<^isub>2 \<equiv> t\<^isub>2 : T\<^isub>1"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   423
  shows "\<Gamma> \<turnstile> App (Lam [x]. s\<^isub>1) s\<^isub>2 \<equiv> t\<^isub>1[x::=t\<^isub>2] : T\<^isub>2"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   424
using a b
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   425
apply(rule_tac Q_Beta)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   426
apply(auto)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   427
apply(drule def_equiv_implies_valid)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   428
apply(drule valid.cases)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   429
apply(auto)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   430
apply(drule test3)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   431
apply(auto simp add: fresh_def supp_prod)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   432
done
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   433
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   434
lemma test1:
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   435
  assumes a: "\<forall>x. x\<sharp>\<Gamma> \<longrightarrow> (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<equiv> App t (Var x) : T\<^isub>2"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   436
  shows "\<Gamma> \<turnstile> s \<equiv> t : T\<^isub>1 \<rightarrow> T\<^isub>2"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   437
using a
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   438
apply -
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   439
apply(generate_fresh "name")
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   440
apply(rule_tac x="c" in Q_Ext)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   441
apply(simp)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   442
apply(simp add: fresh_prod)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   443
done
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   444
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   445
lemma test2:
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   446
  assumes a: "x\<sharp>(\<Gamma>,s,t) \<and> (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<equiv> App t (Var x) : T\<^isub>2"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   447
  shows "\<forall>x. x\<sharp>(\<Gamma>,s,t) \<longrightarrow> (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<equiv> App t (Var x) : T\<^isub>2"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   448
using a
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   449
apply -
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   450
apply(rule allI)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   451
apply(case_tac "xa=x")
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   452
apply(simp)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   453
apply(rule impI)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   454
apply(erule conjE)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   455
apply(drule_tac pi="[(x,xa)]" in def_equiv.eqvt)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   456
apply(simp add: eqvts)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   457
apply(simp add: calc_atm)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   458
apply(perm_simp)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   459
done
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   460
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   461
lemma test2:
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   462
  assumes a: "x\<sharp>(\<Gamma>,s,t) \<and> (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<equiv> App t (Var x) : T\<^isub>2"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   463
  shows "\<forall>x. x\<sharp>\<Gamma> \<longrightarrow> (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<equiv> App t (Var x) : T\<^isub>2"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   464
using a
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   465
apply -
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   466
apply(rule allI)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   467
apply(case_tac "xa=x")
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   468
apply(simp)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   469
apply(rule impI)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   470
apply(erule conjE)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   471
apply(frule test3)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   472
apply(simp add: supp_prod supp_list_cons supp_atm trm.supp)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   473
apply(subgoal_tac "xa\<sharp>(s,t)")
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   474
apply(drule_tac pi="[(x,xa)]" in def_equiv.eqvt)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   475
apply(simp add: eqvts)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   476
apply(simp add: calc_atm)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   477
apply(perm_simp)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   478
apply(simp add: fresh_def supp_prod)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   479
apply(auto)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   480
done
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   481
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   482
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   483
section {* Weak Head Reduction *}
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   484
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23450
diff changeset
   485
inductive
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   486
  whr_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<leadsto> _" [80,80] 80) 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   487
where
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   488
  QAR_Beta[intro]: "App (Lam [x]. t\<^isub>1) t\<^isub>2 \<leadsto> t\<^isub>1[x::=t\<^isub>2]"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   489
| QAR_App[intro]:  "t\<^isub>1 \<leadsto> t\<^isub>1' \<Longrightarrow> App t\<^isub>1 t\<^isub>2 \<leadsto> App t\<^isub>1' t\<^isub>2"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   490
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   491
declare trm.inject  [simp add]
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   492
declare ty.inject  [simp add]
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   493
24088
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   494
inductive_cases whr_inv_auto[elim]: 
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   495
  "t \<leadsto> t'"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   496
  "Lam [x].t \<leadsto> t'"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   497
  "App (Lam [x].t12) t2 \<leadsto> t"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   498
  "Var x \<leadsto> t"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   499
  "Const n \<leadsto> t"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   500
  "App p q \<leadsto> t"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   501
  "t \<leadsto> Const n"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   502
  "t \<leadsto> Var x"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   503
  "t \<leadsto> App p q"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   504
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   505
declare trm.inject  [simp del]
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   506
declare ty.inject  [simp del]
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   507
22531
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22501
diff changeset
   508
equivariance whr_def
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   509
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   510
section {* Weak Head Normalisation *}
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   511
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   512
abbreviation 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   513
 nf :: "trm \<Rightarrow> bool" ("_ \<leadsto>|" [100] 100)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   514
where
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   515
  "t\<leadsto>|  \<equiv> \<not>(\<exists> u. t \<leadsto> u)" 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   516
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23450
diff changeset
   517
inductive
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   518
  whn_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80) 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   519
where
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   520
  QAN_Reduce[intro]: "\<lbrakk>s \<leadsto> t; t \<Down> u\<rbrakk> \<Longrightarrow> s \<Down> u"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   521
| QAN_Normal[intro]: "t\<leadsto>|  \<Longrightarrow> t \<Down> t"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   522
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   523
declare trm.inject[simp]
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   524
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23450
diff changeset
   525
inductive_cases whn_inv_auto[elim]: "t \<Down> t'"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   526
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   527
declare trm.inject[simp del]
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   528
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   529
lemma whn_eqvt[eqvt]:
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   530
  fixes pi::"name prm"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   531
  assumes a: "t \<Down> t'"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   532
  shows "(pi\<bullet>t) \<Down> (pi\<bullet>t')"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   533
using a
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   534
apply(induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   535
apply(rule QAN_Reduce)
22542
8279a25ad0ae - Renamed <predicate>_eqvt to <predicate>.eqvt
berghofe
parents: 22538
diff changeset
   536
apply(rule whr_def.eqvt)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   537
apply(assumption)+
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   538
apply(rule QAN_Normal)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   539
apply(auto)
22542
8279a25ad0ae - Renamed <predicate>_eqvt to <predicate>.eqvt
berghofe
parents: 22538
diff changeset
   540
apply(drule_tac pi="rev pi" in whr_def.eqvt)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   541
apply(perm_simp)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   542
done
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   543
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   544
lemma red_unicity : 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   545
  assumes a: "x \<leadsto> a" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   546
  and     b: "x \<leadsto> b"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   547
  shows "a=b"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   548
  using a b
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   549
apply (induct arbitrary: b)
24088
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   550
apply (erule whr_inv_auto(3))
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   551
apply (clarify)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   552
apply (rule subst_fun_eq)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   553
apply (simp)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   554
apply (force)
24088
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   555
apply (erule whr_inv_auto(6))
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   556
apply (blast)+
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   557
done
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   558
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   559
lemma nf_unicity :
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   560
  assumes "x \<Down> a" and "x \<Down> b"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   561
  shows "a=b"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   562
  using assms 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   563
proof (induct arbitrary: b)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   564
  case (QAN_Reduce x t a b)
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
   565
  have h:"x \<leadsto> t" "t \<Down> a" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   566
  have ih:"\<And>b. t \<Down> b \<Longrightarrow> a = b" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   567
  have "x \<Down> b" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   568
  then obtain t' where "x \<leadsto> t'" and hl:"t' \<Down> b" using h by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   569
  then have "t=t'" using h red_unicity by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   570
  then show "a=b" using ih hl by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   571
qed (auto)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   572
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   573
lemma test4a:
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   574
  shows "s \<leadsto> t \<Longrightarrow> (supp t) \<subseteq> ((supp s)::name set)"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   575
apply(induct s t rule: whr_def.induct)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   576
apply(insert test3b)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   577
apply(simp add: trm.supp supp_prod abs_supp)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   578
apply(auto simp add: trm.supp)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   579
done
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   580
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   581
lemma test4b:
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   582
  shows "s \<Down> t \<Longrightarrow> (supp t) \<subseteq> ((supp s)::name set)"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   583
apply(induct s t rule: whn_def.induct)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   584
apply(auto dest: test4a)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   585
done
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   586
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   587
section {* Algorithmic Term Equivalence and Algorithmic Path Equivalence *}
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   588
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23450
diff changeset
   589
inductive
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: 22609
diff changeset
   590
  alg_equiv :: "Ctxt\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<Leftrightarrow> _ : _" [60,60,60,60] 60) 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   591
and
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: 22609
diff changeset
   592
  alg_path_equiv :: "Ctxt\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<leftrightarrow> _ : _" [60,60,60,60] 60) 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   593
where
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: 22609
diff changeset
   594
  QAT_Base[intro]:  "\<lbrakk>s \<Down> p; t \<Down> q; \<Gamma> \<turnstile> p \<leftrightarrow> q : TBase\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : TBase"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   595
| QAT_Arrow[intro]: "\<lbrakk>x\<sharp>(\<Gamma>,s,t); (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2\<rbrakk> 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   596
                     \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1 \<rightarrow> T\<^isub>2"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   597
| QAT_One[intro]:   "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : TUnit"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   598
| QAP_Var[intro]:   "\<lbrakk>valid \<Gamma>;(x,T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x \<leftrightarrow> Var x : T"
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: 22609
diff changeset
   599
| QAP_App[intro]:   "\<lbrakk>\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^isub>1 \<rightarrow> T\<^isub>2; \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App p s \<leftrightarrow> App q t : T\<^isub>2"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   600
| QAP_Const[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Const n \<leftrightarrow> Const n : TBase"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   601
22730
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
   602
equivariance alg_equiv
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
   603
22531
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22501
diff changeset
   604
nominal_inductive alg_equiv
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22501
diff changeset
   605
  avoids QAT_Arrow: x
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22501
diff changeset
   606
  by simp_all
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   607
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   608
declare trm.inject [simp add]
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   609
declare ty.inject  [simp add]
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   610
24088
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   611
inductive_cases alg_equiv_inv_auto[elim]: 
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   612
  "\<Gamma> \<turnstile> s\<Leftrightarrow>t : TBase"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   613
  "\<Gamma> \<turnstile> s\<Leftrightarrow>t : T\<^isub>1 \<rightarrow> T\<^isub>2"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   614
  "\<Gamma> \<turnstile> s\<leftrightarrow>t : TBase"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   615
  "\<Gamma> \<turnstile> s\<leftrightarrow>t : TUnit"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   616
  "\<Gamma> \<turnstile> s\<leftrightarrow>t : T\<^isub>1 \<rightarrow> T\<^isub>2"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   617
24088
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   618
  "\<Gamma> \<turnstile> Var x \<leftrightarrow> t : T"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   619
  "\<Gamma> \<turnstile> Var x \<leftrightarrow> t : T'"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   620
  "\<Gamma> \<turnstile> s \<leftrightarrow> Var x : T"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   621
  "\<Gamma> \<turnstile> s \<leftrightarrow> Var x : T'"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   622
  "\<Gamma> \<turnstile> Const n \<leftrightarrow> t : T"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   623
  "\<Gamma> \<turnstile> s \<leftrightarrow> Const n : T"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   624
  "\<Gamma> \<turnstile> App p s \<leftrightarrow> t : T"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   625
  "\<Gamma> \<turnstile> s \<leftrightarrow> App q t : T"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   626
  "\<Gamma> \<turnstile> Lam[x].s \<leftrightarrow> t : T"
c2d8270e53a5 undo a change in last commit : give a single name to the inversion lemmas for the same inductive type
narboux
parents: 24070
diff changeset
   627
  "\<Gamma> \<turnstile> t \<leftrightarrow> Lam[x].s : T"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   628
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   629
declare trm.inject [simp del]
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   630
declare ty.inject [simp del]
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   631
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   632
lemma Q_Arrow_strong_inversion:
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   633
  assumes fs: "x\<sharp>\<Gamma>" "x\<sharp>t" "x\<sharp>u" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   634
  and h: "\<Gamma> \<turnstile> t \<Leftrightarrow> u : T\<^isub>1\<rightarrow>T\<^isub>2"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   635
  shows "(x,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2"
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   636
proof -
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   637
  obtain y where fs2: "y\<sharp>(\<Gamma>,t,u)" and "(y,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var y) \<Leftrightarrow> App u (Var y) : T\<^isub>2" 
22082
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
   638
    using h by auto
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   639
  then have "([(x,y)]\<bullet>((y,T\<^isub>1)#\<Gamma>)) \<turnstile> [(x,y)]\<bullet> App t (Var y) \<Leftrightarrow> [(x,y)]\<bullet> App u (Var y) : T\<^isub>2" 
22542
8279a25ad0ae - Renamed <predicate>_eqvt to <predicate>.eqvt
berghofe
parents: 22538
diff changeset
   640
    using  alg_equiv.eqvt[simplified] by blast
22082
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
   641
  then show ?thesis using fs fs2 by (perm_simp)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   642
qed
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   643
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   644
(*
22594
33a690455f88 add a few details in the Fst and Snd cases of unicity proof
narboux
parents: 22542
diff changeset
   645
Warning this lemma is false:
33a690455f88 add a few details in the Fst and Snd cases of unicity proof
narboux
parents: 22542
diff changeset
   646
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   647
lemma algorithmic_type_unicity:
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   648
  shows "\<lbrakk> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T ; \<Gamma> \<turnstile> s \<Leftrightarrow> u : T' \<rbrakk> \<Longrightarrow> T = T'"
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   649
  and "\<lbrakk> \<Gamma> \<turnstile> s \<leftrightarrow> t : T ; \<Gamma> \<turnstile> s \<leftrightarrow> u : T' \<rbrakk> \<Longrightarrow> T = T'"
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   650
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   651
Here is the counter example : 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   652
\<Gamma> \<turnstile> Const n \<Leftrightarrow> Const n : Tbase and \<Gamma> \<turnstile> Const n \<Leftrightarrow> Const n : TUnit
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   653
*)
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   654
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   655
lemma algorithmic_path_type_unicity:
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   656
  shows "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<turnstile> s \<leftrightarrow> u : T' \<Longrightarrow> T = T'"
22082
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
   657
proof (induct arbitrary:  u T' 
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
   658
       rule: alg_equiv_alg_path_equiv.inducts(2) [of _ _ _ _ _  "%a b c d . True"    ])
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   659
  case (QAP_Var \<Gamma> x T u T')
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   660
  have "\<Gamma> \<turnstile> Var x \<leftrightarrow> u : T'" by fact
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   661
  then have "u=Var x" and "(x,T') \<in> set \<Gamma>" by auto
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
   662
  moreover have "valid \<Gamma>" "(x,T) \<in> set \<Gamma>" by fact+
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   663
  ultimately show "T=T'" using type_unicity_in_context by auto
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   664
next
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   665
  case (QAP_App \<Gamma> p q T\<^isub>1 T\<^isub>2 s t u T\<^isub>2')
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   666
  have ih:"\<And>u T. \<Gamma> \<turnstile> p \<leftrightarrow> u : T \<Longrightarrow> T\<^isub>1\<rightarrow>T\<^isub>2 = T" by fact
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   667
  have "\<Gamma> \<turnstile> App p s \<leftrightarrow> u : T\<^isub>2'" by fact
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   668
  then obtain r t T\<^isub>1' where "u = App r t"  "\<Gamma> \<turnstile> p \<leftrightarrow> r : T\<^isub>1' \<rightarrow> T\<^isub>2'" by auto
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
   669
  with ih have "T\<^isub>1\<rightarrow>T\<^isub>2 = T\<^isub>1' \<rightarrow> T\<^isub>2'" by auto
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   670
  then show "T\<^isub>2=T\<^isub>2'" using ty.inject by auto
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   671
qed (auto)
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   672
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   673
lemma alg_path_equiv_implies_valid:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   674
  shows  "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<Longrightarrow> valid \<Gamma>" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   675
  and    "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> valid \<Gamma>"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   676
by (induct rule : alg_equiv_alg_path_equiv.inducts, auto)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   677
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   678
lemma algorithmic_symmetry:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   679
  shows "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<turnstile> t \<Leftrightarrow> s : T" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   680
  and   "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<turnstile> t \<leftrightarrow> s : T"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   681
by (induct rule: alg_equiv_alg_path_equiv.inducts) 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   682
   (auto simp add: fresh_prod)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   683
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   684
lemma algorithmic_transitivity:
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   685
  shows "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<turnstile> t \<Leftrightarrow> u : T \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> u : T"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   686
  and   "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<turnstile> t \<leftrightarrow> u : T \<Longrightarrow> \<Gamma> \<turnstile> s \<leftrightarrow> u : T"
22531
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22501
diff changeset
   687
proof (nominal_induct \<Gamma> s t T and \<Gamma> s t T avoiding: u rule: alg_equiv_alg_path_equiv.strong_inducts)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   688
  case (QAT_Base s p t q \<Gamma> u)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   689
  have "\<Gamma> \<turnstile> t \<Leftrightarrow> u : TBase" by fact
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   690
  then obtain r' q' where b1: "t \<Down> q'" and b2: "u \<Down> r'" and b3: "\<Gamma> \<turnstile> q' \<leftrightarrow> r' : TBase" by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   691
  have ih: "\<Gamma> \<turnstile> q \<leftrightarrow> r' : TBase \<Longrightarrow> \<Gamma> \<turnstile> p \<leftrightarrow> r' : TBase" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   692
  have "t \<Down> q" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   693
  with b1 have eq: "q=q'" by (simp add: nf_unicity)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   694
  with ih b3 have "\<Gamma> \<turnstile> p \<leftrightarrow> r' : TBase" by simp
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   695
  moreover
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   696
  have "s \<Down> p" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   697
  ultimately show "\<Gamma> \<turnstile> s \<Leftrightarrow> u : TBase" using b2 by auto
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   698
next
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   699
  case (QAT_Arrow  x \<Gamma> s t T\<^isub>1 T\<^isub>2 u)
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   700
  have ih:"(x,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   701
                                   \<Longrightarrow> (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2" by fact
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
   702
  have fs: "x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" "x\<sharp>u" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   703
  have "\<Gamma> \<turnstile> t \<Leftrightarrow> u : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   704
  then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2" using fs 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   705
    by (simp add: Q_Arrow_strong_inversion)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   706
  with ih have "(x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2" by simp
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   707
  then show "\<Gamma> \<turnstile> s \<Leftrightarrow> u : T\<^isub>1\<rightarrow>T\<^isub>2" using fs by (auto simp add: fresh_prod)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   708
next
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   709
  case (QAP_App \<Gamma> p q T\<^isub>1 T\<^isub>2 s t u)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   710
  have "\<Gamma> \<turnstile> App q t \<leftrightarrow> u : T\<^isub>2" by fact
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   711
  then obtain r T\<^isub>1' v where ha: "\<Gamma> \<turnstile> q \<leftrightarrow> r : T\<^isub>1'\<rightarrow>T\<^isub>2" and hb: "\<Gamma> \<turnstile> t \<Leftrightarrow> v : T\<^isub>1'" and eq: "u = App r v" 
22082
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
   712
    by auto
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   713
  have ih1: "\<Gamma> \<turnstile> q \<leftrightarrow> r : T\<^isub>1\<rightarrow>T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> p \<leftrightarrow> r : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   714
  have ih2:"\<Gamma> \<turnstile> t \<Leftrightarrow> v : T\<^isub>1 \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> v : T\<^isub>1" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   715
  have "\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   716
  then have "\<Gamma> \<turnstile> q \<leftrightarrow> p : T\<^isub>1\<rightarrow>T\<^isub>2" by (simp add: algorithmic_symmetry)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   717
  with ha have "T\<^isub>1'\<rightarrow>T\<^isub>2 = T\<^isub>1\<rightarrow>T\<^isub>2" using algorithmic_path_type_unicity by simp
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   718
  then have "T\<^isub>1' = T\<^isub>1" by (simp add: ty.inject) 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   719
  then have "\<Gamma> \<turnstile> s \<Leftrightarrow> v : T\<^isub>1" "\<Gamma> \<turnstile> p \<leftrightarrow> r : T\<^isub>1\<rightarrow>T\<^isub>2" using ih1 ih2 ha hb by auto
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   720
  then show "\<Gamma> \<turnstile> App p s \<leftrightarrow> u : T\<^isub>2" using eq by auto
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   721
qed (auto)
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   722
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   723
lemma algorithmic_weak_head_closure:
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   724
  shows "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<Longrightarrow> s' \<leadsto> s \<Longrightarrow> t' \<leadsto> t \<Longrightarrow> \<Gamma> \<turnstile> s' \<Leftrightarrow> t' : T"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   725
apply (nominal_induct \<Gamma> s t T avoiding: s' t'  
22531
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22501
diff changeset
   726
    rule: alg_equiv_alg_path_equiv.strong_inducts(1) [of _ _ _ _ "%a b c d e. True"])
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   727
apply(auto intro!: QAT_Arrow)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   728
done
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   729
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   730
lemma algorithmic_monotonicity:
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: 22609
diff changeset
   731
  shows "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<subseteq> \<Gamma>' \<Longrightarrow> valid \<Gamma>' \<Longrightarrow> \<Gamma>' \<turnstile> s \<Leftrightarrow> t : T"
0c5b22076fb3 tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents: 22609
diff changeset
   732
  and   "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<subseteq> \<Gamma>' \<Longrightarrow> valid \<Gamma>' \<Longrightarrow> \<Gamma>' \<turnstile> s \<leftrightarrow> t : T"
22531
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22501
diff changeset
   733
proof (nominal_induct \<Gamma> s t T and \<Gamma> s t T avoiding: \<Gamma>' rule: alg_equiv_alg_path_equiv.strong_inducts)
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   734
 case (QAT_Arrow x \<Gamma> s t T\<^isub>1 T\<^isub>2 \<Gamma>')
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
   735
  have fs:"x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" "x\<sharp>\<Gamma>'" 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: 22609
diff changeset
   736
  have h2:"\<Gamma> \<subseteq> \<Gamma>'" by fact
0c5b22076fb3 tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents: 22609
diff changeset
   737
  have ih:"\<And>\<Gamma>'. \<lbrakk>(x,T\<^isub>1)#\<Gamma> \<subseteq> \<Gamma>'; valid \<Gamma>'\<rbrakk>  \<Longrightarrow> \<Gamma>' \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2" by fact
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   738
  have "valid \<Gamma>'" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   739
  then have "valid ((x,T\<^isub>1)#\<Gamma>')" using fs by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   740
  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: 22609
diff changeset
   741
  have sub: "(x,T\<^isub>1)#\<Gamma> \<subseteq> (x,T\<^isub>1)#\<Gamma>'" using h2 by auto
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   742
  ultimately have "(x,T\<^isub>1)#\<Gamma>' \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2" using ih by simp
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   743
  then show "\<Gamma>' \<turnstile> s \<Leftrightarrow> t : T\<^isub>1\<rightarrow>T\<^isub>2" using fs by (auto simp add: fresh_prod)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   744
qed (auto)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   745
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   746
lemma path_equiv_implies_nf:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   747
  assumes "\<Gamma> \<turnstile> s \<leftrightarrow> t : T"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   748
  shows "s \<leadsto>|" and "t \<leadsto>|"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   749
using assms
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   750
by (induct rule: alg_equiv_alg_path_equiv.inducts(2)) (simp, auto)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   751
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   752
section {* Logical Equivalence *}
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   753
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   754
function log_equiv :: "(Ctxt \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool)" ("_ \<turnstile> _ is _ : _" [60,60,60,60] 60) 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   755
where    
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   756
   "\<Gamma> \<turnstile> s is t : TUnit = True"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   757
 | "\<Gamma> \<turnstile> s is t : TBase = \<Gamma> \<turnstile> s \<Leftrightarrow> t : TBase"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   758
 | "\<Gamma> \<turnstile> s is t : (T\<^isub>1 \<rightarrow> T\<^isub>2) =  
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: 22609
diff changeset
   759
    (\<forall>\<Gamma>' s' t'. \<Gamma>\<subseteq>\<Gamma>' \<longrightarrow> valid \<Gamma>' \<longrightarrow> \<Gamma>' \<turnstile> s' is t' : T\<^isub>1 \<longrightarrow>  (\<Gamma>' \<turnstile> (App s s') is (App t t') : T\<^isub>2))"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   760
apply (auto simp add: ty.inject)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   761
apply (subgoal_tac "(\<exists>T\<^isub>1 T\<^isub>2. b=T\<^isub>1 \<rightarrow> T\<^isub>2) \<or> b=TUnit \<or> b=TBase" )
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   762
apply (force)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   763
apply (rule ty_cases)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   764
done
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   765
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   766
termination
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   767
apply(relation "measure (\<lambda>(_,_,_,T). size T)")
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   768
apply(auto)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   769
done
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   770
24231
85fb973a8207 added type constraints to resolve syntax ambiguities;
wenzelm
parents: 24088
diff changeset
   771
lemma logical_monotonicity:
85fb973a8207 added type constraints to resolve syntax ambiguities;
wenzelm
parents: 24088
diff changeset
   772
 fixes \<Gamma> \<Gamma>' :: Ctxt
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   773
 assumes a1: "\<Gamma> \<turnstile> s is t : T" 
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: 22609
diff changeset
   774
 and     a2: "\<Gamma> \<subseteq> \<Gamma>'" 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   775
 and     a3: "valid \<Gamma>'"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   776
 shows "\<Gamma>' \<turnstile> s is t : T"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   777
using a1 a2 a3
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   778
proof (induct arbitrary: \<Gamma>' rule: log_equiv.induct)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   779
  case (2 \<Gamma> s t \<Gamma>')
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   780
  then show "\<Gamma>' \<turnstile> s is t : TBase" using algorithmic_monotonicity by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   781
next
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   782
  case (3 \<Gamma> s t T\<^isub>1 T\<^isub>2 \<Gamma>')
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   783
  have "\<Gamma> \<turnstile> s is t : T\<^isub>1\<rightarrow>T\<^isub>2" 
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   784
  and  "\<Gamma> \<subseteq> \<Gamma>'" 
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
   785
  and  "valid \<Gamma>'" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   786
  then show "\<Gamma>' \<turnstile> s is t : T\<^isub>1\<rightarrow>T\<^isub>2" by simp
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   787
qed (auto)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   788
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   789
lemma main_lemma:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   790
  shows "\<Gamma> \<turnstile> s is t : T \<Longrightarrow> valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   791
    and "\<Gamma> \<turnstile> p \<leftrightarrow> q : T \<Longrightarrow> \<Gamma> \<turnstile> p is q : T"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   792
proof (nominal_induct T arbitrary: \<Gamma> s t p q rule: ty.induct)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   793
  case (Arrow T\<^isub>1 T\<^isub>2)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   794
  { 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   795
    case (1 \<Gamma> s t)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   796
    have ih1:"\<And>\<Gamma> s t. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>2; valid \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>2" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   797
    have ih2:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s \<leftrightarrow> t : T\<^isub>1 \<Longrightarrow> \<Gamma> \<turnstile> s is t : T\<^isub>1" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   798
    have h:"\<Gamma> \<turnstile> s is t : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   799
    obtain x::name where fs:"x\<sharp>(\<Gamma>,s,t)" by (erule exists_fresh[OF fs_name1])
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   800
    have "valid \<Gamma>" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   801
    then have v: "valid ((x,T\<^isub>1)#\<Gamma>)" using fs by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   802
    then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> Var x \<leftrightarrow> Var x : T\<^isub>1" by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   803
    then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> Var x is Var x : T\<^isub>1" using ih2 by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   804
    then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) is App t (Var x) : T\<^isub>2" using h v by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   805
    then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2" using ih1 v by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   806
    then show "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1\<rightarrow>T\<^isub>2" using fs by (auto simp add: fresh_prod)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   807
  next
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   808
    case (2 \<Gamma> p q)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   809
    have h: "\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   810
    have ih1:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s \<leftrightarrow> t : T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> s is t : T\<^isub>2" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   811
    have ih2:"\<And>\<Gamma> s t. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>1; valid \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   812
    {
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   813
      fix \<Gamma>' s t
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: 22609
diff changeset
   814
      assume "\<Gamma> \<subseteq> \<Gamma>'" and hl:"\<Gamma>' \<turnstile> s is t : T\<^isub>1" and hk: "valid \<Gamma>'"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   815
      then have "\<Gamma>' \<turnstile> p \<leftrightarrow> q : T\<^isub>1 \<rightarrow> T\<^isub>2" using h algorithmic_monotonicity by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   816
      moreover have "\<Gamma>' \<turnstile> s \<Leftrightarrow> t : T\<^isub>1" using ih2 hl hk by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   817
      ultimately have "\<Gamma>' \<turnstile> App p s \<leftrightarrow> App q t : T\<^isub>2" by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   818
      then have "\<Gamma>' \<turnstile> App p s is App q t : T\<^isub>2" using ih1 by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   819
    }
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   820
    then show "\<Gamma> \<turnstile> p is q : T\<^isub>1\<rightarrow>T\<^isub>2"  by simp
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   821
  }
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   822
next
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   823
  case TBase
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   824
  { case 2
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   825
    have h:"\<Gamma> \<turnstile> s \<leftrightarrow> t : TBase" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   826
    then have "s \<leadsto>|" and "t \<leadsto>|" using path_equiv_implies_nf by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   827
    then have "s \<Down> s" and "t \<Down> t" by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   828
    then have "\<Gamma> \<turnstile> s \<Leftrightarrow> t : TBase" using h by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   829
    then show "\<Gamma> \<turnstile> s is t : TBase" by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   830
  }
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   831
qed (auto elim: alg_path_equiv_implies_valid) 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   832
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   833
corollary corollary_main:
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   834
  assumes a: "\<Gamma> \<turnstile> s \<leftrightarrow> t : T"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   835
  shows "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   836
using a main_lemma alg_path_equiv_implies_valid by blast
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   837
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   838
lemma logical_symmetry:
22082
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
   839
  assumes a: "\<Gamma> \<turnstile> s is t : T"
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   840
  shows "\<Gamma> \<turnstile> t is s : T"
22082
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
   841
using a 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   842
by (nominal_induct arbitrary: \<Gamma> s t rule: ty.induct) 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   843
   (auto simp add: algorithmic_symmetry)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   844
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   845
lemma logical_transitivity:
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   846
  assumes "\<Gamma> \<turnstile> s is t : T" "\<Gamma> \<turnstile> t is u : T" 
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   847
  shows "\<Gamma> \<turnstile> s is u : T"
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   848
using assms
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   849
proof (nominal_induct arbitrary: \<Gamma> s t u  rule:ty.induct)
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   850
  case TBase
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   851
  then show "\<Gamma> \<turnstile> s is u : TBase" by (auto elim:  algorithmic_transitivity)
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   852
next 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   853
  case (Arrow T\<^isub>1 T\<^isub>2 \<Gamma> s t u)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   854
  have h1:"\<Gamma> \<turnstile> s is t : T\<^isub>1 \<rightarrow> T\<^isub>2" by fact
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   855
  have h2:"\<Gamma> \<turnstile> t is u : T\<^isub>1 \<rightarrow> T\<^isub>2" by fact
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   856
  have ih1:"\<And>\<Gamma> s t u. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>1; \<Gamma> \<turnstile> t is u : T\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s is u : T\<^isub>1" by fact
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   857
  have ih2:"\<And>\<Gamma> s t u. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>2; \<Gamma> \<turnstile> t is u : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s is u : T\<^isub>2" by fact
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   858
  {
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   859
    fix \<Gamma>' s' u'
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: 22609
diff changeset
   860
    assume hsub:"\<Gamma> \<subseteq> \<Gamma>'" and hl:"\<Gamma>' \<turnstile> s' is u' : T\<^isub>1" and hk: "valid \<Gamma>'"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   861
    then have "\<Gamma>' \<turnstile> u' is s' : T\<^isub>1" using logical_symmetry by blast
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   862
    then have "\<Gamma>' \<turnstile> u' is u' : T\<^isub>1" using ih1 hl by blast
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   863
    then have "\<Gamma>' \<turnstile> App t u' is App u u' : T\<^isub>2" using h2 hsub hk by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   864
    moreover have "\<Gamma>' \<turnstile>  App s s' is App t u' : T\<^isub>2" using h1 hsub hl hk by auto
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   865
    ultimately have "\<Gamma>' \<turnstile>  App s s' is App u u' : T\<^isub>2" using ih2 by blast
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   866
  }
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   867
  then show "\<Gamma> \<turnstile> s is u : T\<^isub>1 \<rightarrow> T\<^isub>2" by auto
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   868
qed (auto)
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   869
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   870
lemma logical_weak_head_closure:
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   871
  assumes a: "\<Gamma> \<turnstile> s is t : T" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   872
  and     b: "s' \<leadsto> s" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   873
  and     c: "t' \<leadsto> t"
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   874
  shows "\<Gamma> \<turnstile> s' is t' : T"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   875
using a b c algorithmic_weak_head_closure 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   876
by (nominal_induct arbitrary: \<Gamma> s t s' t' rule: ty.induct) 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   877
   (auto, blast)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   878
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   879
lemma logical_weak_head_closure':
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   880
  assumes "\<Gamma> \<turnstile> s is t : T" and "s' \<leadsto> s" 
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   881
  shows "\<Gamma> \<turnstile> s' is t : T"
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   882
using assms
22082
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
   883
proof (nominal_induct arbitrary: \<Gamma> s t s' rule: ty.induct)
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   884
  case (TBase  \<Gamma> s t s')
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   885
  then show ?case by force
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   886
next
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   887
  case (TUnit \<Gamma> s t s')
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   888
  then show ?case by auto
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   889
next 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   890
  case (Arrow T\<^isub>1 T\<^isub>2 \<Gamma> s t s')
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   891
  have h1:"s' \<leadsto> s" by fact
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   892
  have ih:"\<And>\<Gamma> s t s'. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>2; s' \<leadsto> s\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s' is t : T\<^isub>2" by fact
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   893
  have h2:"\<Gamma> \<turnstile> s is t : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   894
  then 
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: 22609
diff changeset
   895
  have hb:"\<forall>\<Gamma>' s' t'. \<Gamma>\<subseteq>\<Gamma>' \<longrightarrow> valid \<Gamma>' \<longrightarrow> \<Gamma>' \<turnstile> s' is t' : T\<^isub>1 \<longrightarrow> (\<Gamma>' \<turnstile> (App s s') is (App t t') : T\<^isub>2)" 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   896
    by auto
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   897
  {
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   898
    fix \<Gamma>' s\<^isub>2 t\<^isub>2
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: 22609
diff changeset
   899
    assume "\<Gamma> \<subseteq> \<Gamma>'" and "\<Gamma>' \<turnstile> s\<^isub>2 is t\<^isub>2 : T\<^isub>1" and "valid \<Gamma>'"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   900
    then have "\<Gamma>' \<turnstile> (App s s\<^isub>2) is (App t t\<^isub>2) : T\<^isub>2" using hb by auto
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   901
    moreover have "(App s' s\<^isub>2)  \<leadsto> (App s s\<^isub>2)" using h1 by auto  
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   902
    ultimately have "\<Gamma>' \<turnstile> App s' s\<^isub>2 is App t t\<^isub>2 : T\<^isub>2" using ih by auto
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   903
  }
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   904
  then show "\<Gamma> \<turnstile> s' is t : T\<^isub>1\<rightarrow>T\<^isub>2" by auto
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   905
qed 
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   906
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   907
abbreviation 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   908
 log_equiv_for_psubsts :: "Ctxt \<Rightarrow> Subst \<Rightarrow> Subst \<Rightarrow> Ctxt \<Rightarrow> bool"  ("_ \<turnstile> _ is _ over _" [60,60] 60) 
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   909
where
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   910
  "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma> \<equiv> \<forall>x T. (x,T) \<in> set \<Gamma> \<longrightarrow> \<Gamma>' \<turnstile> \<theta><Var x> is  \<theta>'<Var x> : T"
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   911
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   912
lemma logical_pseudo_reflexivity:
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   913
  assumes "\<Gamma>' \<turnstile> t is s over \<Gamma>"
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   914
  shows "\<Gamma>' \<turnstile> s is s over \<Gamma>" 
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   915
proof -
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   916
  have "\<Gamma>' \<turnstile> t is s over \<Gamma>" by fact
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   917
  moreover then have "\<Gamma>' \<turnstile> s is t over \<Gamma>" using logical_symmetry by blast
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   918
  ultimately show "\<Gamma>' \<turnstile> s is s over \<Gamma>" using logical_transitivity by blast
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   919
qed
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   920
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   921
lemma logical_subst_monotonicity :
24231
85fb973a8207 added type constraints to resolve syntax ambiguities;
wenzelm
parents: 24088
diff changeset
   922
  fixes \<Gamma> \<Gamma>' \<Gamma>'' :: Ctxt
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   923
  assumes a: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>" 
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: 22609
diff changeset
   924
  and     b: "\<Gamma>' \<subseteq> \<Gamma>''"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   925
  and     c: "valid \<Gamma>''"
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   926
  shows "\<Gamma>'' \<turnstile> \<theta> is \<theta>' over \<Gamma>"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   927
using a b c logical_monotonicity by blast
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   928
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   929
lemma equiv_subst_ext :
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   930
  assumes h1: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   931
  and     h2: "\<Gamma>' \<turnstile> s is t : T" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   932
  and     fs: "x\<sharp>\<Gamma>"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   933
  shows "\<Gamma>' \<turnstile> (x,s)#\<theta> is (x,t)#\<theta>' over (x,T)#\<Gamma>"
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   934
using assms
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   935
proof -
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   936
  {
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   937
    fix y U
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   938
    assume "(y,U) \<in> set ((x,T)#\<Gamma>)"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   939
    moreover
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   940
    { 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   941
      assume "(y,U) \<in> set [(x,T)]"
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
   942
      with h2 have "\<Gamma>' \<turnstile> ((x,s)#\<theta>)<Var y> is ((x,t)#\<theta>')<Var y> : U" by auto 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   943
    }
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   944
    moreover
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   945
    {
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   946
      assume hl:"(y,U) \<in> set \<Gamma>"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   947
      then have "\<not> y\<sharp>\<Gamma>" by (induct \<Gamma>) (auto simp add: fresh_list_cons fresh_atm fresh_prod)
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   948
      then have hf:"x\<sharp> Var y" using fs by (auto simp add: fresh_atm)
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   949
      then have "((x,s)#\<theta>)<Var y> = \<theta><Var y>" "((x,t)#\<theta>')<Var y> = \<theta>'<Var y>" 
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   950
	using fresh_psubst_simp by blast+ 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   951
      moreover have  "\<Gamma>' \<turnstile> \<theta><Var y> is \<theta>'<Var y> : U" using h1 hl by auto
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   952
      ultimately have "\<Gamma>' \<turnstile> ((x,s)#\<theta>)<Var y> is ((x,t)#\<theta>')<Var y> : U" by auto
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   953
    }
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   954
    ultimately have "\<Gamma>' \<turnstile> ((x,s)#\<theta>)<Var y> is ((x,t)#\<theta>')<Var y> : U" by auto
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   955
  }
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   956
  then show "\<Gamma>' \<turnstile> (x,s)#\<theta> is (x,t)#\<theta>' over (x,T)#\<Gamma>" by auto
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   957
qed
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   958
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   959
theorem fundamental_theorem_1:
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   960
  assumes a1: "\<Gamma> \<turnstile> t : T" 
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   961
  and     a2: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   962
  and     a3: "valid \<Gamma>'" 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   963
  shows "\<Gamma>' \<turnstile> \<theta><t> is \<theta>'<t> : T"   
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   964
using a1 a2 a3
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   965
proof (nominal_induct \<Gamma> t T avoiding: \<theta> \<theta>' arbitrary: \<Gamma>' rule: typing.strong_induct)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   966
  case (T_Lam x \<Gamma> T\<^isub>1 t\<^isub>2 T\<^isub>2 \<theta> \<theta>' \<Gamma>')
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
   967
  have vc: "x\<sharp>\<theta>" "x\<sharp>\<theta>'" "x\<sharp>\<Gamma>" by fact+
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   968
  have asm1: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>" by fact
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   969
  have ih:"\<And>\<theta> \<theta>' \<Gamma>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over (x,T\<^isub>1)#\<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><t\<^isub>2> is \<theta>'<t\<^isub>2> : T\<^isub>2" by fact
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   970
  show "\<Gamma>' \<turnstile> \<theta><Lam [x].t\<^isub>2> is \<theta>'<Lam [x].t\<^isub>2> : T\<^isub>1\<rightarrow>T\<^isub>2" using vc
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   971
  proof (simp, intro strip)
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   972
    fix \<Gamma>'' s' t'
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   973
    assume sub: "\<Gamma>' \<subseteq> \<Gamma>''" 
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   974
    and    asm2: "\<Gamma>''\<turnstile> s' is t' : T\<^isub>1" 
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   975
    and    val: "valid \<Gamma>''"
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   976
    from asm1 val sub have "\<Gamma>'' \<turnstile> \<theta> is \<theta>' over \<Gamma>" using logical_subst_monotonicity by blast
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   977
    with asm2 vc have "\<Gamma>'' \<turnstile> (x,s')#\<theta> is (x,t')#\<theta>' over (x,T\<^isub>1)#\<Gamma>" using equiv_subst_ext by blast
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   978
    with ih val have "\<Gamma>'' \<turnstile> ((x,s')#\<theta>)<t\<^isub>2> is ((x,t')#\<theta>')<t\<^isub>2> : T\<^isub>2" by auto
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   979
    with vc have "\<Gamma>''\<turnstile>\<theta><t\<^isub>2>[x::=s'] is \<theta>'<t\<^isub>2>[x::=t'] : T\<^isub>2" by (simp add: psubst_subst_psubst)
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   980
    moreover 
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   981
    have "App (Lam [x].\<theta><t\<^isub>2>) s' \<leadsto> \<theta><t\<^isub>2>[x::=s']" by auto
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   982
    moreover 
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   983
    have "App (Lam [x].\<theta>'<t\<^isub>2>) t' \<leadsto> \<theta>'<t\<^isub>2>[x::=t']" by auto
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   984
    ultimately show "\<Gamma>''\<turnstile> App (Lam [x].\<theta><t\<^isub>2>) s' is App (Lam [x].\<theta>'<t\<^isub>2>) t' : T\<^isub>2" 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   985
      using logical_weak_head_closure by auto
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   986
  qed
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   987
qed (auto)
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   988
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
   989
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
   990
theorem fundamental_theorem_2:
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   991
  assumes h1: "\<Gamma> \<turnstile> s \<equiv> t : T" 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   992
  and     h2: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   993
  and     h3: "valid \<Gamma>'"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   994
  shows "\<Gamma>' \<turnstile> \<theta><s> is \<theta>'<t> : T"
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   995
using h1 h2 h3
22531
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22501
diff changeset
   996
proof (nominal_induct \<Gamma> s t T avoiding: \<Gamma>' \<theta> \<theta>' rule: def_equiv.strong_induct)
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   997
  case (Q_Refl \<Gamma> t T \<Gamma>' \<theta> \<theta>')
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
   998
  have "\<Gamma> \<turnstile> t : T" 
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
   999
  and  "valid \<Gamma>'" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1000
  moreover 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1001
  have "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1002
  ultimately show "\<Gamma>' \<turnstile> \<theta><t> is \<theta>'<t> : T" using fundamental_theorem_1 by blast
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1003
next
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1004
  case (Q_Symm \<Gamma> t s T \<Gamma>' \<theta> \<theta>')
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1005
  have "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>" 
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
  1006
  and "valid \<Gamma>'" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1007
  moreover 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1008
  have ih: "\<And> \<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><t> is \<theta>'<s> : T" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1009
  ultimately show "\<Gamma>' \<turnstile> \<theta><s> is \<theta>'<t> : T" using logical_symmetry by blast
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1010
next
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1011
  case (Q_Trans \<Gamma> s t T u \<Gamma>' \<theta> \<theta>')
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1012
  have ih1: "\<And> \<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><s> is \<theta>'<t> : T" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1013
  have ih2: "\<And> \<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><t> is \<theta>'<u> : T" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1014
  have h: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>" 
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
  1015
  and  v: "valid \<Gamma>'" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1016
  then have "\<Gamma>' \<turnstile> \<theta>' is \<theta>' over \<Gamma>" using logical_pseudo_reflexivity by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1017
  then have "\<Gamma>' \<turnstile> \<theta>'<t> is \<theta>'<u> : T" using ih2 v by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1018
  moreover have "\<Gamma>' \<turnstile> \<theta><s> is \<theta>'<t> : T" using ih1 h v by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1019
  ultimately show "\<Gamma>' \<turnstile> \<theta><s> is \<theta>'<u> : T" using logical_transitivity by blast
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1020
next
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1021
  case (Q_Abs x \<Gamma> T\<^isub>1 s\<^isub>2 t\<^isub>2 T\<^isub>2 \<Gamma>' \<theta> \<theta>')
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1022
  have fs:"x\<sharp>\<Gamma>" by fact
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
  1023
  have fs2: "x\<sharp>\<theta>" "x\<sharp>\<theta>'" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1024
  have h2: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>" 
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
  1025
  and  h3: "valid \<Gamma>'" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1026
  have ih:"\<And>\<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over (x,T\<^isub>1)#\<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><s\<^isub>2> is \<theta>'<t\<^isub>2> : T\<^isub>2" by fact
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1027
  {
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1028
    fix \<Gamma>'' s' t'
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: 22609
diff changeset
  1029
    assume "\<Gamma>' \<subseteq> \<Gamma>''" and hl:"\<Gamma>''\<turnstile> s' is t' : T\<^isub>1" and hk: "valid \<Gamma>''"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1030
    then have "\<Gamma>'' \<turnstile> \<theta> is \<theta>' over \<Gamma>" using h2 logical_subst_monotonicity by blast
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1031
    then have "\<Gamma>'' \<turnstile> (x,s')#\<theta> is (x,t')#\<theta>' over (x,T\<^isub>1)#\<Gamma>" using equiv_subst_ext hl fs by blast
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
  1032
    then have "\<Gamma>'' \<turnstile> ((x,s')#\<theta>)<s\<^isub>2> is ((x,t')#\<theta>')<t\<^isub>2> : T\<^isub>2" using ih hk by blast
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1033
    then have "\<Gamma>''\<turnstile> \<theta><s\<^isub>2>[x::=s'] is \<theta>'<t\<^isub>2>[x::=t'] : T\<^isub>2" using fs2 psubst_subst_psubst by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1034
    moreover have "App (Lam [x]. \<theta><s\<^isub>2>) s' \<leadsto>  \<theta><s\<^isub>2>[x::=s']" 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1035
              and "App (Lam [x].\<theta>'<t\<^isub>2>) t' \<leadsto> \<theta>'<t\<^isub>2>[x::=t']" by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1036
    ultimately have "\<Gamma>'' \<turnstile> App (Lam [x]. \<theta><s\<^isub>2>) s' is App (Lam [x].\<theta>'<t\<^isub>2>) t' : T\<^isub>2" 
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1037
      using logical_weak_head_closure by auto
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1038
  }
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
  1039
  moreover have "valid \<Gamma>'" by fact
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1040
  ultimately have "\<Gamma>' \<turnstile> Lam [x].\<theta><s\<^isub>2> is Lam [x].\<theta>'<t\<^isub>2> : T\<^isub>1\<rightarrow>T\<^isub>2" by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1041
  then show "\<Gamma>' \<turnstile> \<theta><Lam [x].s\<^isub>2> is \<theta>'<Lam [x].t\<^isub>2> : T\<^isub>1\<rightarrow>T\<^isub>2" using fs2 by auto
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1042
next
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1043
  case (Q_App \<Gamma> s\<^isub>1 t\<^isub>1 T\<^isub>1 T\<^isub>2 s\<^isub>2 t\<^isub>2 \<Gamma>' \<theta> \<theta>')
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1044
  then show "\<Gamma>' \<turnstile> \<theta><App s\<^isub>1 s\<^isub>2> is \<theta>'<App t\<^isub>1 t\<^isub>2> : T\<^isub>2" by auto 
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1045
next
22531
1cbfb4066e47 Adapted to changes in nominal_inductive.
berghofe
parents: 22501
diff changeset
  1046
  case (Q_Beta x \<Gamma> s\<^isub>2 t\<^isub>2 T\<^isub>1 s12 t12 T\<^isub>2 \<Gamma>' \<theta> \<theta>')
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1047
  have h: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>" 
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
  1048
  and  h': "valid \<Gamma>'" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1049
  have fs: "x\<sharp>\<Gamma>" by fact
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
  1050
  have fs2: " x\<sharp>\<theta>" "x\<sharp>\<theta>'" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1051
  have ih1: "\<And>\<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><s\<^isub>2> is \<theta>'<t\<^isub>2> : T\<^isub>1" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1052
  have ih2: "\<And>\<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over (x,T\<^isub>1)#\<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><s12> is \<theta>'<t12> : T\<^isub>2" by fact
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1053
  have "\<Gamma>' \<turnstile> \<theta><s\<^isub>2> is \<theta>'<t\<^isub>2> : T\<^isub>1" using ih1 h' h by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1054
  then have "\<Gamma>' \<turnstile> (x,\<theta><s\<^isub>2>)#\<theta> is (x,\<theta>'<t\<^isub>2>)#\<theta>' over (x,T\<^isub>1)#\<Gamma>" using equiv_subst_ext h fs by blast
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
  1055
  then have "\<Gamma>' \<turnstile> ((x,\<theta><s\<^isub>2>)#\<theta>)<s12> is ((x,\<theta>'<t\<^isub>2>)#\<theta>')<t12> : T\<^isub>2" using ih2 h' by auto
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1056
  then have "\<Gamma>' \<turnstile> \<theta><s12>[x::=\<theta><s\<^isub>2>] is \<theta>'<t12>[x::=\<theta>'<t\<^isub>2>] : T\<^isub>2" using fs2 psubst_subst_psubst by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1057
  then have "\<Gamma>' \<turnstile> \<theta><s12>[x::=\<theta><s\<^isub>2>] is \<theta>'<t12[x::=t\<^isub>2]> : T\<^isub>2" using fs2 psubst_subst_propagate by auto
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1058
  moreover have "App (Lam [x].\<theta><s12>) (\<theta><s\<^isub>2>) \<leadsto> \<theta><s12>[x::=\<theta><s\<^isub>2>]" by auto 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1059
  ultimately have "\<Gamma>' \<turnstile> App (Lam [x].\<theta><s12>) (\<theta><s\<^isub>2>) is \<theta>'<t12[x::=t\<^isub>2]> : T\<^isub>2" 
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1060
    using logical_weak_head_closure' by auto
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1061
  then show "\<Gamma>' \<turnstile> \<theta><App (Lam [x].s12) s\<^isub>2> is \<theta>'<t12[x::=t\<^isub>2]> : T\<^isub>2" using fs2 by simp
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1062
next
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1063
  case (Q_Ext x \<Gamma> s t T\<^isub>1 T\<^isub>2 \<Gamma>' \<theta> \<theta>')
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1064
  have h2: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>" 
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
  1065
  and  h2': "valid \<Gamma>'" by fact+
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
  1066
  have fs:"x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" by fact+
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1067
  have ih:"\<And>\<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over (x,T\<^isub>1)#\<Gamma>; valid \<Gamma>'\<rbrakk> 
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1068
                          \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><App s (Var x)> is \<theta>'<App t (Var x)> : T\<^isub>2" by fact
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1069
   {
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1070
    fix \<Gamma>'' s' t'
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: 22609
diff changeset
  1071
    assume hsub: "\<Gamma>' \<subseteq> \<Gamma>''" and hl: "\<Gamma>''\<turnstile> s' is t' : T\<^isub>1" and hk: "valid \<Gamma>''"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1072
    then have "\<Gamma>'' \<turnstile> \<theta> is \<theta>' over \<Gamma>" using h2 logical_subst_monotonicity by blast
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1073
    then have "\<Gamma>'' \<turnstile> (x,s')#\<theta> is (x,t')#\<theta>' over (x,T\<^isub>1)#\<Gamma>" using equiv_subst_ext hl fs by blast
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
  1074
    then have "\<Gamma>'' \<turnstile> ((x,s')#\<theta>)<App s (Var x)>  is ((x,t')#\<theta>')<App t (Var x)> : T\<^isub>2" using ih hk by blast
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1075
    then 
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
  1076
    have "\<Gamma>'' \<turnstile> App (((x,s')#\<theta>)<s>) (((x,s')#\<theta>)<(Var x)>) is App (((x,t')#\<theta>')<t>) (((x,t')#\<theta>')<(Var x)>) : T\<^isub>2"
22082
b1be13d32efd tuned a bit the proofs
urbanc
parents: 22073
diff changeset
  1077
      by auto
24070
ff4c715a11cd updated some of the definitions and proofs
urbanc
parents: 23760
diff changeset
  1078
    then have "\<Gamma>'' \<turnstile> App ((x,s')#\<theta>)<s> s'  is App ((x,t')#\<theta>')<t> t' : T\<^isub>2" by auto
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1079
    then have "\<Gamma>'' \<turnstile> App (\<theta><s>) s' is App (\<theta>'<t>) t' : T\<^isub>2" using fs fresh_psubst_simp by auto
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1080
  }
25107
dbf09ca6a80e tuned proofs: avoid implicit prems;
wenzelm
parents: 24231
diff changeset
  1081
  moreover have "valid \<Gamma>'" by fact
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1082
  ultimately show "\<Gamma>' \<turnstile> \<theta><s> is \<theta>'<t> : T\<^isub>1\<rightarrow>T\<^isub>2" by auto
23370
513a8ee192f1 added the Q_Unit rule (was missing) and adjusted the proof accordingly
urbanc
parents: 22829
diff changeset
  1083
next
513a8ee192f1 added the Q_Unit rule (was missing) and adjusted the proof accordingly
urbanc
parents: 22829
diff changeset
  1084
  case (Q_Unit \<Gamma> s t \<Gamma>' \<theta> \<theta>')  
513a8ee192f1 added the Q_Unit rule (was missing) and adjusted the proof accordingly
urbanc
parents: 22829
diff changeset
  1085
  then show "\<Gamma>' \<turnstile> \<theta><s> is \<theta>'<t> : TUnit" by auto
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1086
qed
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1087
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
  1088
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1089
theorem completeness:
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1090
  assumes asm: "\<Gamma> \<turnstile> s \<equiv> t : T"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
  1091
  shows   "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T"
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1092
proof -
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1093
  have val: "valid \<Gamma>" using def_equiv_implies_valid asm by simp
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1094
  moreover
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1095
  {
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1096
    fix x T
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1097
    assume "(x,T) \<in> set \<Gamma>" "valid \<Gamma>"
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1098
    then have "\<Gamma> \<turnstile> Var x is Var x : T" using main_lemma(2) by blast
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1099
  }
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1100
  ultimately have "\<Gamma> \<turnstile> [] is [] over \<Gamma>" by auto 
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1101
  then have "\<Gamma> \<turnstile> []<s> is []<t> : T" using fundamental_theorem_2 val asm by blast
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1102
  then have "\<Gamma> \<turnstile> s is t : T" by simp
22494
b61306c7987a tuned some proofs
urbanc
parents: 22492
diff changeset
  1103
  then show  "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T" using main_lemma(1) val by simp
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1104
qed
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1105
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
  1106
text {* We leave soundness as an exercise - like in the book :-) \\ 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
  1107
 @{prop[mode=IfThen] "\<lbrakk>\<Gamma> \<turnstile> s \<Leftrightarrow> t : T; \<Gamma> \<turnstile> t : T; \<Gamma> \<turnstile> s : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T"} \\
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
  1108
 @{prop "\<lbrakk>\<Gamma> \<turnstile> s \<leftrightarrow> t : T; \<Gamma> \<turnstile> t : T; \<Gamma> \<turnstile> s : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T"} 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
  1109
*}
22073
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1110
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1111
end
c170dcbe6c9d formalisation of Crary's chapter on logical relations
urbanc
parents:
diff changeset
  1112