src/HOL/Nominal/Examples/Weakening.thy
author narboux
Tue, 06 Mar 2007 16:40:32 +0100
changeset 22421 51a18dd1ea86
parent 22418 49e2d9744ae1
child 22533 62c76754da32
permissions -rw-r--r--
correct typo in latex output
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18269
3f36e2165e51 some small tuning
urbanc
parents: 18105
diff changeset
     1
(* $Id$ *)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
     2
19496
79dbe35c6cba Capitalized theory names.
berghofe
parents: 18654
diff changeset
     3
theory Weakening 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
     4
imports "../Nominal" 
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
     5
begin
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
     6
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
     7
section {* Weakening Example for the Simply-Typed Lambda-Calculus *}
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
     8
(*================================================================*)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
     9
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    10
atom_decl name 
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    11
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    12
nominal_datatype lam = 
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    13
    Var "name"
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    14
  | App "lam" "lam"
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    15
  | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    16
18352
urbanc
parents: 18346
diff changeset
    17
nominal_datatype ty =
urbanc
parents: 18346
diff changeset
    18
    TVar "nat"
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    19
  | TArr "ty" "ty" (infix "\<rightarrow>" 200)
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    20
21488
e1b260d204a0 fixed some typos
urbanc
parents: 21487
diff changeset
    21
lemma ty_perm[simp]:
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    22
  fixes pi ::"name prm"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    23
  and   T  ::"ty"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    24
  shows "pi\<bullet>T = T"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    25
by (induct T rule: ty.induct_weak)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    26
   (simp_all add: perm_nat_def)  
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    27
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    28
text {* valid contexts *}
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    29
inductive2
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    30
  valid :: "(name\<times>ty) list \<Rightarrow> bool"
21364
3baf57a27269 inductive2: canonical specification syntax;
wenzelm
parents: 21107
diff changeset
    31
where
3baf57a27269 inductive2: canonical specification syntax;
wenzelm
parents: 21107
diff changeset
    32
    v1[intro]: "valid []"
3baf57a27269 inductive2: canonical specification syntax;
wenzelm
parents: 21107
diff changeset
    33
  | v2[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)"
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    34
22231
f76f187c91f9 added an infrastructure that allows the user to declare lemmas to be equivariance lemmas; the intention is to use these lemmas in automated tools but also can be employed by the user
urbanc
parents: 21488
diff changeset
    35
lemma eqvt_valid[eqvt]:
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    36
  fixes   pi:: "name prm"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    37
  assumes a: "valid \<Gamma>"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    38
  shows   "valid (pi\<bullet>\<Gamma>)"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    39
using a
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    40
by (induct) (auto simp add: fresh_bij)
22231
f76f187c91f9 added an infrastructure that allows the user to declare lemmas to be equivariance lemmas; the intention is to use these lemmas in automated tools but also can be employed by the user
urbanc
parents: 21488
diff changeset
    41
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    42
text{* typing judgements *}
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    43
inductive2
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    44
  typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [80,80,80] 80) 
21364
3baf57a27269 inductive2: canonical specification syntax;
wenzelm
parents: 21107
diff changeset
    45
where
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    46
    t_Var[intro]: "\<lbrakk>valid \<Gamma>; (a,T)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var a : T"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    47
  | t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1\<rightarrow>T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    48
  | t_Lam[intro]: "\<lbrakk>a\<sharp>\<Gamma>;((a,T1)#\<Gamma>) \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [a].t : T1\<rightarrow>T2"
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    49
22231
f76f187c91f9 added an infrastructure that allows the user to declare lemmas to be equivariance lemmas; the intention is to use these lemmas in automated tools but also can be employed by the user
urbanc
parents: 21488
diff changeset
    50
lemma eqvt_typing[eqvt]: 
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    51
  fixes pi:: "name prm"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    52
  assumes a: "\<Gamma> \<turnstile> t : T"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    53
  shows "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>t) : (pi\<bullet>T)"
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    54
using a
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    55
proof (induct)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    56
  case (t_Var \<Gamma> a T)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    57
  have "valid (pi\<bullet>\<Gamma>)" by (rule eqvt_valid)
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    58
  moreover
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    59
  have "(pi\<bullet>(a,T))\<in>(pi\<bullet>set \<Gamma>)" by (rule pt_set_bij2[OF pt_name_inst, OF at_name_inst])
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    60
  ultimately show "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>Var a) : (pi\<bullet>T)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    61
    using typing.intros by (force simp add: set_eqvt)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    62
next 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    63
  case (t_Lam a \<Gamma> T1 t T2)
19972
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19496
diff changeset
    64
  moreover have "(pi\<bullet>a)\<sharp>(pi\<bullet>\<Gamma>)" by (simp add: fresh_bij)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    65
  ultimately show "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>Lam [a].t) :(pi\<bullet>T1\<rightarrow>T2)" by force 
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    66
qed (auto)
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    67
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    68
text {* the strong induction principle needs to be derived manually *}
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    69
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    70
lemma typing_induct[consumes 1, case_names t_Var t_App t_Lam]:
18296
3dcc34f18bfa adapted to the new nominal_induction
urbanc
parents: 18269
diff changeset
    71
  fixes  P :: "'a::fs_name\<Rightarrow>(name\<times>ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow>bool"
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    72
  and    \<Gamma> :: "(name\<times>ty) list"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    73
  and    t :: "lam"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    74
  and    T :: "ty"
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    75
  and    x :: "'a::fs_name"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    76
  assumes a: "\<Gamma> \<turnstile> t : T"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    77
  and a1:    "\<And>\<Gamma> a T x. \<lbrakk>valid \<Gamma>; (a,T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P x \<Gamma> (Var a) T"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    78
  and a2:    "\<And>\<Gamma> T1 T2 t1 t2 x. \<lbrakk>\<And>z. P z \<Gamma> t1 (T1\<rightarrow>T2); \<And>z. P z \<Gamma> t2 T1\<rbrakk>
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    79
              \<Longrightarrow> P x \<Gamma> (App t1 t2) T2"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    80
  and a3:    "\<And>a \<Gamma> T1 T2 t x. \<lbrakk>a\<sharp>x; a\<sharp>\<Gamma>; \<And>z. P z ((a,T1)#\<Gamma>) t T2\<rbrakk>
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    81
              \<Longrightarrow> P x \<Gamma> (Lam [a].t) (T1\<rightarrow>T2)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    82
  shows "P x \<Gamma> t T"
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
    83
proof -
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    84
  from a have "\<And>(pi::name prm) x. P x (pi\<bullet>\<Gamma>) (pi\<bullet>t) (pi\<bullet>T)"
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
    85
  proof (induct)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    86
    case (t_Var \<Gamma> a T)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    87
    have "valid \<Gamma>" by fact
22231
f76f187c91f9 added an infrastructure that allows the user to declare lemmas to be equivariance lemmas; the intention is to use these lemmas in automated tools but also can be employed by the user
urbanc
parents: 21488
diff changeset
    88
    then have "valid (pi\<bullet>\<Gamma>)" by (rule eqvt)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    89
    moreover
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    90
    have "(a,T)\<in>set \<Gamma>" by fact
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    91
    then have "pi\<bullet>(a,T)\<in>pi\<bullet>(set \<Gamma>)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst])  
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    92
    then have "(pi\<bullet>a,T)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: set_eqvt)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    93
    ultimately show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Var a)) (pi\<bullet>T)" using a1 by simp
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
    94
  next
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    95
    case (t_App \<Gamma> t1 T1 T2 t2)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    96
    thus "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(App t1 t2)) (pi\<bullet>T2)" using a2 
22231
f76f187c91f9 added an infrastructure that allows the user to declare lemmas to be equivariance lemmas; the intention is to use these lemmas in automated tools but also can be employed by the user
urbanc
parents: 21488
diff changeset
    97
      by (simp only: eqvt) (blast)
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
    98
  next
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
    99
    case (t_Lam a \<Gamma> T1 t T2)
21487
45f9163d79e7 tuned the proof of the strong induction principle
urbanc
parents: 21405
diff changeset
   100
    obtain c::"name" where fs: "c\<sharp>(pi\<bullet>a,pi\<bullet>t,pi\<bullet>\<Gamma>,x)" by (rule exists_fresh[OF fs_name1])
45f9163d79e7 tuned the proof of the strong induction principle
urbanc
parents: 21405
diff changeset
   101
    let ?sw="[(pi\<bullet>a,c)]"
45f9163d79e7 tuned the proof of the strong induction principle
urbanc
parents: 21405
diff changeset
   102
    let ?pi'="?sw@pi"
45f9163d79e7 tuned the proof of the strong induction principle
urbanc
parents: 21405
diff changeset
   103
    have f1: "a\<sharp>\<Gamma>" by fact
45f9163d79e7 tuned the proof of the strong induction principle
urbanc
parents: 21405
diff changeset
   104
    have f2: "(pi\<bullet>a)\<sharp>(pi\<bullet>\<Gamma>)" using f1 by (simp add: fresh_bij)
45f9163d79e7 tuned the proof of the strong induction principle
urbanc
parents: 21405
diff changeset
   105
    have f3: "c\<sharp>?pi'\<bullet>\<Gamma>" using f1 by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   106
    have ih1: "\<And>x. P x (?pi'\<bullet>((a,T1)#\<Gamma>)) (?pi'\<bullet>t) (?pi'\<bullet>T2)" by fact
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   107
    then have "\<And>x. P x ((c,T1)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>t) (?pi'\<bullet>T2)" by (simp add: calc_atm)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   108
    then have "P x (?pi'\<bullet>\<Gamma>) (Lam [c].(?pi'\<bullet>t)) (T1\<rightarrow>T2)" using a3 f3 fs by simp
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   109
    then have "P x (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>(Lam [(pi\<bullet>a)].(pi\<bullet>t))) (T1\<rightarrow>T2)" 
21487
45f9163d79e7 tuned the proof of the strong induction principle
urbanc
parents: 21405
diff changeset
   110
      by (simp del: append_Cons add: calc_atm pt_name2)
45f9163d79e7 tuned the proof of the strong induction principle
urbanc
parents: 21405
diff changeset
   111
    moreover have "(?sw\<bullet>(pi\<bullet>\<Gamma>)) = (pi\<bullet>\<Gamma>)" 
45f9163d79e7 tuned the proof of the strong induction principle
urbanc
parents: 21405
diff changeset
   112
      by (rule perm_fresh_fresh) (simp_all add: fs f2)
45f9163d79e7 tuned the proof of the strong induction principle
urbanc
parents: 21405
diff changeset
   113
    moreover have "(?sw\<bullet>(Lam [(pi\<bullet>a)].(pi\<bullet>t))) = Lam [(pi\<bullet>a)].(pi\<bullet>t)" 
45f9163d79e7 tuned the proof of the strong induction principle
urbanc
parents: 21405
diff changeset
   114
      by (rule perm_fresh_fresh) (simp_all add: fs f2 abs_fresh)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   115
    ultimately show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Lam [a].t)) (pi\<bullet>T1\<rightarrow>T2)" by (simp)
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   116
  qed
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   117
  hence "P x (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>t) (([]::name prm)\<bullet>T)" by blast
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   118
  thus "P x \<Gamma> t T" by simp
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   119
qed
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   120
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   121
text {* definition of a subcontext *}
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   122
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   123
abbreviation
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21377
diff changeset
   124
  "sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (" _ \<lless> _ " [80,80] 80) where
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   125
  "\<Gamma>1 \<lless> \<Gamma>2 \<equiv> \<forall>a \<sigma>. (a,\<sigma>)\<in>set \<Gamma>1 \<longrightarrow> (a,\<sigma>)\<in>set \<Gamma>2"
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   126
21488
e1b260d204a0 fixed some typos
urbanc
parents: 21487
diff changeset
   127
text {* now it comes: The Weakening Lemma *}
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   128
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   129
lemma weakening_version1: 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   130
  assumes a: "\<Gamma>1 \<turnstile> t : T" 
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   131
  and     b: "valid \<Gamma>2" 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   132
  and     c: "\<Gamma>1 \<lless> \<Gamma>2"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   133
  shows "\<Gamma>2 \<turnstile> t : T"
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   134
using a b c
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   135
by (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing_induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   136
   (auto | atomize)+
21488
e1b260d204a0 fixed some typos
urbanc
parents: 21487
diff changeset
   137
(* FIXME: meta-quantifiers seem to be not as "automatic" as object-quantifiers *)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   138
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   139
lemma weakening_version2: 
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   140
  fixes \<Gamma>1::"(name\<times>ty) list"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   141
  and   t ::"lam"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   142
  and   \<tau> ::"ty"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   143
  assumes a: "\<Gamma>1 \<turnstile> t:T"
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   144
  and     b: "valid \<Gamma>2" 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   145
  and     c: "\<Gamma>1 \<lless> \<Gamma>2"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   146
  shows "\<Gamma>2 \<turnstile> t:T"
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   147
using a b c
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   148
proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing_induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   149
  case (t_Var \<Gamma>1 a T)  (* variable case *)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   150
  have "\<Gamma>1 \<lless> \<Gamma>2" by fact 
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   151
  moreover  
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   152
  have "valid \<Gamma>2" by fact 
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   153
  moreover 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   154
  have "(a,T)\<in> set \<Gamma>1" by fact
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   155
  ultimately show "\<Gamma>2 \<turnstile> Var a : T" by auto
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   156
next
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   157
  case (t_Lam a \<Gamma>1 T1 T2 t) (* lambda case *)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   158
  have vc: "a\<sharp>\<Gamma>2" by fact (* variable convention *)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   159
  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; ((a,T1)#\<Gamma>1) \<lless> \<Gamma>3\<rbrakk> \<Longrightarrow>  \<Gamma>3 \<turnstile> t:T2" by fact
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   160
  have "\<Gamma>1 \<lless> \<Gamma>2" by fact
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   161
  then have "((a,T1)#\<Gamma>1) \<lless> ((a,T1)#\<Gamma>2)" by simp
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   162
  moreover
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   163
  have "valid \<Gamma>2" by fact
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   164
  then have "valid ((a,T1)#\<Gamma>2)" using vc by (simp add: v2)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   165
  ultimately have "((a,T1)#\<Gamma>2) \<turnstile> t:T2" using ih by simp
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   166
  with vc show "\<Gamma>2 \<turnstile> (Lam [a].t) : T1\<rightarrow>T2" by auto
21488
e1b260d204a0 fixed some typos
urbanc
parents: 21487
diff changeset
   167
qed (auto) (* app case *)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   168
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   169
lemma weakening_version3: 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   170
  assumes a: "\<Gamma>1 \<turnstile> t : T"
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   171
  and     b: "valid \<Gamma>2" 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   172
  and     c: "\<Gamma>1 \<lless> \<Gamma>2"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   173
  shows "\<Gamma>2 \<turnstile> t : T"
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   174
using a b c
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   175
proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing_induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   176
  case (t_Lam a \<Gamma>1 T1 T2 t) (* lambda case *)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   177
  have vc: "a\<sharp>\<Gamma>2" by fact (* variable convention *)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   178
  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; ((a,T1)#\<Gamma>1) \<lless> \<Gamma>3\<rbrakk> \<Longrightarrow>  \<Gamma>3 \<turnstile> t : T2" by fact
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   179
  have "\<Gamma>1 \<lless> \<Gamma>2" by fact
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   180
  then have "((a,T1)#\<Gamma>1) \<lless> ((a,T1)#\<Gamma>2)" by simp
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   181
  moreover
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   182
  have "valid \<Gamma>2" by fact
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   183
  then have "valid ((a,T1)#\<Gamma>2)" using vc by (simp add: v2)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   184
  ultimately have "((a,T1)#\<Gamma>2) \<turnstile> t : T2" using ih by simp
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   185
  with vc show "\<Gamma>2 \<turnstile> (Lam [a].t) : T1 \<rightarrow> T2" by auto
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   186
qed (auto) (* app and var case *)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   187
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   188
text{* The original induction principle for the typing relation
21488
e1b260d204a0 fixed some typos
urbanc
parents: 21487
diff changeset
   189
       is not strong enough - even this simple lemma fails:     *}
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   190
lemma weakening_too_weak: 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   191
  assumes a: "\<Gamma>1 \<turnstile> t : T"
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   192
  and     b: "valid \<Gamma>2" 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   193
  and     c: "\<Gamma>1 \<lless> \<Gamma>2"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   194
  shows "\<Gamma>2 \<turnstile> t : T"
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   195
using a b c
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19972
diff changeset
   196
proof (induct arbitrary: \<Gamma>2)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   197
  case (t_Var \<Gamma>1 a T) (* variable case *)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   198
  have "\<Gamma>1 \<lless> \<Gamma>2" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   199
  moreover
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   200
  have "valid \<Gamma>2" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   201
  moreover
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   202
  have "(a,T) \<in> (set \<Gamma>1)" by fact 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   203
  ultimately show "\<Gamma>2 \<turnstile> Var a : T" by auto
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   204
next
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   205
  case (t_Lam a \<Gamma>1 T1 t T2) (* lambda case *)
21488
e1b260d204a0 fixed some typos
urbanc
parents: 21487
diff changeset
   206
  (* all assumptions available in this case*)
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   207
  have a0: "a\<sharp>\<Gamma>1" by fact
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   208
  have a1: "((a,T1)#\<Gamma>1) \<turnstile> t : T2" by fact
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   209
  have a2: "\<Gamma>1 \<lless> \<Gamma>2" by fact
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   210
  have a3: "valid \<Gamma>2" by fact
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   211
  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; ((a,T1)#\<Gamma>1) \<lless> \<Gamma>3\<rbrakk>  \<Longrightarrow>  \<Gamma>3 \<turnstile> t : T2" by fact
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   212
  have "((a,T1)#\<Gamma>1) \<lless> ((a,T1)#\<Gamma>2)" using a2 by simp
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   213
  moreover
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22231
diff changeset
   214
  have "valid ((a,T1)#\<Gamma>2)" using v2 (* fails *) 
19496
79dbe35c6cba Capitalized theory names.
berghofe
parents: 18654
diff changeset
   215
    oops
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   216
19496
79dbe35c6cba Capitalized theory names.
berghofe
parents: 18654
diff changeset
   217
end