src/HOL/Nominal/Examples/Weakening.thy
author wenzelm
Fri, 17 Nov 2006 02:20:03 +0100
changeset 21404 eb85850d3eb7
parent 21377 c29146dc14f1
child 21405 26b51f724fe6
permissions -rw-r--r--
more robust syntax for definition/abbreviation/notation;
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 
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
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
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    21
lemma [simp]:
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    22
  fixes pi ::"name prm"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    23
  and   \<tau>  ::"ty"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    24
  shows "pi\<bullet>\<tau> = \<tau>"
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    25
by (induct \<tau> rule: ty.induct_weak)
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
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    35
lemma eqvt_valid:
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
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    40
by (induct)
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    41
   (auto simp add: fresh_bij)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    42
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    43
text{* typing judgements *}
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    44
inductive2
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    45
  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
    46
where
3baf57a27269 inductive2: canonical specification syntax;
wenzelm
parents: 21107
diff changeset
    47
    t_Var[intro]: "\<lbrakk>valid \<Gamma>; (a,\<tau>)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var a : \<tau>"
3baf57a27269 inductive2: canonical specification syntax;
wenzelm
parents: 21107
diff changeset
    48
  | t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma>; \<Gamma> \<turnstile> t2 : \<tau>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : \<sigma>"
3baf57a27269 inductive2: canonical specification syntax;
wenzelm
parents: 21107
diff changeset
    49
  | t_Lam[intro]: "\<lbrakk>a\<sharp>\<Gamma>;((a,\<tau>)#\<Gamma>) \<turnstile> t : \<sigma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [a].t : \<tau>\<rightarrow>\<sigma>"
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    50
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    51
lemma eqvt_typing: 
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    52
  fixes pi:: "name prm"
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    53
  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    54
  shows "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>t) : \<tau>"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    55
using a
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    56
proof (induct)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    57
  case (t_Var \<Gamma> a \<tau>)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    58
  have "valid (pi\<bullet>\<Gamma>)" by (rule eqvt_valid)
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    59
  moreover
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    60
  have "(pi\<bullet>(a,\<tau>))\<in>((pi::name prm)\<bullet>set \<Gamma>)" by (rule pt_set_bij2[OF pt_name_inst, OF at_name_inst])
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    61
  ultimately show "(pi\<bullet>\<Gamma>) \<turnstile> ((pi::name prm)\<bullet>Var a) : \<tau>"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    62
    using typing.intros by (force simp add: pt_list_set_pi[OF pt_name_inst, symmetric])
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    63
next 
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    64
  case (t_Lam a \<Gamma> \<tau> t \<sigma>)
19972
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19496
diff changeset
    65
  moreover have "(pi\<bullet>a)\<sharp>(pi\<bullet>\<Gamma>)" by (simp add: fresh_bij)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    66
  ultimately show "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>Lam [a].t) :\<tau>\<rightarrow>\<sigma>" by force 
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    67
qed (auto)
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    68
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    69
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
    70
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    71
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
    72
  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
    73
  and    \<Gamma> :: "(name\<times>ty) list"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    74
  and    t :: "lam"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    75
  and    \<tau> :: "ty"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    76
  and    x :: "'a::fs_name"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
    77
  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    78
  and a1:    "\<And>\<Gamma> a \<tau> x. \<lbrakk>valid \<Gamma>; (a,\<tau>) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P x \<Gamma> (Var a) \<tau>"
18296
3dcc34f18bfa adapted to the new nominal_induction
urbanc
parents: 18269
diff changeset
    79
  and a2:    "\<And>\<Gamma> \<tau> \<sigma> t1 t2 x. 
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    80
              \<lbrakk>\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma>; (\<And>z. P z \<Gamma> t1 (\<tau>\<rightarrow>\<sigma>)); \<Gamma> \<turnstile> t2 : \<tau>; (\<And>z. P z \<Gamma> t2 \<tau>)\<rbrakk>
18296
3dcc34f18bfa adapted to the new nominal_induction
urbanc
parents: 18269
diff changeset
    81
              \<Longrightarrow> P x \<Gamma> (App t1 t2) \<sigma>"
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    82
  and a3:    "\<And>a \<Gamma> \<tau> \<sigma> t x. \<lbrakk>a\<sharp>x; a\<sharp>\<Gamma>; ((a,\<tau>)#\<Gamma>) \<turnstile> t : \<sigma>; (\<And>z. P z ((a,\<tau>)#\<Gamma>) t \<sigma>)\<rbrakk>
18296
3dcc34f18bfa adapted to the new nominal_induction
urbanc
parents: 18269
diff changeset
    83
              \<Longrightarrow> P x \<Gamma> (Lam [a].t) (\<tau>\<rightarrow>\<sigma>)"
3dcc34f18bfa adapted to the new nominal_induction
urbanc
parents: 18269
diff changeset
    84
  shows "P x \<Gamma> t \<tau>"
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
    85
proof -
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
    86
  from a have "\<And>(pi::name prm) x. P x (pi\<bullet>\<Gamma>) (pi\<bullet>t) \<tau>"
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
    87
  proof (induct)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    88
    case (t_Var \<Gamma> a \<tau>)
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    89
    have "valid \<Gamma>" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    90
    then have "valid (pi\<bullet>\<Gamma>)" by (rule eqvt_valid)
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    91
    moreover
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    92
    have "(a,\<tau>)\<in>set \<Gamma>" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    93
    then have "pi\<bullet>(a,\<tau>)\<in>pi\<bullet>(set \<Gamma>)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst])  
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    94
    then have "(pi\<bullet>a,\<tau>)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: pt_list_set_pi[OF pt_name_inst])
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    95
    ultimately show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Var a)) \<tau>" using a1 by simp
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
    96
  next
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    97
    case (t_App \<Gamma> t1 \<tau> \<sigma> t2)
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
    98
    thus "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(App t1 t2)) \<sigma>" using a2 by (simp, blast intro: eqvt_typing)
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
    99
  next
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   100
    case (t_Lam a \<Gamma> \<tau> t \<sigma>)
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   101
    have k1: "a\<sharp>\<Gamma>" by fact
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   102
    have k2: "((a,\<tau>)#\<Gamma>)\<turnstile>t:\<sigma>" by fact
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   103
    have k3: "\<And>(pi::name prm) (x::'a::fs_name). P x (pi \<bullet>((a,\<tau>)#\<Gamma>)) (pi\<bullet>t) \<sigma>" by fact
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   104
    have f: "\<exists>c::name. c\<sharp>(pi\<bullet>a,pi\<bullet>t,pi\<bullet>\<Gamma>,x)"
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21364
diff changeset
   105
      by (rule exists_fresh', simp add: fs_name1)
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   106
    then obtain c::"name" 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   107
      where f1: "c\<noteq>(pi\<bullet>a)" and f2: "c\<sharp>x" and f3: "c\<sharp>(pi\<bullet>t)" and f4: "c\<sharp>(pi\<bullet>\<Gamma>)"
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   108
      by (force simp add: fresh_prod fresh_atm)
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   109
    from k1 have k1a: "(pi\<bullet>a)\<sharp>(pi\<bullet>\<Gamma>)"  by (simp add: fresh_bij)
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   110
    have l1: "(([(c,pi\<bullet>a)]@pi)\<bullet>\<Gamma>) = (pi\<bullet>\<Gamma>)" using f4 k1a 
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   111
      by (simp only: pt_name2, rule perm_fresh_fresh)
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   112
    have "\<And>x. P x (([(c,pi\<bullet>a)]@pi)\<bullet>((a,\<tau>)#\<Gamma>)) (([(c,pi\<bullet>a)]@pi)\<bullet>t) \<sigma>" using k3 by force
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   113
    hence l2: "\<And>x. P x ((c, \<tau>)#(pi\<bullet>\<Gamma>)) (([(c,pi\<bullet>a)]@pi)\<bullet>t) \<sigma>" using f1 l1
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   114
      by (force simp add: pt_name2  calc_atm)
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   115
    have "(([(c,pi\<bullet>a)]@pi)\<bullet>((a,\<tau>)#\<Gamma>)) \<turnstile> (([(c,pi\<bullet>a)]@pi)\<bullet>t) : \<sigma>" using k2 by (rule eqvt_typing)
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   116
    hence l3: "((c, \<tau>)#(pi\<bullet>\<Gamma>)) \<turnstile> (([(c,pi\<bullet>a)]@pi)\<bullet>t) : \<sigma>" using l1 f1 
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   117
      by (force simp add: pt_name2  calc_atm)
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   118
    have l4: "P x (pi\<bullet>\<Gamma>) (Lam [c].(([(c,pi\<bullet>a)]@pi)\<bullet>t)) (\<tau> \<rightarrow> \<sigma>)" using f2 f4 l2 l3 a3 by auto
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   119
    have alpha: "(Lam [c].([(c,pi\<bullet>a)]\<bullet>(pi\<bullet>t))) = (Lam [(pi\<bullet>a)].(pi\<bullet>t))" using f1 f3
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   120
      by (simp add: lam.inject alpha)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   121
    show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Lam [a].t)) (\<tau> \<rightarrow> \<sigma>)" using l4 alpha by (simp only: pt_name2, simp)
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   122
  qed
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   123
  hence "P x (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>t) \<tau>" by blast
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   124
  thus "P x \<Gamma> t \<tau>" by simp
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   125
qed
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   126
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   127
lemma typing_induct_test[consumes 1, case_names t_Var t_App t_Lam]:
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   128
  fixes  P :: "'a::fs_name\<Rightarrow>(name\<times>ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow>bool"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   129
  and    \<Gamma> :: "(name\<times>ty) list"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   130
  and    t :: "lam"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   131
  and    \<tau> :: "ty"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   132
  and    x :: "'a::fs_name"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   133
  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   134
  and a1:    "\<And>\<Gamma> a \<tau> x. \<lbrakk>valid \<Gamma>; (a,\<tau>) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P x \<Gamma> (Var a) \<tau>"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   135
  and a2:    "\<And>\<Gamma> \<tau> \<sigma> t1 t2 x. 
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   136
              \<lbrakk>\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma>; \<And>z. P z \<Gamma> t1 (\<tau>\<rightarrow>\<sigma>); \<Gamma> \<turnstile> t2 : \<tau>; \<And>z. P z \<Gamma> t2 \<tau>\<rbrakk>
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   137
              \<Longrightarrow> P x \<Gamma> (App t1 t2) \<sigma>"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   138
  and a3:    "\<And>a \<Gamma> \<tau> \<sigma> t x. \<lbrakk>a\<sharp>x; a\<sharp>\<Gamma>; ((a,\<tau>)#\<Gamma>) \<turnstile> t : \<sigma>; \<And>z. P z ((a,\<tau>)#\<Gamma>) t \<sigma>\<rbrakk>
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   139
              \<Longrightarrow> P x \<Gamma> (Lam [a].t) (\<tau>\<rightarrow>\<sigma>)"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   140
  shows "P x \<Gamma> t \<tau>"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   141
proof -
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   142
  from a have "\<And>(pi::name prm) x. P x (pi\<bullet>\<Gamma>) (pi\<bullet>t) \<tau>"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   143
  proof (induct)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   144
    case (t_Var \<Gamma> a \<tau>)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   145
    have "valid \<Gamma>" by fact
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   146
    then have "valid (pi\<bullet>\<Gamma>)" by (rule eqvt_valid)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   147
    moreover
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   148
    have "(a,\<tau>)\<in>set \<Gamma>" by fact
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   149
    then have "pi\<bullet>(a,\<tau>)\<in>pi\<bullet>(set \<Gamma>)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst])  
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   150
    then have "(pi\<bullet>a,\<tau>)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: pt_list_set_pi[OF pt_name_inst])
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   151
    ultimately show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Var a)) \<tau>" using a1 by simp
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   152
  next
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   153
    case (t_App \<Gamma> t1 \<tau> \<sigma> t2)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   154
    thus "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(App t1 t2)) \<sigma>" using a2 by (simp, blast intro: eqvt_typing)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   155
  next
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   156
    case (t_Lam a \<Gamma> \<tau> t \<sigma> pi x)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   157
    have p1: "((a,\<tau>)#\<Gamma>)\<turnstile>t:\<sigma>" by fact
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   158
    have ih1: "\<And>(pi::name prm) x. P x (pi\<bullet>((a,\<tau>)#\<Gamma>)) (pi\<bullet>t) \<sigma>" by fact
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   159
    have f: "a\<sharp>\<Gamma>" by fact
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   160
    then have f': "(pi\<bullet>a)\<sharp>(pi\<bullet>\<Gamma>)"  by (simp add: fresh_bij)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   161
    have "\<exists>c::name. c\<sharp>(pi\<bullet>a,pi\<bullet>t,pi\<bullet>\<Gamma>,x)"
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21364
diff changeset
   162
      by (rule exists_fresh', simp add: fs_name1)
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   163
    then obtain c::"name" 
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   164
      where fs: "c\<noteq>(pi\<bullet>a)" "c\<sharp>x" "c\<sharp>(pi\<bullet>t)" "c\<sharp>(pi\<bullet>\<Gamma>)"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   165
      by (force simp add: fresh_prod fresh_atm)    
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   166
    let ?pi'="[(pi\<bullet>a,c)]@pi"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   167
    have eq: "((pi\<bullet>a,c)#pi)\<bullet>a = c" by (simp add: calc_atm)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   168
    have p1': "(?pi'\<bullet>((a,\<tau>)#\<Gamma>))\<turnstile>(?pi'\<bullet>t):\<sigma>" using p1 by (simp only: eqvt_typing)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   169
    have ih1': "\<And>x. P x (?pi'\<bullet>((a,\<tau>)#\<Gamma>)) (?pi'\<bullet>t) \<sigma>" using ih1 by simp
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   170
    have "P x (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>(Lam [a].t)) (\<tau>\<rightarrow>\<sigma>)" using f f' fs p1' ih1' eq
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   171
      apply -
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   172
      apply(simp del: append_Cons)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   173
      apply(rule a3)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   174
      apply(simp_all add: fresh_left calc_atm pt_name2)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   175
      done
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   176
    then have "P x ([(pi\<bullet>a,c)]\<bullet>(pi\<bullet>\<Gamma>)) ([(pi\<bullet>a,c)]\<bullet>(Lam [(pi\<bullet>a)].(pi\<bullet>t))) (\<tau>\<rightarrow>\<sigma>)" 
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   177
      by (simp del: append_Cons add: pt_name2)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   178
    then show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Lam [a].t)) (\<tau> \<rightarrow> \<sigma>)" using f f' fs      
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   179
      apply -
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   180
      apply(subgoal_tac "c\<sharp>Lam [(pi\<bullet>a)].(pi\<bullet>t)")
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   181
      apply(subgoal_tac "(pi\<bullet>a)\<sharp>Lam [(pi\<bullet>a)].(pi\<bullet>t)")
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   182
      apply(simp only: perm_fresh_fresh)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   183
      apply(simp)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   184
      apply(simp add: abs_fresh)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   185
      apply(simp add: abs_fresh)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   186
      done 
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   187
  qed
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   188
  hence "P x (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>t) \<tau>" by blast
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   189
  thus "P x \<Gamma> t \<tau>" by simp
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   190
qed
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   191
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 21052
diff changeset
   192
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   193
text {* definition of a subcontext *}
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   194
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   195
abbreviation
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21377
diff changeset
   196
  "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
   197
  "\<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
   198
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   199
text {* Now it comes: The Weakening Lemma *}
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   200
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   201
lemma weakening_version1: 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   202
  assumes a: "\<Gamma>1 \<turnstile> t : \<sigma>" 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   203
  and     b: "valid \<Gamma>2" 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   204
  and     c: "\<Gamma>1 \<lless> \<Gamma>2"
18296
3dcc34f18bfa adapted to the new nominal_induction
urbanc
parents: 18269
diff changeset
   205
  shows "\<Gamma>2 \<turnstile> t:\<sigma>"
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   206
using a b c
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   207
apply(nominal_induct \<Gamma>1 t \<sigma> avoiding: \<Gamma>2 rule: typing_induct)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   208
apply(auto | atomize)+
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   209
(* FIXME: meta-quantifiers seem to not ba as "automatic" as object-quantifiers *)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   210
done
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   211
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   212
lemma weakening_version2: 
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   213
  fixes \<Gamma>1::"(name\<times>ty) list"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   214
  and   t ::"lam"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   215
  and   \<tau> ::"ty"
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   216
  assumes a: "\<Gamma>1 \<turnstile> t:\<sigma>"
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   217
  and     b: "valid \<Gamma>2" 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   218
  and     c: "\<Gamma>1 \<lless> \<Gamma>2"
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   219
  shows "\<Gamma>2 \<turnstile> t:\<sigma>"
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   220
using a b c
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   221
proof (nominal_induct \<Gamma>1 t \<sigma> avoiding: \<Gamma>2 rule: typing_induct)
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   222
  case (t_Var \<Gamma>1 a \<tau>)  (* variable case *)
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   223
  have "\<Gamma>1 \<lless> \<Gamma>2" by fact 
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   224
  moreover  
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   225
  have "valid \<Gamma>2" by fact 
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   226
  moreover 
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   227
  have "(a,\<tau>)\<in> set \<Gamma>1" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   228
  ultimately show "\<Gamma>2 \<turnstile> Var a : \<tau>" by auto
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   229
next
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   230
  case (t_Lam a \<Gamma>1 \<tau> \<sigma> t) (* lambda case *)
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   231
  have vc: "a\<sharp>\<Gamma>2" by fact (* variable convention *)
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   232
  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; ((a,\<tau>)#\<Gamma>1) \<lless> \<Gamma>3\<rbrakk> \<Longrightarrow>  \<Gamma>3 \<turnstile> t:\<sigma>" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   233
  have "\<Gamma>1 \<lless> \<Gamma>2" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   234
  then have "((a,\<tau>)#\<Gamma>1) \<lless> ((a,\<tau>)#\<Gamma>2)" by simp
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   235
  moreover
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   236
  have "valid \<Gamma>2" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   237
  then have "valid ((a,\<tau>)#\<Gamma>2)" using vc v2 by simp
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   238
  ultimately have "((a,\<tau>)#\<Gamma>2) \<turnstile> t:\<sigma>" using ih by simp
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   239
  with vc show "\<Gamma>2 \<turnstile> (Lam [a].t) : \<tau> \<rightarrow> \<sigma>" by auto
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   240
qed (auto)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   241
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   242
lemma weakening_version3: 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   243
  assumes a: "\<Gamma>1 \<turnstile> t:\<sigma>"
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   244
  and     b: "valid \<Gamma>2" 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   245
  and     c: "\<Gamma>1 \<lless> \<Gamma>2"
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   246
  shows "\<Gamma>2 \<turnstile> t:\<sigma>"
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   247
using a b c
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   248
proof (nominal_induct \<Gamma>1 t \<sigma> avoiding: \<Gamma>2 rule: typing_induct)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   249
  case (t_Lam a \<Gamma>1 \<tau> \<sigma> t) (* lambda case *)
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   250
  have vc: "a\<sharp>\<Gamma>2" by fact (* variable convention *)
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   251
  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; ((a,\<tau>)#\<Gamma>1) \<lless> \<Gamma>3\<rbrakk> \<Longrightarrow>  \<Gamma>3 \<turnstile> t:\<sigma>" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   252
  have "\<Gamma>1 \<lless> \<Gamma>2" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   253
  then have "((a,\<tau>)#\<Gamma>1) \<lless> ((a,\<tau>)#\<Gamma>2)" by simp
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   254
  moreover
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   255
  have "valid \<Gamma>2" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   256
  then have "valid ((a,\<tau>)#\<Gamma>2)" using vc v2 by simp
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   257
  ultimately have "((a,\<tau>)#\<Gamma>2) \<turnstile> t:\<sigma>" using ih by simp
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   258
  with vc show "\<Gamma>2 \<turnstile> (Lam [a].t) : \<tau> \<rightarrow> \<sigma>" by auto
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   259
qed (auto) (* app and var case *)
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   260
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   261
text{* The original induction principle for the typing relation
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   262
       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
   263
lemma weakening_too_weak: 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   264
  assumes a: "\<Gamma>1 \<turnstile> t:\<sigma>"
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   265
  and     b: "valid \<Gamma>2" 
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   266
  and     c: "\<Gamma>1 \<lless> \<Gamma>2"
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   267
  shows "\<Gamma>2 \<turnstile> t:\<sigma>"
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   268
using a b c
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19972
diff changeset
   269
proof (induct arbitrary: \<Gamma>2)
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   270
  case (t_Var \<Gamma>1 a \<tau>) (* variable case *)
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   271
  have "\<Gamma>1 \<lless> \<Gamma>2" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   272
  moreover
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   273
  have "valid \<Gamma>2" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   274
  moreover
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   275
  have "(a,\<tau>) \<in> (set \<Gamma>1)" by fact 
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   276
  ultimately show "\<Gamma>2 \<turnstile> Var a : \<tau>" by auto
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   277
next
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   278
  case (t_Lam a \<Gamma>1 \<tau> t \<sigma>) (* lambda case *)
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   279
  (* all assumption in this case*)
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   280
  have a0: "a\<sharp>\<Gamma>1" by fact
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   281
  have a1: "((a,\<tau>)#\<Gamma>1) \<turnstile> t : \<sigma>" by fact
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   282
  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
   283
  have a3: "valid \<Gamma>2" by fact
21052
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   284
  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; ((a,\<tau>)#\<Gamma>1) \<lless> \<Gamma>3\<rbrakk>  \<Longrightarrow>  \<Gamma>3 \<turnstile> t:\<sigma>" by fact
ec5531061ed6 adapted to Stefan's new inductive package and cleaning up
urbanc
parents: 20955
diff changeset
   285
  have "((a,\<tau>)#\<Gamma>1) \<lless> ((a,\<tau>)#\<Gamma>2)" using a2 by simp
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   286
  moreover
18311
b83b00cbaecf changed "fresh:" to "avoiding:" and cleaned up the weakening example
urbanc
parents: 18296
diff changeset
   287
  have "valid ((a,\<tau>)#\<Gamma>2)" using v2 (* fails *) 
19496
79dbe35c6cba Capitalized theory names.
berghofe
parents: 18654
diff changeset
   288
    oops
18105
4c9c081a416b Initial commit of the theory "Weakening".
urbanc
parents:
diff changeset
   289
19496
79dbe35c6cba Capitalized theory names.
berghofe
parents: 18654
diff changeset
   290
end